src/HOL/IntDef.thy
author huffman
Sat, 09 Jun 2007 02:38:51 +0200
changeset 23299 292b1cbd05dc
parent 23282 dfc459989d24
child 23303 6091e530ff77
permissions -rw-r--r--
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
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
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    26
  int :: "nat \<Rightarrow> int"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    27
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    28
  [code func del]: "int m = Abs_Integ (intrel `` {(m, 0)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    29
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    30
instance int :: zero
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
    31
  Zero_int_def: "0 \<equiv> Abs_Integ (intrel `` {(0, 0)})" ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    32
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    33
instance int :: one
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
    34
  One_int_def: "1 \<equiv> Abs_Integ (intrel `` {(1, 0)})" ..
23164
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 :: plus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
  add_int_def: "z + w \<equiv> Abs_Integ
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
    (\<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
    39
      intrel `` {(x + u, y + v)})" ..
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 :: minus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    42
  minus_int_def:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    43
    "- 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
    44
  diff_int_def:  "z - w \<equiv> z + (-w)" ..
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 :: times
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
  mult_int_def: "z * w \<equiv>  Abs_Integ
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    48
    (\<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
    49
      intrel `` {(x*u + y*v, x*v + y*u)})" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    50
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    51
instance int :: ord
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    52
  le_int_def:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    53
   "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
    54
  less_int_def: "z < w \<equiv> z \<le> w \<and> z \<noteq> w" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    55
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    56
lemmas [code func del] = Zero_int_def One_int_def add_int_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    57
  minus_int_def mult_int_def le_int_def less_int_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    58
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    59
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    60
subsection{*Construction of the Integers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    61
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    62
subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    63
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    64
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
    65
by (simp add: intrel_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    66
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    67
lemma equiv_intrel: "equiv UNIV intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    68
by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    69
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    70
text{*Reduces equality of equivalence classes to the @{term intrel} relation:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    71
  @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    72
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
    73
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    74
text{*All equivalence classes belong to set of representatives*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    75
lemma [simp]: "intrel``{(x,y)} \<in> Integ"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    76
by (auto simp add: Integ_def intrel_def quotient_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    77
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    78
text{*Reduces equality on abstractions to equality on representatives:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    79
  @{prop "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    80
declare Abs_Integ_inject [simp]  Abs_Integ_inverse [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    81
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    82
text{*Case analysis on the representation of an integer as an equivalence
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    83
      class of pairs of naturals.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    84
lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    85
     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
apply (rule Abs_Integ_cases [of z]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    87
apply (auto simp add: Integ_def quotient_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    88
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    89
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    90
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    91
subsubsection{*Integer Unary Negation*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    93
lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    94
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    95
  have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    96
    by (simp add: congruent_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    98
    by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    99
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   101
lemma zminus_zminus: "- (- z) = (z::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
  by (cases z) (simp add: minus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   104
lemma zminus_0: "- 0 = (0::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   105
  by (simp add: Zero_int_def minus)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   106
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   107
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
subsection{*Integer Addition*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
lemma add:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
     "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
      Abs_Integ (intrel``{(x+u, y+v)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
  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
   115
        respects2 intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   116
    by (simp add: congruent2_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   118
    by (simp add: add_int_def UN_UN_split_split_eq
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   119
                  UN_equiv_class2 [OF equiv_intrel equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   120
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   121
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   122
lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
  by (cases z, cases w) (simp add: minus add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   124
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   125
lemma zadd_commute: "(z::int) + w = w + z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   126
  by (cases z, cases w) (simp add: add_ac add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   127
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   128
lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   129
  by (cases z1, cases z2, cases z3) (simp add: add add_assoc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   130
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   131
(*For AC rewriting*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   132
lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   133
  apply (rule mk_left_commute [of "op +"])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   134
  apply (rule zadd_assoc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   135
  apply (rule zadd_commute)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   136
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   137
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   138
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   139
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   140
lemmas zmult_ac = OrderedGroup.mult_ac
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   141
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   142
(*also for the instance declaration int :: comm_monoid_add*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   143
lemma zadd_0: "(0::int) + z = z"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   144
apply (simp add: Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   145
apply (cases z, simp add: add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   146
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   147
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   148
lemma zadd_0_right: "z + (0::int) = z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   149
by (rule trans [OF zadd_commute zadd_0])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   150
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   151
lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   152
by (cases z, simp add: Zero_int_def minus add)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   153
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{*Integer Multiplication*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   156
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   157
text{*Congruence property for multiplication*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   158
lemma mult_congruent2:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   159
     "(%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
   160
      respects2 intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   161
apply (rule equiv_intrel [THEN congruent2_commuteI])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   162
 apply (force simp add: mult_ac, clarify) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   163
apply (simp add: congruent_def mult_ac)  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   164
apply (rename_tac u v w x y z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   165
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
   166
apply (simp add: mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   167
apply (simp add: add_mult_distrib [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   168
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   169
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   170
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   171
lemma mult:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   172
     "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   173
      Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   174
by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   175
              UN_equiv_class2 [OF equiv_intrel equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   176
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   177
lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   178
by (cases z, cases w, simp add: minus mult add_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   179
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   180
lemma zmult_commute: "(z::int) * w = w * z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   181
by (cases z, cases w, simp add: mult add_ac mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   182
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   183
lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   184
by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   185
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   186
lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   187
by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   188
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   189
lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   190
by (simp add: zmult_commute [of w] zadd_zmult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   191
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   193
by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   195
lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   196
by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   197
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   198
lemmas int_distrib =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   199
  zadd_zmult_distrib zadd_zmult_distrib2
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   200
  zdiff_zmult_distrib zdiff_zmult_distrib2
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   201
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   202
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   203
lemma zmult_1: "(1::int) * z = z"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   204
by (cases z, simp add: One_int_def mult)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   205
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   206
lemma zmult_1_right: "z * (1::int) = z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   207
by (rule trans [OF zmult_commute zmult_1])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   208
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   209
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   210
text{*The integers form a @{text comm_ring_1}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   211
instance int :: comm_ring_1
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   212
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   213
  fix i j k :: int
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   214
  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   215
  show "i + j = j + i" by (simp add: zadd_commute)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   216
  show "0 + i = i" by (rule zadd_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   217
  show "- i + i = 0" by (rule zadd_zminus_inverse2)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   218
  show "i - j = i + (-j)" by (simp add: diff_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   219
  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   220
  show "i * j = j * i" by (rule zmult_commute)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   221
  show "1 * i = i" by (rule zmult_1) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   222
  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   223
  show "0 \<noteq> (1::int)" by (simp add: Zero_int_def One_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   224
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   225
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   226
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   227
subsection{*The @{text "\<le>"} Ordering*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   228
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   229
lemma le:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   230
  "(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
   231
by (force simp add: le_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   232
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   233
lemma less:
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   234
  "(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
   235
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
   236
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   237
lemma zle_refl: "w \<le> (w::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   238
by (cases w, simp add: le)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   239
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   240
lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   241
by (cases i, cases j, cases k, simp add: le)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   242
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   243
lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   244
by (cases w, cases z, simp add: le)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   245
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   246
instance int :: order
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   247
  by intro_classes
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   248
    (assumption |
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   249
      rule zle_refl zle_trans zle_anti_sym less_int_def [THEN meta_eq_to_obj_eq])+
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 zle_linear: "(z::int) \<le> w \<or> w \<le> z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   252
by (cases z, cases w) (simp add: le linorder_linear)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   253
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   254
instance int :: linorder
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   255
  by intro_classes (rule zle_linear)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   256
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   257
lemmas zless_linear = linorder_less_linear [where 'a = int]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   258
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   259
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   260
lemma int_0_less_1: "0 < (1::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   261
by (simp add: Zero_int_def One_int_def less)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   262
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   263
lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   264
by (rule int_0_less_1 [THEN less_imp_neq])
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   265
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   266
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   267
subsection{*Monotonicity results*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   268
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   269
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
   270
proof
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   271
  fix a b c :: int
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   272
  show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   273
    by (cases a, cases b, cases c, simp add: le add)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   274
qed
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   275
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   276
lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   277
by (rule add_left_mono)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   278
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   279
lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   280
by (rule add_strict_right_mono)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   281
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   282
lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   283
by (rule add_less_le_mono)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   284
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   285
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   286
subsection{*Strict Monotonicity of Multiplication*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   287
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   288
text{*strict, in 1st argument; proof is by induction on k>0*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   289
lemma zmult_zless_mono2_lemma:
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   290
     "(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
   291
apply (induct "k", simp)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   292
apply (simp add: left_distrib)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   293
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
   294
apply (simp_all add: add_strict_mono)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   295
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   296
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   297
lemma int_of_nat_def: "of_nat m = Abs_Integ (intrel `` {(m, 0)})"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   298
by (induct m, simp_all add: Zero_int_def One_int_def add)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   300
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
   301
apply (cases k)
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   302
apply (auto simp add: le add int_of_nat_def Zero_int_def)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   303
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
   304
done
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   305
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   306
lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = of_nat n"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   307
apply (cases k)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   308
apply (simp add: less int_of_nat_def Zero_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   309
apply (rule_tac x="x-y" in exI, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   310
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   311
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   312
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
   313
apply (drule zero_less_imp_eq_int)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   314
apply (auto simp add: zmult_zless_mono2_lemma)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   315
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   316
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   317
instance int :: minus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   318
  zabs_def: "\<bar>i\<Colon>int\<bar> \<equiv> if i < 0 then - i else i" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   319
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   320
instance int :: distrib_lattice
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   321
  "inf \<equiv> min"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   322
  "sup \<equiv> max"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   323
  by intro_classes
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   324
    (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
   325
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   326
text{*The integers form an ordered integral domain*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   327
instance int :: ordered_idom
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   328
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   329
  fix i j k :: int
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   330
  show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   331
  show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   332
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   333
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   334
lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   335
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
   336
apply (simp add: less le add One_int_def)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   337
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   338
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   339
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   340
subsection{*@{term int}: Embedding the Naturals into the Integers*}
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   341
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   342
lemma inj_int: "inj int"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   343
by (simp add: inj_on_def int_def)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   344
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   345
lemma int_int_eq [iff]: "(int m = int n) = (m = n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   346
by (fast elim!: inj_int [THEN injD])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   347
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   348
lemma zadd_int: "(int m) + (int n) = int (m + n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   349
  by (simp add: int_def add)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   350
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   351
lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   352
  by (simp add: zadd_int zadd_assoc [symmetric])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   353
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   354
lemma int_mult: "int (m * n) = (int m) * (int n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   355
by (simp add: int_def mult)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   356
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   357
text{*Compatibility binding*}
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   358
lemmas zmult_int = int_mult [symmetric]
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   359
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   360
lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   361
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   362
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   363
lemma zless_int [simp]: "(int m < int n) = (m<n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   364
by (simp add: le add int_def linorder_not_le [symmetric]) 
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   365
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   366
lemma int_less_0_conv [simp]: "~ (int k < 0)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   367
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   368
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   369
lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   370
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   371
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   372
lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   373
by (simp add: linorder_not_less [symmetric])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   374
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   375
lemma zero_zle_int [simp]: "(0 \<le> int n)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   376
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   377
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   378
lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   379
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   380
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   381
lemma int_0 [simp]: "int 0 = (0::int)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   382
by (simp add: Zero_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   383
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   384
lemma int_1 [simp]: "int 1 = 1"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   385
by (simp add: One_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   386
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   387
lemma int_Suc0_eq_1: "int (Suc 0) = 1"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   388
by (simp add: One_int_def [folded int_def])
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   389
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   390
lemma int_Suc: "int (Suc m) = 1 + (int m)"
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   391
by (simp add: One_int_def [folded int_def] zadd_int)
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   392
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   393
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   394
subsection{*Magnitude of an Integer, as a Natural Number: @{term nat}*}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   395
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   396
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   397
  nat :: "int \<Rightarrow> nat"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   398
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   399
  [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
   400
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   401
lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   402
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   403
  have "(\<lambda>(x,y). {x-y}) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   404
    by (simp add: congruent_def) arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   405
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   406
    by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   407
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   408
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   409
lemma nat_int [simp]: "nat(int n) = n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   410
by (simp add: nat int_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   411
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   412
lemma nat_zero [simp]: "nat 0 = 0"
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   413
by (simp only: Zero_int_def [folded int_def] nat_int)
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   414
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   415
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   416
by (cases z, simp add: nat le int_def Zero_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   417
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   418
corollary nat_0_le: "0 \<le> z ==> int (nat z) = z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   419
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   420
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   421
lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   422
by (cases z, simp add: nat le int_def Zero_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   423
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   424
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
   425
apply (cases w, cases z) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   426
apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   427
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   428
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   429
text{*An alternative condition is @{term "0 \<le> w"} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   430
corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   431
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   432
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   433
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
   434
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   435
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   436
lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   437
apply (cases w, cases z) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   438
apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   439
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   440
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   441
lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   442
by (blast dest: nat_0_le sym)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   443
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   444
lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   445
by (cases w, simp add: nat le int_def Zero_int_def, arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   446
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   447
corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   448
by (simp only: eq_commute [of m] nat_eq_iff) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   449
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   450
lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   451
apply (cases w)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   452
apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   453
done
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 int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   456
by (auto simp add: nat_eq_iff2)
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 zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   459
by (insert zless_nat_conj [of 0], auto)
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 nat_add_distrib:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   462
     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   463
by (cases z, cases z', simp add: nat add le int_def Zero_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   464
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   465
lemma nat_diff_distrib:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   466
     "[| (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
   467
by (cases z, cases z', 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   468
    simp add: nat add minus diff_minus le int_def Zero_int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   469
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   470
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   471
lemma nat_zminus_int [simp]: "nat (- (int n)) = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   472
by (simp add: int_def minus nat Zero_int_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   473
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   474
lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   475
by (cases z, simp add: nat le int_def  linorder_not_le [symmetric], arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   476
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   477
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   478
subsection{*Lemmas about the Function @{term int} and Orderings*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   479
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   480
lemma negative_zless_0: "- (int (Suc n)) < 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   481
by (simp add: order_less_le)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   482
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   483
lemma negative_zless [iff]: "- (int (Suc n)) < int m"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   484
by (rule negative_zless_0 [THEN order_less_le_trans], simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   485
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   486
lemma negative_zle_0: "- int n \<le> 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   487
by (simp add: minus_le_iff)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   488
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   489
lemma negative_zle [iff]: "- int n \<le> int m"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   490
by (rule order_trans [OF negative_zle_0 zero_zle_int])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   491
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   492
lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   493
by (subst le_minus_iff, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   494
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   495
lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   496
by (simp add: int_def le minus Zero_int_def) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   497
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   498
lemma not_int_zless_negative [simp]: "~ (int n < - int m)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   499
by (simp add: linorder_not_less)
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 negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   502
by (force simp add: order_eq_iff [of "- int n"] int_zle_neg)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   503
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   504
lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   505
proof (cases w, cases z, simp add: le add int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   506
  fix a b c d
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   507
  assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   508
  show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   509
  proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   510
    assume "a + d \<le> c + b" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   511
    thus "\<exists>n. c + b = a + n + d" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   512
      by (auto intro!: exI [where x="c+b - (a+d)"])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   513
  next    
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   514
    assume "\<exists>n. c + b = a + n + d" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   515
    then obtain n where "c + b = a + n + d" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   516
    thus "a + d \<le> c + b" by arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   517
  qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   518
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   519
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   520
lemma abs_int_eq [simp]: "abs (int m) = int m"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   521
by (simp add: abs_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   522
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   523
text{*This version is proved for all ordered rings, not just integers!
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   524
      It is proved here because attribute @{text arith_split} is not available
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   525
      in theory @{text Ring_and_Field}.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   526
      But is it really better than just rewriting with @{text abs_if}?*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   527
lemma abs_split [arith_split]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   528
     "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
   529
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
   530
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   531
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   532
subsection {* Constants @{term neg} and @{term iszero} *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   533
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   534
definition
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   535
  neg  :: "'a\<Colon>ordered_idom \<Rightarrow> bool"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   536
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   537
  [code inline]: "neg Z \<longleftrightarrow> Z < 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   538
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   539
definition (*for simplifying equalities*)
23276
a70934b63910 generalize of_nat and related constants to class semiring_1
huffman
parents: 23164
diff changeset
   540
  iszero :: "'a\<Colon>semiring_1 \<Rightarrow> bool"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   541
where
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   542
  "iszero z \<longleftrightarrow> z = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   543
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   544
lemma not_neg_int [simp]: "~ neg(int n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   545
by (simp add: neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   546
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   547
lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   548
by (simp add: neg_def neg_less_0_iff_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   549
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   550
lemmas neg_eq_less_0 = neg_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   551
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   552
lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   553
by (simp add: neg_def linorder_not_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   554
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   555
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   556
subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   557
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   558
lemma not_neg_0: "~ neg 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   559
by (simp add: One_int_def neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   560
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   561
lemma not_neg_1: "~ neg 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   562
by (simp add: neg_def linorder_not_less zero_le_one)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   563
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   564
lemma iszero_0: "iszero 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   565
by (simp add: iszero_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   566
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   567
lemma not_iszero_1: "~ iszero 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   568
by (simp add: iszero_def eq_commute)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   569
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   570
lemma neg_nat: "neg z ==> nat z = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   571
by (simp add: neg_def order_less_imp_le) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   572
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   573
lemma not_neg_nat: "~ neg z ==> int (nat z) = z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   574
by (simp add: linorder_not_less neg_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   575
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   576
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   577
subsection{*The Set of Natural Numbers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   578
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   579
constdefs
23276
a70934b63910 generalize of_nat and related constants to class semiring_1
huffman
parents: 23164
diff changeset
   580
  Nats  :: "'a::semiring_1 set"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   581
  "Nats == range of_nat"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   582
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   583
notation (xsymbols)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   584
  Nats  ("\<nat>")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   585
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   586
lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   587
by (simp add: Nats_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   588
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   589
lemma Nats_0 [simp]: "0 \<in> Nats"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   590
apply (simp add: Nats_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   591
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   592
apply (rule of_nat_0 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   593
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   594
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   595
lemma Nats_1 [simp]: "1 \<in> Nats"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   596
apply (simp add: Nats_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   597
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   598
apply (rule of_nat_1 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   599
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   600
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   601
lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   602
apply (auto simp add: Nats_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   603
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   604
apply (rule of_nat_add [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   605
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   606
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   607
lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   608
apply (auto simp add: Nats_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   609
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   610
apply (rule of_nat_mult [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   611
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   612
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   613
text{*Agreement with the specific embedding for the integers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   614
lemma int_eq_of_nat: "int = (of_nat :: nat => int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   615
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   616
  fix n
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   617
  show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   618
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   619
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   620
lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   621
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   622
  fix n
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   623
  show "of_nat n = id n"  by (induct n, simp_all)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   624
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   625
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   626
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   627
subsection{*Embedding of the Integers into any @{text ring_1}:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   628
@{term of_int}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   629
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   630
constdefs
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   631
   of_int :: "int => 'a::ring_1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   632
   "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   633
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   634
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   635
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
   636
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   637
  have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   638
    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   639
            del: of_nat_add) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   640
  thus ?thesis
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   641
    by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   642
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   643
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   644
lemma of_int_0 [simp]: "of_int 0 = 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   645
by (simp add: of_int Zero_int_def int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   646
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   647
lemma of_int_1 [simp]: "of_int 1 = 1"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   648
by (simp add: of_int One_int_def int_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   649
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   650
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
   651
by (cases w, cases z, simp add: compare_rls of_int add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   652
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   653
lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   654
by (cases z, simp add: compare_rls of_int minus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   655
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   656
lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   657
by (simp add: diff_minus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   658
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   659
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
   660
apply (cases w, cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   661
apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   662
                 mult add_ac)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   663
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   664
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   665
lemma of_int_le_iff [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   666
     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   667
apply (cases w)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   668
apply (cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   669
apply (simp add: compare_rls of_int le diff_int_def add minus
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   670
                 of_nat_add [symmetric]   del: of_nat_add)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   671
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   672
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   673
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   674
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
   675
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
   676
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   677
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   678
lemma of_int_less_iff [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   679
     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   680
by (simp add: linorder_not_le [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   681
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   682
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   683
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
   684
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
   685
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   686
text{*Class for unital rings with characteristic zero.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   687
 Includes non-ordered rings like the complex numbers.*}
23282
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   688
axclass ring_char_0 < ring_1, semiring_char_0
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   689
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   690
lemma of_int_eq_iff [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   691
     "(of_int w = (of_int z::'a::ring_char_0)) = (w = z)"
23282
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   692
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
   693
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
   694
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
   695
done
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   696
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   697
text{*Every @{text ordered_idom} has characteristic zero.*}
23282
dfc459989d24 add axclass semiring_char_0 for types where of_nat is injective
huffman
parents: 23276
diff changeset
   698
instance ordered_idom < ring_char_0 ..
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   699
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   700
text{*Special cases where either operand is zero*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   701
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
   702
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
   703
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   704
lemma of_int_eq_id [simp]: "of_int = (id :: int => int)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   705
proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   706
  fix z
23299
292b1cbd05dc remove dependencies of proofs on constant int::nat=>int, preparing to remove it
huffman
parents: 23282
diff changeset
   707
  show "of_int z = id z"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   708
    by (cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   709
      (simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   710
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   711
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   712
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   713
subsection{*The Set of Integers*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   714
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   715
constdefs
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   716
  Ints  :: "'a::ring_1 set"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   717
  "Ints == range of_int"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   718
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   719
notation (xsymbols)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   720
  Ints  ("\<int>")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   721
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   722
lemma Ints_0 [simp]: "0 \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   723
apply (simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   724
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   725
apply (rule of_int_0 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   726
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   727
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   728
lemma Ints_1 [simp]: "1 \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   729
apply (simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   730
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   731
apply (rule of_int_1 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   732
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   733
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   734
lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   735
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   736
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   737
apply (rule of_int_add [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   738
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   739
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   740
lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   741
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   742
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   743
apply (rule of_int_minus [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   744
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   745
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   746
lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   747
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   748
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   749
apply (rule of_int_diff [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   750
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   751
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   752
lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   753
apply (auto simp add: Ints_def)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   754
apply (rule range_eqI)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   755
apply (rule of_int_mult [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   756
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   757
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   758
text{*Collapse nested embeddings*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   759
lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   760
by (induct n, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   761
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   762
lemma of_int_int_eq [simp]: "of_int (int n) = of_nat n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   763
by (simp add: int_eq_of_nat)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   764
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   765
lemma Ints_cases [cases set: Ints]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   766
  assumes "q \<in> \<int>"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   767
  obtains (of_int) z where "q = of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   768
  unfolding Ints_def
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   769
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   770
  from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   771
  then obtain z where "q = of_int z" ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   772
  then show thesis ..
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   773
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   774
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   775
lemma Ints_induct [case_names of_int, induct set: Ints]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   776
  "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   777
  by (rule Ints_cases) auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   778
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   779
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   780
(* int (Suc n) = 1 + int n *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   781
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   782
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   783
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   784
subsection{*More Properties of @{term setsum} and  @{term setprod}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   785
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   786
text{*By Jeremy Avigad*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   787
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   788
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   789
lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   790
  apply (cases "finite A")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   791
  apply (erule finite_induct, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   792
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   793
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   794
lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   795
  apply (cases "finite A")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   796
  apply (erule finite_induct, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   797
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   798
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   799
lemma int_setsum: "int (setsum f A) = (\<Sum>x\<in>A. int(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   800
  by (simp add: int_eq_of_nat of_nat_setsum)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   801
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   802
lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   803
  apply (cases "finite A")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   804
  apply (erule finite_induct, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   805
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   806
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   807
lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   808
  apply (cases "finite A")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   809
  apply (erule finite_induct, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   810
  done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   811
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   812
lemma int_setprod: "int (setprod f A) = (\<Prod>x\<in>A. int(f x))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   813
  by (simp add: int_eq_of_nat of_nat_setprod)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   814
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   815
lemma setprod_nonzero_nat:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   816
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   817
  by (rule setprod_nonzero, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   818
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   819
lemma setprod_zero_eq_nat:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   820
    "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   821
  by (rule setprod_zero_eq, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   822
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   823
lemma setprod_nonzero_int:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   824
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   825
  by (rule setprod_nonzero, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   826
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   827
lemma setprod_zero_eq_int:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   828
    "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   829
  by (rule setprod_zero_eq, auto)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   830
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   831
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   832
subsection {* Further properties *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   833
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   834
text{*Now we replace the case analysis rule by a more conventional one:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   835
whether an integer is negative or not.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   836
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   837
lemma zless_iff_Suc_zadd:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   838
    "(w < z) = (\<exists>n. z = w + int(Suc n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   839
apply (cases z, cases w)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   840
apply (auto simp add: le add int_def linorder_not_le [symmetric]) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   841
apply (rename_tac a b c d) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   842
apply (rule_tac x="a+d - Suc(c+b)" in exI) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   843
apply arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   844
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   845
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   846
lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   847
apply (cases x)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   848
apply (auto simp add: le minus Zero_int_def int_def order_less_le) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   849
apply (rule_tac x="y - Suc x" in exI, arith)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   850
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   851
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   852
theorem int_cases [cases type: int, case_names nonneg neg]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   853
     "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   854
apply (cases "z < 0", blast dest!: negD)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   855
apply (simp add: linorder_not_less)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   856
apply (blast dest: nat_0_le [THEN sym])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   857
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   858
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   859
theorem int_induct [induct type: int, case_names nonneg neg]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   860
     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   861
  by (cases z) auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   862
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   863
text{*Contributed by Brian Huffman*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   864
theorem int_diff_cases [case_names diff]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   865
assumes prem: "!!m n. z = int m - int n ==> P" shows "P"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   866
 apply (rule_tac z=z in int_cases)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   867
  apply (rule_tac m=n and n=0 in prem, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   868
 apply (rule_tac m=0 and n="Suc n" in prem, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   869
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   870
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   871
lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   872
apply (cases z)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   873
apply (simp_all add: not_zle_0_negative del: int_Suc)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   874
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   875
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   876
lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   877
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   878
lemmas [simp] = int_Suc
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   879
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   880
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   881
subsection {* Legacy ML bindings *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   882
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   883
ML {*
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   884
val of_nat_0 = @{thm of_nat_0};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   885
val of_nat_1 = @{thm of_nat_1};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   886
val of_nat_Suc = @{thm of_nat_Suc};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   887
val of_nat_add = @{thm of_nat_add};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   888
val of_nat_mult = @{thm of_nat_mult};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   889
val of_int_0 = @{thm of_int_0};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   890
val of_int_1 = @{thm of_int_1};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   891
val of_int_add = @{thm of_int_add};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   892
val of_int_mult = @{thm of_int_mult};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   893
val int_eq_of_nat = @{thm int_eq_of_nat};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   894
val zle_int = @{thm zle_int};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   895
val int_int_eq = @{thm int_int_eq};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   896
val diff_int_def = @{thm diff_int_def};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   897
val zadd_ac = @{thms zadd_ac};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   898
val zless_int = @{thm zless_int};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   899
val zadd_int = @{thm zadd_int};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   900
val zmult_int = @{thm zmult_int};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   901
val nat_0_le = @{thm nat_0_le};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   902
val int_0 = @{thm int_0};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   903
val int_1 = @{thm int_1};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   904
val abs_split = @{thm abs_split};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   905
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   906
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   907
end