src/HOL/IntDef.thy
author haftmann
Tue, 30 Oct 2007 08:45:54 +0100
changeset 25230 022029099a83
parent 25193 e2e1a4b00de3
child 25349 0d46bea01741
permissions -rw-r--r--
continued localization
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      IntDef.thy
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     6
*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     7
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     8
header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     9
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    10
theory IntDef
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    11
imports Equiv_Relations Nat
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    12
begin
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    13
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    14
text {* the equivalence relation underlying the integers *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    16
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    17
  intrel :: "((nat \<times> nat) \<times> (nat \<times> nat)) set"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    18
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    19
  "intrel = {((x, y), (u, v)) | x y u v. x + v = u +y }"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    20
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    21
typedef (Integ)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    22
  int = "UNIV//intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    23
  by (auto simp add: quotient_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    24
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    25
instance int :: zero
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
    26
  Zero_int_def: "0 \<equiv> Abs_Integ (intrel `` {(0, 0)})" ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    27
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    28
instance int :: one
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
    29
  One_int_def: "1 \<equiv> Abs_Integ (intrel `` {(1, 0)})" ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    30
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    31
instance int :: plus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    32
  add_int_def: "z + w \<equiv> Abs_Integ
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    33
    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u, v) \<in> Rep_Integ w.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    34
      intrel `` {(x + u, y + v)})" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    35
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    36
instance int :: minus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
  minus_int_def:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
    "- z \<equiv> Abs_Integ (\<Union>(x, y) \<in> Rep_Integ z. intrel `` {(y, x)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    39
  diff_int_def:  "z - w \<equiv> z + (-w)" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    40
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    41
instance int :: times
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    42
  mult_int_def: "z * w \<equiv>  Abs_Integ
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    43
    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u,v ) \<in> Rep_Integ w.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    44
      intrel `` {(x*u + y*v, x*v + y*u)})" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    45
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    46
instance int :: ord
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
  le_int_def:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    48
   "z \<le> w \<equiv> \<exists>x y u v. x+v \<le> u+y \<and> (x, y) \<in> Rep_Integ z \<and> (u, v) \<in> Rep_Integ w"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    49
  less_int_def: "z < w \<equiv> z \<le> w \<and> z \<noteq> w" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    50
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    51
lemmas [code func del] = Zero_int_def One_int_def add_int_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    52
  minus_int_def mult_int_def le_int_def less_int_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    53
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    54
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    55
subsection{*Construction of the Integers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    56
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    57
lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    58
by (simp add: intrel_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    59
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    60
lemma equiv_intrel: "equiv UNIV intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    61
by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    62
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    63
text{*Reduces equality of equivalence classes to the @{term intrel} relation:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    64
  @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    65
lemmas equiv_intrel_iff [simp] = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    66
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    67
text{*All equivalence classes belong to set of representatives*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    68
lemma [simp]: "intrel``{(x,y)} \<in> Integ"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    69
by (auto simp add: Integ_def intrel_def quotient_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    70
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    71
text{*Reduces equality on abstractions to equality on representatives:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    72
  @{prop "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24196
diff changeset
    73
declare Abs_Integ_inject [simp,noatp]  Abs_Integ_inverse [simp,noatp]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    74
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    75
text{*Case analysis on the representation of an integer as an equivalence
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    76
      class of pairs of naturals.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    77
lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    78
     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    79
apply (rule Abs_Integ_cases [of z]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    80
apply (auto simp add: Integ_def quotient_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    81
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    82
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    83
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
    84
subsection{*Arithmetic Operations*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    85
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    87
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    88
  have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    89
    by (simp add: congruent_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    90
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    91
    by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    93
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    94
lemma add:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    95
     "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    96
      Abs_Integ (intrel``{(x+u, y+v)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    98
  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    99
        respects2 intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
    by (simp add: congruent2_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   101
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
    by (simp add: add_int_def UN_UN_split_split_eq
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
                  UN_equiv_class2 [OF equiv_intrel equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   104
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   105
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   106
text{*Congruence property for multiplication*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   107
lemma mult_congruent2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
     "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
      respects2 intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
apply (rule equiv_intrel [THEN congruent2_commuteI])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
 apply (force simp add: mult_ac, clarify) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
apply (simp add: congruent_def mult_ac)  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
apply (rename_tac u v w x y z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
apply (subgoal_tac "u*y + x*y = w*y + v*y  &  u*z + x*z = w*z + v*z")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   115
apply (simp add: mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   116
apply (simp add: add_mult_distrib [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   118
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   119
lemma mult:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   120
     "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   121
      Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   122
by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
              UN_equiv_class2 [OF equiv_intrel equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   124
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   125
text{*The integers form a @{text comm_ring_1}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   126
instance int :: comm_ring_1
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   127
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   128
  fix i j k :: int
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   129
  show "(i + j) + k = i + (j + k)"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   130
    by (cases i, cases j, cases k) (simp add: add add_assoc)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   131
  show "i + j = j + i" 
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   132
    by (cases i, cases j) (simp add: add_ac add)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   133
  show "0 + i = i"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   134
    by (cases i) (simp add: Zero_int_def add)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   135
  show "- i + i = 0"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   136
    by (cases i) (simp add: Zero_int_def minus add)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   137
  show "i - j = i + - j"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   138
    by (simp add: diff_int_def)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   139
  show "(i * j) * k = i * (j * k)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23438
diff changeset
   140
    by (cases i, cases j, cases k) (simp add: mult ring_simps)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   141
  show "i * j = j * i"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23438
diff changeset
   142
    by (cases i, cases j) (simp add: mult ring_simps)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   143
  show "1 * i = i"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   144
    by (cases i) (simp add: One_int_def mult)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   145
  show "(i + j) * k = i * k + j * k"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23438
diff changeset
   146
    by (cases i, cases j, cases k) (simp add: add mult ring_simps)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   147
  show "0 \<noteq> (1::int)"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   148
    by (simp add: Zero_int_def One_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   149
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   150
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   151
lemma int_def: "of_nat m = Abs_Integ (intrel `` {(m, 0)})"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   152
by (induct m, simp_all add: Zero_int_def One_int_def add)
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   153
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   154
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   155
subsection{*The @{text "\<le>"} Ordering*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   156
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   157
lemma le:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   158
  "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   159
by (force simp add: le_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   160
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   161
lemma less:
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   162
  "(Abs_Integ(intrel``{(x,y)}) < Abs_Integ(intrel``{(u,v)})) = (x+v < u+y)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   163
by (simp add: less_int_def le order_less_le)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   164
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   165
instance int :: linorder
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   166
proof
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   167
  fix i j k :: int
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   168
  show "(i < j) = (i \<le> j \<and> i \<noteq> j)"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   169
    by (simp add: less_int_def)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   170
  show "i \<le> i"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   171
    by (cases i) (simp add: le)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   172
  show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   173
    by (cases i, cases j, cases k) (simp add: le)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   174
  show "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   175
    by (cases i, cases j) (simp add: le)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   176
  show "i \<le> j \<or> j \<le> i"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   177
    by (cases i, cases j) (simp add: le linorder_linear)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   178
qed
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   179
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   180
instance int :: pordered_cancel_ab_semigroup_add
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   181
proof
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   182
  fix i j k :: int
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   183
  show "i \<le> j \<Longrightarrow> k + i \<le> k + j"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   184
    by (cases i, cases j, cases k) (simp add: le add)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   185
qed
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   186
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   187
text{*Strict Monotonicity of Multiplication*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   188
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   189
text{*strict, in 1st argument; proof is by induction on k>0*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   190
lemma zmult_zless_mono2_lemma:
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   191
     "(i::int)<j ==> 0<k ==> of_nat k * i < of_nat k * j"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
apply (induct "k", simp)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   193
apply (simp add: left_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
apply (case_tac "k=0")
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   195
apply (simp_all add: add_strict_mono)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   196
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   197
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   198
lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = of_nat n"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   199
apply (cases k)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   200
apply (auto simp add: le add int_def Zero_int_def)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   201
apply (rule_tac x="x-y" in exI, simp)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   202
done
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   203
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   204
lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = of_nat n"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   205
apply (cases k)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   206
apply (simp add: less int_def Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   207
apply (rule_tac x="x-y" in exI, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   208
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   209
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   210
lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   211
apply (drule zero_less_imp_eq_int)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   212
apply (auto simp add: zmult_zless_mono2_lemma)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   213
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   214
23879
4776af8be741 split class abs from class minus
haftmann
parents: 23852
diff changeset
   215
instance int :: abs
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   216
  zabs_def: "\<bar>i\<Colon>int\<bar> \<equiv> if i < 0 then - i else i" ..
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24355
diff changeset
   217
instance int :: sgn
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24355
diff changeset
   218
  zsgn_def: "sgn(i\<Colon>int) \<equiv> (if i=0 then 0 else if 0<i then 1 else - 1)" ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   219
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   220
instance int :: distrib_lattice
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   221
  "inf \<equiv> min"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   222
  "sup \<equiv> max"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   223
  by intro_classes
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   224
    (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   225
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   226
text{*The integers form an ordered integral domain*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   227
instance int :: ordered_idom
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   228
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   229
  fix i j k :: int
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   230
  show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   231
    by (rule zmult_zless_mono2)
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   232
  show "\<bar>i\<bar> = (if i < 0 then -i else i)"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   233
    by (simp only: zabs_def)
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24355
diff changeset
   234
  show "sgn(i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24355
diff changeset
   235
    by (simp only: zsgn_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   236
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   237
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   238
lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   239
apply (cases w, cases z) 
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   240
apply (simp add: less le add One_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   241
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   242
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   243
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   244
subsection{*Magnitude of an Integer, as a Natural Number: @{term nat}*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   245
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   246
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   247
  nat :: "int \<Rightarrow> nat"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   248
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   249
  [code func del]: "nat z = contents (\<Union>(x, y) \<in> Rep_Integ z. {x-y})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   250
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   251
lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   252
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   253
  have "(\<lambda>(x,y). {x-y}) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   254
    by (simp add: congruent_def) arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   255
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   256
    by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   257
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   258
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   259
lemma nat_int [simp]: "nat (of_nat n) = n"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   260
by (simp add: nat int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   261
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   262
lemma nat_zero [simp]: "nat 0 = 0"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   263
by (simp add: Zero_int_def nat)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   264
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   265
lemma int_nat_eq [simp]: "of_nat (nat z) = (if 0 \<le> z then z else 0)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   266
by (cases z, simp add: nat le int_def Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   267
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   268
corollary nat_0_le: "0 \<le> z ==> of_nat (nat z) = z"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   269
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   270
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   271
lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   272
by (cases z, simp add: nat le Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   273
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   274
lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   275
apply (cases w, cases z) 
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   276
apply (simp add: nat le linorder_not_le [symmetric] Zero_int_def, arith)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   277
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   278
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   279
text{*An alternative condition is @{term "0 \<le> w"} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   280
corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   281
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   282
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   283
corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   284
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   285
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   286
lemma zless_nat_conj [simp]: "(nat w < nat z) = (0 < z & w < z)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   287
apply (cases w, cases z) 
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   288
apply (simp add: nat le Zero_int_def linorder_not_le [symmetric], arith)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   289
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   290
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   291
lemma nonneg_eq_int:
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   292
  fixes z :: int
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   293
  assumes "0 \<le> z" and "\<And>m. z = of_nat m \<Longrightarrow> P"
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   294
  shows P
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   295
  using assms by (blast dest: nat_0_le sym)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   296
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   297
lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = of_nat m else m=0)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   298
by (cases w, simp add: nat le int_def Zero_int_def, arith)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   299
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   300
corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = of_nat m else m=0)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   301
by (simp only: eq_commute [of m] nat_eq_iff)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   302
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   303
lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < of_nat m)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   304
apply (cases w)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   305
apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   306
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   307
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   308
lemma int_eq_iff: "(of_nat m = z) = (m = nat z & 0 \<le> z)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   309
by (auto simp add: nat_eq_iff2)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   310
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   311
lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   312
by (insert zless_nat_conj [of 0], auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   313
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   314
lemma nat_add_distrib:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   315
     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   316
by (cases z, cases z', simp add: nat add le Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   317
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   318
lemma nat_diff_distrib:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   319
     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   320
by (cases z, cases z', 
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   321
    simp add: nat add minus diff_minus le Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   322
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   323
lemma nat_zminus_int [simp]: "nat (- (of_nat n)) = 0"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   324
by (simp add: int_def minus nat Zero_int_def) 
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   325
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   326
lemma zless_nat_eq_int_zless: "(m < nat z) = (of_nat m < z)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   327
by (cases z, simp add: nat less int_def, arith)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   328
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   329
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   330
subsection{*Lemmas about the Function @{term of_nat} and Orderings*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   331
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   332
lemma negative_zless_0: "- (of_nat (Suc n)) < (0 \<Colon> int)"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   333
by (simp add: order_less_le del: of_nat_Suc)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   334
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   335
lemma negative_zless [iff]: "- (of_nat (Suc n)) < (of_nat m \<Colon> int)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   336
by (rule negative_zless_0 [THEN order_less_le_trans], simp)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   337
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   338
lemma negative_zle_0: "- of_nat n \<le> (0 \<Colon> int)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   339
by (simp add: minus_le_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   340
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   341
lemma negative_zle [iff]: "- of_nat n \<le> (of_nat m \<Colon> int)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   342
by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   343
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   344
lemma not_zle_0_negative [simp]: "~ (0 \<le> - (of_nat (Suc n) \<Colon> int))"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   345
by (subst le_minus_iff, simp del: of_nat_Suc)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   346
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   347
lemma int_zle_neg: "((of_nat n \<Colon> int) \<le> - of_nat m) = (n = 0 & m = 0)"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   348
by (simp add: int_def le minus Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   349
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   350
lemma not_int_zless_negative [simp]: "~ ((of_nat n \<Colon> int) < - of_nat m)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   351
by (simp add: linorder_not_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   352
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   353
lemma negative_eq_positive [simp]: "((- of_nat n \<Colon> int) = of_nat m) = (n = 0 & m = 0)"
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   354
by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   355
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   356
lemma zle_iff_zadd: "(w\<Colon>int) \<le> z \<longleftrightarrow> (\<exists>n. z = w + of_nat n)"
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   357
proof -
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   358
  have "(w \<le> z) = (0 \<le> z - w)"
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   359
    by (simp only: le_diff_eq add_0_left)
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   360
  also have "\<dots> = (\<exists>n. z - w = of_nat n)"
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   361
    by (auto elim: zero_le_imp_eq_int)
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   362
  also have "\<dots> = (\<exists>n. z = w + of_nat n)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23438
diff changeset
   363
    by (simp only: group_simps)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   364
  finally show ?thesis .
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   365
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   366
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   367
lemma zadd_int_left: "of_nat m + (of_nat n + z) = of_nat (m + n) + (z\<Colon>int)"
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   368
by simp
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   369
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   370
lemma int_Suc0_eq_1: "of_nat (Suc 0) = (1\<Colon>int)"
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   371
by simp
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   372
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   373
text{*This version is proved for all ordered rings, not just integers!
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   374
      It is proved here because attribute @{text arith_split} is not available
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   375
      in theory @{text Ring_and_Field}.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   376
      But is it really better than just rewriting with @{text abs_if}?*}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24196
diff changeset
   377
lemma abs_split [arith_split,noatp]:
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   378
     "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   379
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   380
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   381
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   382
subsection {* Constants @{term neg} and @{term iszero} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   383
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   384
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   385
  neg  :: "'a\<Colon>ordered_idom \<Rightarrow> bool"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   386
where
25164
0fcb4775cbfb dropped superfluous inlining rule
haftmann
parents: 24728
diff changeset
   387
  "neg Z \<longleftrightarrow> Z < 0"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   388
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   389
definition (*for simplifying equalities*)
23276
a70934b63910 generalize of_nat and related constants to class semiring_1
huffman
parents: 23164
diff changeset
   390
  iszero :: "'a\<Colon>semiring_1 \<Rightarrow> bool"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   391
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   392
  "iszero z \<longleftrightarrow> z = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   393
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   394
lemma not_neg_int [simp]: "~ neg (of_nat n)"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   395
by (simp add: neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   396
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   397
lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   398
by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   399
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   400
lemmas neg_eq_less_0 = neg_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   401
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   402
lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   403
by (simp add: neg_def linorder_not_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   404
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   405
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   406
text{*To simplify inequalities when Numeral1 can get simplified to 1*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   407
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   408
lemma not_neg_0: "~ neg 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   409
by (simp add: One_int_def neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   410
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   411
lemma not_neg_1: "~ neg 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   412
by (simp add: neg_def linorder_not_less zero_le_one)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   413
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   414
lemma iszero_0: "iszero 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   415
by (simp add: iszero_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   416
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   417
lemma not_iszero_1: "~ iszero 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   418
by (simp add: iszero_def eq_commute)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   419
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   420
lemma neg_nat: "neg z ==> nat z = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   421
by (simp add: neg_def order_less_imp_le) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   422
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   423
lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   424
by (simp add: linorder_not_less neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   425
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   426
23852
3736cdf9398b moved set Nats to Nat.thy
haftmann
parents: 23705
diff changeset
   427
subsection{*Embedding of the Integers into any @{text ring_1}: @{term of_int}*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   428
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   429
context ring_1
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   430
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   431
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   432
term of_nat
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   433
23950
f54c0e339061 dropped axclass
haftmann
parents: 23879
diff changeset
   434
definition
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   435
  of_int :: "int \<Rightarrow> 'a"
23950
f54c0e339061 dropped axclass
haftmann
parents: 23879
diff changeset
   436
where
f54c0e339061 dropped axclass
haftmann
parents: 23879
diff changeset
   437
  "of_int z = contents (\<Union>(i, j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
23852
3736cdf9398b moved set Nats to Nat.thy
haftmann
parents: 23705
diff changeset
   438
lemmas [code func del] = of_int_def
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   439
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   440
lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   441
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   442
  have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   443
    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   444
            del: of_nat_add) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   445
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   446
    by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   447
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   448
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   449
lemma of_int_0 [simp]: "of_int 0 = 0"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   450
by (simp add: of_int Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   451
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   452
lemma of_int_1 [simp]: "of_int 1 = 1"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   453
by (simp add: of_int One_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   454
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   455
lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   456
by (cases w, cases z, simp add: compare_rls of_int add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   457
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   458
lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   459
by (cases z, simp add: compare_rls of_int minus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   460
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   461
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   462
apply (cases w, cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   463
apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   464
                 mult add_ac of_nat_mult)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   465
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   466
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   467
text{*Collapse nested embeddings*}
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   468
lemma of_int_of_nat_eq [simp]: "of_int (Nat.of_nat n) = of_nat n"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   469
  by (induct n, auto)
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   470
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   471
end
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   472
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   473
lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   474
by (simp add: diff_minus)
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   475
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   476
lemma of_int_le_iff [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   477
     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   478
apply (cases w)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   479
apply (cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   480
apply (simp add: compare_rls of_int le diff_int_def add minus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   481
                 of_nat_add [symmetric]   del: of_nat_add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   482
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   483
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   484
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   485
lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   486
lemmas of_int_le_0_iff [simp] = of_int_le_iff [of _ 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   487
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   488
lemma of_int_less_iff [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   489
     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   490
by (simp add: linorder_not_le [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   491
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   492
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   493
lemmas of_int_0_less_iff [simp] = of_int_less_iff [of 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   494
lemmas of_int_less_0_iff [simp] = of_int_less_iff [of _ 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   495
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   496
text{*Class for unital rings with characteristic zero.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   497
 Includes non-ordered rings like the complex numbers.*}
23950
f54c0e339061 dropped axclass
haftmann
parents: 23879
diff changeset
   498
class ring_char_0 = ring_1 + semiring_char_0
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   499
begin
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   500
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   501
lemma of_int_eq_iff [simp]:
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   502
   "of_int w = of_int z \<longleftrightarrow> w = z"
23282
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   503
apply (cases w, cases z, simp add: of_int)
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   504
apply (simp only: diff_eq_eq diff_add_eq eq_diff_eq)
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   505
apply (simp only: of_nat_add [symmetric] of_nat_eq_iff)
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   506
done
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   507
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   508
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   509
lemmas of_int_0_eq_iff [simp] = of_int_eq_iff [of 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   510
lemmas of_int_eq_0_iff [simp] = of_int_eq_iff [of _ 0, simplified]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   511
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   512
end
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   513
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   514
text{*Every @{text ordered_idom} has characteristic zero.*}
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   515
instance ordered_idom \<subseteq> ring_char_0 ..
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   516
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   517
lemma of_int_eq_id [simp]: "of_int = id"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   518
proof
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   519
  fix z show "of_int z = id z"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   520
    by (cases z) (simp add: of_int add minus int_def diff_minus)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   521
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   522
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   523
context ring_1
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   524
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   525
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   526
lemma of_nat_nat: "0 \<le> z \<Longrightarrow> of_nat (nat z) = of_int z"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   527
  by (cases z rule: eq_Abs_Integ)
23438
dd824e86fa8a remove simp attribute from of_nat_diff, for backward compatibility with zdiff_int
huffman
parents: 23431
diff changeset
   528
   (simp add: nat le of_int Zero_int_def of_nat_diff)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   529
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   530
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   531
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   532
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   533
subsection{*The Set of Integers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   534
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   535
context ring_1
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   536
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   537
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   538
definition
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   539
  Ints  :: "'a set"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   540
where
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   541
  "Ints = range of_int"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   542
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   543
end
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   544
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   545
notation (xsymbols)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   546
  Ints  ("\<int>")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   547
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   548
context ring_1
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   549
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   550
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   551
lemma Ints_0 [simp]: "0 \<in> \<int>"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   552
apply (simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   553
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   554
apply (rule of_int_0 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   555
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   556
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   557
lemma Ints_1 [simp]: "1 \<in> \<int>"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   558
apply (simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   559
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   560
apply (rule of_int_1 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   561
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   562
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   563
lemma Ints_add [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a + b \<in> \<int>"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   564
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   565
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   566
apply (rule of_int_add [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   567
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   568
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   569
lemma Ints_minus [simp]: "a \<in> \<int> \<Longrightarrow> -a \<in> \<int>"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   570
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   571
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   572
apply (rule of_int_minus [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   573
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   574
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   575
lemma Ints_mult [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a * b \<in> \<int>"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   576
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   577
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   578
apply (rule of_int_mult [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   579
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   580
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   581
lemma Ints_cases [cases set: Ints]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   582
  assumes "q \<in> \<int>"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   583
  obtains (of_int) z where "q = of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   584
  unfolding Ints_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   585
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   586
  from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   587
  then obtain z where "q = of_int z" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   588
  then show thesis ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   589
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   590
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   591
lemma Ints_induct [case_names of_int, induct set: Ints]:
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   592
  "q \<in> \<int> \<Longrightarrow> (\<And>z. P (of_int z)) \<Longrightarrow> P q"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   593
  by (rule Ints_cases) auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   594
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   595
end
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   596
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   597
lemma Ints_diff [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a-b \<in> \<int>"
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   598
apply (auto simp add: Ints_def)
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   599
apply (rule range_eqI)
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   600
apply (rule of_int_diff [symmetric])
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   601
done
e2e1a4b00de3 various localizations
haftmann
parents: 25164
diff changeset
   602
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   603
24728
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   604
subsection {* @{term setsum} and @{term setprod} *}
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   605
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   606
text {*By Jeremy Avigad*}
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   607
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   608
lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   609
  apply (cases "finite A")
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   610
  apply (erule finite_induct, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   611
  done
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   612
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   613
lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   614
  apply (cases "finite A")
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   615
  apply (erule finite_induct, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   616
  done
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   617
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   618
lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   619
  apply (cases "finite A")
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   620
  apply (erule finite_induct, auto simp add: of_nat_mult)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   621
  done
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   622
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   623
lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   624
  apply (cases "finite A")
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   625
  apply (erule finite_induct, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   626
  done
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   627
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   628
lemma setprod_nonzero_nat:
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   629
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   630
  by (rule setprod_nonzero, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   631
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   632
lemma setprod_zero_eq_nat:
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   633
    "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   634
  by (rule setprod_zero_eq, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   635
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   636
lemma setprod_nonzero_int:
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   637
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   638
  by (rule setprod_nonzero, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   639
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   640
lemma setprod_zero_eq_int:
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   641
    "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   642
  by (rule setprod_zero_eq, auto)
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   643
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   644
lemmas int_setsum = of_nat_setsum [where 'a=int]
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   645
lemmas int_setprod = of_nat_setprod [where 'a=int]
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   646
e2b3a1065676 moved Finite_Set before Datatype
haftmann
parents: 24506
diff changeset
   647
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   648
subsection {* Further properties *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   649
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   650
text{*Now we replace the case analysis rule by a more conventional one:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   651
whether an integer is negative or not.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   652
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   653
lemma zless_iff_Suc_zadd:
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   654
  "(w \<Colon> int) < z \<longleftrightarrow> (\<exists>n. z = w + of_nat (Suc n))"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   655
apply (cases z, cases w)
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   656
apply (auto simp add: less add int_def)
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   657
apply (rename_tac a b c d) 
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   658
apply (rule_tac x="a+d - Suc(c+b)" in exI) 
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   659
apply arith
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   660
done
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   661
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   662
lemma negD: "(x \<Colon> int) < 0 \<Longrightarrow> \<exists>n. x = - (of_nat (Suc n))"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   663
apply (cases x)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   664
apply (auto simp add: le minus Zero_int_def int_def order_less_le)
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   665
apply (rule_tac x="y - Suc x" in exI, arith)
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   666
done
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   667
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   668
theorem int_cases [cases type: int, case_names nonneg neg]:
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   669
  "[|!! n. (z \<Colon> int) = of_nat n ==> P;  !! n. z =  - (of_nat (Suc n)) ==> P |] ==> P"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   670
apply (cases "z < 0", blast dest!: negD)
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   671
apply (simp add: linorder_not_less del: of_nat_Suc)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   672
apply (blast dest: nat_0_le [THEN sym])
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   673
done
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   674
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   675
theorem int_induct [induct type: int, case_names nonneg neg]:
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   676
     "[|!! n. P (of_nat n \<Colon> int);  !!n. P (- (of_nat (Suc n))) |] ==> P z"
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   677
  by (cases z rule: int_cases) auto
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   678
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   679
text{*Contributed by Brian Huffman*}
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   680
theorem int_diff_cases [case_names diff]:
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   681
assumes prem: "!!m n. (z\<Colon>int) = of_nat m - of_nat n ==> P" shows "P"
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   682
apply (cases z rule: eq_Abs_Integ)
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   683
apply (rule_tac m=x and n=y in prem)
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   684
apply (simp add: int_def diff_def minus add)
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   685
done
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   686
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   687
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   688
subsection {* Legacy theorems *}
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   689
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   690
lemmas zminus_zminus = minus_minus [of "?z::int"]
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   691
lemmas zminus_0 = minus_zero [where 'a=int]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   692
lemmas zminus_zadd_distrib = minus_add_distrib [of "?z::int" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   693
lemmas zadd_commute = add_commute [of "?z::int" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   694
lemmas zadd_assoc = add_assoc [of "?z1.0::int" "?z2.0" "?z3.0"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   695
lemmas zadd_left_commute = add_left_commute [of "?x::int" "?y" "?z"]
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   696
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   697
lemmas zmult_ac = OrderedGroup.mult_ac
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   698
lemmas zadd_0 = OrderedGroup.add_0_left [of "?z::int"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   699
lemmas zadd_0_right = OrderedGroup.add_0_left [of "?z::int"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   700
lemmas zadd_zminus_inverse2 = left_minus [of "?z::int"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   701
lemmas zmult_zminus = mult_minus_left [of "?z::int" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   702
lemmas zmult_commute = mult_commute [of "?z::int" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   703
lemmas zmult_assoc = mult_assoc [of "?z1.0::int" "?z2.0" "?z3.0"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   704
lemmas zadd_zmult_distrib = left_distrib [of "?z1.0::int" "?z2.0" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   705
lemmas zadd_zmult_distrib2 = right_distrib [of "?w::int" "?z1.0" "?z2.0"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   706
lemmas zdiff_zmult_distrib = left_diff_distrib [of "?z1.0::int" "?z2.0" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   707
lemmas zdiff_zmult_distrib2 = right_diff_distrib [of "?w::int" "?z1.0" "?z2.0"]
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   708
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   709
lemmas int_distrib =
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   710
  zadd_zmult_distrib zadd_zmult_distrib2
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   711
  zdiff_zmult_distrib zdiff_zmult_distrib2
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   712
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   713
lemmas zmult_1 = mult_1_left [of "?z::int"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   714
lemmas zmult_1_right = mult_1_right [of "?z::int"]
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   715
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   716
lemmas zle_refl = order_refl [of "?w::int"]
23402
6472c689664f renamed vars in zle_trans for compatibility
nipkow
parents: 23372
diff changeset
   717
lemmas zle_trans = order_trans [where 'a=int and x="?i" and y="?j" and z="?k"]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   718
lemmas zle_anti_sym = order_antisym [of "?z::int" "?w"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   719
lemmas zle_linear = linorder_linear [of "?z::int" "?w"]
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   720
lemmas zless_linear = linorder_less_linear [where 'a = int]
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   721
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   722
lemmas zadd_left_mono = add_left_mono [of "?i::int" "?j" "?k"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   723
lemmas zadd_strict_right_mono = add_strict_right_mono [of "?i::int" "?j" "?k"]
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   724
lemmas zadd_zless_mono = add_less_le_mono [of "?w'::int" "?w" "?z'" "?z"]
23372
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   725
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   726
lemmas int_0_less_1 = zero_less_one [where 'a=int]
0035be079bee clean up instance proofs; reorganize section headings
huffman
parents: 23365
diff changeset
   727
lemmas int_0_neq_1 = zero_neq_one [where 'a=int]
23303
6091e530ff77 add abbreviation int_of_nat for of_nat::nat=>int;
huffman
parents: 23299
diff changeset
   728
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   729
lemmas inj_int = inj_of_nat [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   730
lemmas int_int_eq = of_nat_eq_iff [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   731
lemmas zadd_int = of_nat_add [where 'a=int, symmetric]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   732
lemmas int_mult = of_nat_mult [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   733
lemmas zmult_int = of_nat_mult [where 'a=int, symmetric]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   734
lemmas int_eq_0_conv = of_nat_eq_0_iff [where 'a=int and m="?n"]
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   735
lemmas zless_int = of_nat_less_iff [where 'a=int]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   736
lemmas int_less_0_conv = of_nat_less_0_iff [where 'a=int and m="?k"]
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   737
lemmas zero_less_int_conv = of_nat_0_less_iff [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   738
lemmas zle_int = of_nat_le_iff [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   739
lemmas zero_zle_int = of_nat_0_le_iff [where 'a=int]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   740
lemmas int_le_0_conv = of_nat_le_0_iff [where 'a=int and m="?n"]
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   741
lemmas int_0 = of_nat_0 [where 'a=int]
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   742
lemmas int_1 = of_nat_1 [where 'a=int]
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   743
lemmas int_Suc = of_nat_Suc [where 'a=int]
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23402
diff changeset
   744
lemmas abs_int_eq = abs_of_nat [where 'a=int and n="?m"]
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   745
lemmas of_int_int_eq = of_int_of_nat_eq [where 'a=int]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   746
lemmas zdiff_int = of_nat_diff [where 'a=int, symmetric]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   747
lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq]
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   748
lemmas int_eq_of_nat = TrueI
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   749
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23308
diff changeset
   750
abbreviation
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   751
  int :: "nat \<Rightarrow> int"
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   752
where
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   753
  "int \<equiv> of_nat"
f1dbfd7e3223 localized of_nat
haftmann
parents: 23950
diff changeset
   754
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   755
end