src/HOL/Integ/IntDef.thy
author haftmann
Fri, 02 Mar 2007 15:43:23 +0100
changeset 22392 35f54980d4cc
parent 21911 e29bcab0c81c
child 22456 6070e48ecb78
permissions -rw-r--r--
tuned code theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
     1
(*  Title:      IntDef.thy
691c70898518 new files in Integ
paulson
parents:
diff changeset
     2
    ID:         $Id$
691c70898518 new files in Integ
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
691c70898518 new files in Integ
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
691c70898518 new files in Integ
paulson
parents:
diff changeset
     5
691c70898518 new files in Integ
paulson
parents:
diff changeset
     6
*)
691c70898518 new files in Integ
paulson
parents:
diff changeset
     7
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
     8
header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} 
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15047
diff changeset
    10
theory IntDef
21243
afffe1f72143 removed theory NatArith (now part of Nat);
wenzelm
parents: 21238
diff changeset
    11
imports Equiv_Relations Nat
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15047
diff changeset
    12
begin
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    13
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    14
constdefs
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    15
  intrel :: "((nat * nat) * (nat * nat)) set"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    16
    --{*the equivalence relation underlying the integers*}
14496
paulson
parents: 14485
diff changeset
    17
    "intrel == {((x,y),(u,v)) | x y u v. x+v = u+y}"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    18
691c70898518 new files in Integ
paulson
parents:
diff changeset
    19
typedef (Integ)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    20
  int = "UNIV//intrel"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    21
    by (auto simp add: quotient_def)
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    22
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
    23
instance int :: "{ord, zero, one, plus, times, minus}" ..
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    24
691c70898518 new files in Integ
paulson
parents:
diff changeset
    25
constdefs
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    26
  int :: "nat => int"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    27
  "int m == Abs_Integ(intrel `` {(m,0)})"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    28
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    29
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14259
diff changeset
    30
defs (overloaded)
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    31
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    32
  Zero_int_def:  "0 == int 0"
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    33
  One_int_def:   "1 == int 1"
8937
7328d7c8d201 defining 0::int to be (int 0)
paulson
parents: 7375
diff changeset
    34
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    35
  minus_int_def:
14532
paulson
parents: 14496
diff changeset
    36
    "- z == Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. intrel``{(y,x)})"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    37
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    38
  add_int_def:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    39
   "z + w ==
14532
paulson
parents: 14496
diff changeset
    40
       Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w.
paulson
parents: 14496
diff changeset
    41
		 intrel``{(x+u, y+v)})"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    42
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    43
  diff_int_def:  "z - (w::int) == z + (-w)"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    44
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    45
  mult_int_def:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    46
   "z * w ==
14532
paulson
parents: 14496
diff changeset
    47
       Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w.
paulson
parents: 14496
diff changeset
    48
		  intrel``{(x*u + y*v, x*v + y*u)})"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    49
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    50
  le_int_def:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    51
   "z \<le> (w::int) == 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    52
    \<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Integ z & (u,v) \<in> Rep_Integ w"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    53
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    54
  less_int_def: "(z < (w::int)) == (z \<le> w & z \<noteq> w)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    55
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    56
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    57
subsection{*Construction of the Integers*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    58
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    59
subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    60
14496
paulson
parents: 14485
diff changeset
    61
lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    62
by (simp add: intrel_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    63
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    64
lemma equiv_intrel: "equiv UNIV intrel"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    65
by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    66
14496
paulson
parents: 14485
diff changeset
    67
text{*Reduces equality of equivalence classes to the @{term intrel} relation:
paulson
parents: 14485
diff changeset
    68
  @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
paulson
parents: 14485
diff changeset
    69
lemmas equiv_intrel_iff = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I]
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    70
14496
paulson
parents: 14485
diff changeset
    71
declare equiv_intrel_iff [simp]
paulson
parents: 14485
diff changeset
    72
paulson
parents: 14485
diff changeset
    73
paulson
parents: 14485
diff changeset
    74
text{*All equivalence classes belong to set of representatives*}
14532
paulson
parents: 14496
diff changeset
    75
lemma [simp]: "intrel``{(x,y)} \<in> Integ"
14496
paulson
parents: 14485
diff changeset
    76
by (auto simp add: Integ_def intrel_def quotient_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    77
15413
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
    78
text{*Reduces equality on abstractions to equality on representatives:
14496
paulson
parents: 14485
diff changeset
    79
  @{term "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
15413
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
    80
declare Abs_Integ_inject [simp]  Abs_Integ_inverse [simp]
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    81
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    82
text{*Case analysis on the representation of an integer as an equivalence
14485
ea2707645af8 new material from Avigad
paulson
parents: 14479
diff changeset
    83
      class of pairs of naturals.*}
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    84
lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    85
     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
15413
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
    86
apply (rule Abs_Integ_cases [of z]) 
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
    87
apply (auto simp add: Integ_def quotient_def) 
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    88
done
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    89
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    90
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    91
subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    92
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    93
lemma inj_int: "inj int"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
    94
by (simp add: inj_on_def int_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    95
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    96
lemma int_int_eq [iff]: "(int m = int n) = (m = n)"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    97
by (fast elim!: inj_int [THEN injD])
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    98
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    99
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   100
subsubsection{*Integer Unary Negation*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   101
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   102
lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   103
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   104
  have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   105
    by (simp add: congruent_def) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   106
  thus ?thesis
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   107
    by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   108
qed
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   109
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   110
lemma zminus_zminus: "- (- z) = (z::int)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   111
  by (cases z) (simp add: minus)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   112
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   113
lemma zminus_0: "- 0 = (0::int)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   114
  by (simp add: int_def Zero_int_def minus)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   115
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   116
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   117
subsection{*Integer Addition*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   118
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   119
lemma add:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   120
     "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   121
      Abs_Integ (intrel``{(x+u, y+v)})"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   122
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   123
  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   124
        respects2 intrel"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   125
    by (simp add: congruent2_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   126
  thus ?thesis
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   127
    by (simp add: add_int_def UN_UN_split_split_eq
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14532
diff changeset
   128
                  UN_equiv_class2 [OF equiv_intrel equiv_intrel])
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   129
qed
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   130
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   131
lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   132
  by (cases z, cases w) (simp add: minus add)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   133
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   134
lemma zadd_commute: "(z::int) + w = w + z"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   135
  by (cases z, cases w) (simp add: add_ac add)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   136
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   137
lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   138
  by (cases z1, cases z2, cases z3) (simp add: add add_assoc)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   139
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   140
(*For AC rewriting*)
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   141
lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   142
  apply (rule mk_left_commute [of "op +"])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   143
  apply (rule zadd_assoc)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   144
  apply (rule zadd_commute)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   145
  done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   146
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   147
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   148
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   149
lemmas zmult_ac = OrderedGroup.mult_ac
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   150
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   151
lemma zadd_int: "(int m) + (int n) = int (m + n)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   152
  by (simp add: int_def add)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   153
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   154
lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   155
  by (simp add: zadd_int zadd_assoc [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   156
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   157
lemma int_Suc: "int (Suc m) = 1 + (int m)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   158
  by (simp add: One_int_def zadd_int)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   159
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   160
(*also for the instance declaration int :: comm_monoid_add*)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   161
lemma zadd_0: "(0::int) + z = z"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   162
apply (simp add: Zero_int_def int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   163
apply (cases z, simp add: add)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   164
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   165
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   166
lemma zadd_0_right: "z + (0::int) = z"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   167
by (rule trans [OF zadd_commute zadd_0])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   168
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   169
lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   170
by (cases z, simp add: int_def Zero_int_def minus add)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   171
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   172
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   173
subsection{*Integer Multiplication*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   174
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   175
text{*Congruence property for multiplication*}
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   176
lemma mult_congruent2:
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   177
     "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   178
      respects2 intrel"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   179
apply (rule equiv_intrel [THEN congruent2_commuteI])
14532
paulson
parents: 14496
diff changeset
   180
 apply (force simp add: mult_ac, clarify) 
paulson
parents: 14496
diff changeset
   181
apply (simp add: congruent_def mult_ac)  
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   182
apply (rename_tac u v w x y z)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   183
apply (subgoal_tac "u*y + x*y = w*y + v*y  &  u*z + x*z = w*z + v*z")
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16642
diff changeset
   184
apply (simp add: mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   185
apply (simp add: add_mult_distrib [symmetric])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   186
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   187
14532
paulson
parents: 14496
diff changeset
   188
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   189
lemma mult:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   190
     "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   191
      Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   192
by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
14658
b1293d0f8d5f congruent2 now allows different equiv relations
paulson
parents: 14532
diff changeset
   193
              UN_equiv_class2 [OF equiv_intrel equiv_intrel])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   194
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   195
lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   196
by (cases z, cases w, simp add: minus mult add_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   197
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   198
lemma zmult_commute: "(z::int) * w = w * z"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   199
by (cases z, cases w, simp add: mult add_ac mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   200
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   201
lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   202
by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   203
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   204
lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   205
by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   206
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   207
lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   208
by (simp add: zmult_commute [of w] zadd_zmult_distrib)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   209
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   210
lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
14496
paulson
parents: 14485
diff changeset
   211
by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   212
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   213
lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   214
by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   215
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   216
lemmas int_distrib =
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   217
  zadd_zmult_distrib zadd_zmult_distrib2
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   218
  zdiff_zmult_distrib zdiff_zmult_distrib2
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   219
16413
47ffc49c7d7b a few new integer lemmas
paulson
parents: 15921
diff changeset
   220
lemma int_mult: "int (m * n) = (int m) * (int n)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   221
by (simp add: int_def mult)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   222
16413
47ffc49c7d7b a few new integer lemmas
paulson
parents: 15921
diff changeset
   223
text{*Compatibility binding*}
47ffc49c7d7b a few new integer lemmas
paulson
parents: 15921
diff changeset
   224
lemmas zmult_int = int_mult [symmetric]
47ffc49c7d7b a few new integer lemmas
paulson
parents: 15921
diff changeset
   225
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   226
lemma zmult_1: "(1::int) * z = z"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   227
by (cases z, simp add: One_int_def int_def mult)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   228
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   229
lemma zmult_1_right: "z * (1::int) = z"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   230
by (rule trans [OF zmult_commute zmult_1])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   231
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   232
14740
c8e1937110c2 fixed latex problems
nipkow
parents: 14738
diff changeset
   233
text{*The integers form a @{text comm_ring_1}*}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   234
instance int :: comm_ring_1
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   235
proof
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   236
  fix i j k :: int
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   237
  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   238
  show "i + j = j + i" by (simp add: zadd_commute)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   239
  show "0 + i = i" by (rule zadd_0)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   240
  show "- i + i = 0" by (rule zadd_zminus_inverse2)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   241
  show "i - j = i + (-j)" by (simp add: diff_int_def)
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   242
  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   243
  show "i * j = j * i" by (rule zmult_commute)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   244
  show "1 * i = i" by (rule zmult_1) 
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   245
  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   246
  show "0 \<noteq> (1::int)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   247
    by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   248
qed
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   249
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   250
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   251
subsection{*The @{text "\<le>"} Ordering*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   252
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   253
lemma le:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   254
  "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   255
by (force simp add: le_int_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   256
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   257
lemma zle_refl: "w \<le> (w::int)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   258
by (cases w, simp add: le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   259
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   260
lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   261
by (cases i, cases j, cases k, simp add: le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   262
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   263
lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   264
by (cases w, cases z, simp add: le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   265
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   266
instance int :: order
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
   267
  by intro_classes
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
   268
    (assumption |
22392
35f54980d4cc tuned code theorems
haftmann
parents: 21911
diff changeset
   269
      rule zle_refl zle_trans zle_anti_sym less_int_def [THEN meta_eq_to_obj_eq])+
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   270
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   271
(* Axiom 'linorder_linear' of class 'linorder': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   272
lemma zle_linear: "(z::int) \<le> w | w \<le> z"
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
   273
by (cases z, cases w) (simp add: le linorder_linear)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   274
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   275
instance int :: linorder
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14658
diff changeset
   276
  by intro_classes (rule zle_linear)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   277
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   278
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   279
lemmas zless_linear = linorder_less_linear [where 'a = int]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   280
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   281
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   282
lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   283
by (simp add: Zero_int_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   284
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   285
lemma zless_int [simp]: "(int m < int n) = (m<n)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   286
by (simp add: le add int_def linorder_not_le [symmetric]) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   287
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   288
lemma int_less_0_conv [simp]: "~ (int k < 0)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   289
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   290
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   291
lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   292
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   293
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   294
lemma int_0_less_1: "0 < (1::int)"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   295
by (simp only: Zero_int_def One_int_def One_nat_def zless_int)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   296
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   297
lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)"
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   298
by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   299
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   300
lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   301
by (simp add: linorder_not_less [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   302
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   303
lemma zero_zle_int [simp]: "(0 \<le> int n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   304
by (simp add: Zero_int_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   305
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   306
lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   307
by (simp add: Zero_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   308
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   309
lemma int_0 [simp]: "int 0 = (0::int)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   310
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   311
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   312
lemma int_1 [simp]: "int 1 = 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   313
by (simp add: One_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   314
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   315
lemma int_Suc0_eq_1: "int (Suc 0) = 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   316
by (simp add: One_int_def One_nat_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   317
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   318
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   319
subsection{*Monotonicity results*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   320
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   321
lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   322
by (cases i, cases j, cases k, simp add: le add)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   323
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   324
lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   325
apply (cases i, cases j, cases k)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   326
apply (simp add: linorder_not_le [where 'a = int, symmetric]
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   327
                 linorder_not_le [where 'a = nat]  le add)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   328
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   329
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   330
lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   331
by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono])
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   332
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   333
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   334
subsection{*Strict Monotonicity of Multiplication*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   335
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   336
text{*strict, in 1st argument; proof is by induction on k>0*}
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15169
diff changeset
   337
lemma zmult_zless_mono2_lemma:
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15169
diff changeset
   338
     "i<j ==> 0<k ==> int k * i < int k * j"
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15169
diff changeset
   339
apply (induct "k", simp)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   340
apply (simp add: int_Suc)
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15169
diff changeset
   341
apply (case_tac "k=0")
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   342
apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   343
apply (simp add: zadd_zless_mono int_Suc0_eq_1 order_le_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   344
done
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   345
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   346
lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   347
apply (cases k)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   348
apply (auto simp add: le add int_def Zero_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   349
apply (rule_tac x="x-y" in exI, simp)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   350
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   351
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   352
lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   353
apply (frule order_less_imp_le [THEN zero_le_imp_eq_int])
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   354
apply (auto simp add: zmult_zless_mono2_lemma)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   355
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   356
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   357
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   358
defs (overloaded)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   359
    zabs_def:  "abs(i::int) == if i < 0 then -i else i"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   360
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   361
14740
c8e1937110c2 fixed latex problems
nipkow
parents: 14738
diff changeset
   362
text{*The integers form an ordered @{text comm_ring_1}*}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   363
instance int :: ordered_idom
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   364
proof
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   365
  fix i j k :: int
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   366
  show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   367
  show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   368
  show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   369
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   370
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   371
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   372
lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   373
apply (cases w, cases z) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   374
apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   375
done
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   376
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   377
subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   379
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   380
   nat  :: "int => nat"
14532
paulson
parents: 14496
diff changeset
   381
    "nat z == contents (\<Union>(x,y) \<in> Rep_Integ z. { x-y })"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   382
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   383
lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   384
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   385
  have "(\<lambda>(x,y). {x-y}) respects intrel"
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20355
diff changeset
   386
    by (simp add: congruent_def) arith
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   387
  thus ?thesis
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   388
    by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   389
qed
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   390
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   391
lemma nat_int [simp]: "nat(int n) = n"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   392
by (simp add: nat int_def) 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   393
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   394
lemma nat_zero [simp]: "nat 0 = 0"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   395
by (simp only: Zero_int_def nat_int)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   396
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   397
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   398
by (cases z, simp add: nat le int_def Zero_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   399
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   400
corollary nat_0_le: "0 \<le> z ==> int (nat z) = z"
15413
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   401
by simp
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   402
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   403
lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   404
by (cases z, simp add: nat le int_def Zero_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   405
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   406
lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   407
apply (cases w, cases z) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   408
apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   409
done
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   410
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   411
text{*An alternative condition is @{term "0 \<le> w"} *}
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   412
corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   413
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   414
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   415
corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   416
by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   417
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   418
lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   419
apply (cases w, cases z) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   420
apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   421
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   422
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   423
lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   424
by (blast dest: nat_0_le sym)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   425
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   426
lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   427
by (cases w, simp add: nat le int_def Zero_int_def, arith)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   428
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   429
corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   430
by (simp only: eq_commute [of m] nat_eq_iff) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   431
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   432
lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   433
apply (cases w)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   434
apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   435
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   436
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   437
lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   438
by (auto simp add: nat_eq_iff2)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   439
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   440
lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   441
by (insert zless_nat_conj [of 0], auto)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   442
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   443
lemma nat_add_distrib:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   444
     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   445
by (cases z, cases z', simp add: nat add le int_def Zero_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   446
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   447
lemma nat_diff_distrib:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   448
     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   449
by (cases z, cases z', 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   450
    simp add: nat add minus diff_minus le int_def Zero_int_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   451
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   452
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   453
lemma nat_zminus_int [simp]: "nat (- (int n)) = 0"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   454
by (simp add: int_def minus nat Zero_int_def) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   455
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   456
lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   457
by (cases z, simp add: nat le int_def  linorder_not_le [symmetric], arith)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   458
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   459
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   460
subsection{*Lemmas about the Function @{term int} and Orderings*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   461
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   462
lemma negative_zless_0: "- (int (Suc n)) < 0"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   463
by (simp add: order_less_le)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   464
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   465
lemma negative_zless [iff]: "- (int (Suc n)) < int m"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   466
by (rule negative_zless_0 [THEN order_less_le_trans], simp)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   467
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   468
lemma negative_zle_0: "- int n \<le> 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   469
by (simp add: minus_le_iff)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   470
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   471
lemma negative_zle [iff]: "- int n \<le> int m"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   472
by (rule order_trans [OF negative_zle_0 zero_zle_int])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   473
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   474
lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   475
by (subst le_minus_iff, simp)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   476
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   477
lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   478
by (simp add: int_def le minus Zero_int_def) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   479
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   480
lemma not_int_zless_negative [simp]: "~ (int n < - int m)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   481
by (simp add: linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   482
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   483
lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   484
by (force simp add: order_eq_iff [of "- int n"] int_zle_neg)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   485
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   486
lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
15413
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   487
proof (cases w, cases z, simp add: le add int_def)
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   488
  fix a b c d
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   489
  assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})"
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   490
  show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)"
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   491
  proof
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   492
    assume "a + d \<le> c + b" 
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   493
    thus "\<exists>n. c + b = a + n + d" 
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   494
      by (auto intro!: exI [where x="c+b - (a+d)"])
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   495
  next    
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   496
    assume "\<exists>n. c + b = a + n + d" 
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   497
    then obtain n where "c + b = a + n + d" ..
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   498
    thus "a + d \<le> c + b" by arith
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   499
  qed
901d1bfedf09 removal of archaic Abs/Rep proofs
paulson
parents: 15409
diff changeset
   500
qed
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   501
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   502
lemma abs_int_eq [simp]: "abs (int m) = int m"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14740
diff changeset
   503
by (simp add: abs_if)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   504
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   505
text{*This version is proved for all ordered rings, not just integers!
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   506
      It is proved here because attribute @{text arith_split} is not available
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   507
      in theory @{text Ring_and_Field}.
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   508
      But is it really better than just rewriting with @{text abs_if}?*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   509
lemma abs_split [arith_split]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   510
     "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   511
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   512
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   513
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   514
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   515
subsection{*The Constants @{term neg} and @{term iszero}*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   516
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   517
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   518
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   519
  neg   :: "'a::ordered_idom => bool"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   520
  "neg(Z) == Z < 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   521
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   522
  (*For simplifying equalities*)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   523
  iszero :: "'a::comm_semiring_1_cancel => bool"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   524
  "iszero z == z = (0)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   525
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   526
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   527
lemma not_neg_int [simp]: "~ neg(int n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   528
by (simp add: neg_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   529
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   530
lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   531
by (simp add: neg_def neg_less_0_iff_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   532
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   533
lemmas neg_eq_less_0 = neg_def
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   534
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   535
lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   536
by (simp add: neg_def linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   537
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   538
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   539
subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   540
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   541
lemma not_neg_0: "~ neg 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   542
by (simp add: One_int_def neg_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   543
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   544
lemma not_neg_1: "~ neg 1"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   545
by (simp add: neg_def linorder_not_less zero_le_one)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   546
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   547
lemma iszero_0: "iszero 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   548
by (simp add: iszero_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   549
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   550
lemma not_iszero_1: "~ iszero 1"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   551
by (simp add: iszero_def eq_commute)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   552
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   553
lemma neg_nat: "neg z ==> nat z = 0"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   554
by (simp add: neg_def order_less_imp_le) 
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   555
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   556
lemma not_neg_nat: "~ neg z ==> int (nat z) = z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   557
by (simp add: linorder_not_less neg_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   558
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   559
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   560
subsection{*The Set of Natural Numbers*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   561
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   562
constdefs
21238
c46bc715bdfd generalized types of of_nat and of_int to work with non-commutative types
huffman
parents: 21210
diff changeset
   563
  Nats  :: "'a::semiring_1_cancel set"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   564
  "Nats == range of_nat"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   565
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 21191
diff changeset
   566
notation (xsymbols)
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19601
diff changeset
   567
  Nats  ("\<nat>")
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   568
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   569
lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   570
by (simp add: Nats_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   571
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   572
lemma Nats_0 [simp]: "0 \<in> Nats"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   573
apply (simp add: Nats_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   574
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   575
apply (rule of_nat_0 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   576
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   577
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   578
lemma Nats_1 [simp]: "1 \<in> Nats"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   579
apply (simp add: Nats_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   580
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   581
apply (rule of_nat_1 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   582
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   583
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   584
lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   585
apply (auto simp add: Nats_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   586
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   587
apply (rule of_nat_add [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   588
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   589
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   590
lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   591
apply (auto simp add: Nats_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   592
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   593
apply (rule of_nat_mult [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   594
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   595
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   596
text{*Agreement with the specific embedding for the integers*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   597
lemma int_eq_of_nat: "int = (of_nat :: nat => int)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   598
proof
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   599
  fix n
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   600
  show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   601
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   602
14496
paulson
parents: 14485
diff changeset
   603
lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)"
paulson
parents: 14485
diff changeset
   604
proof
paulson
parents: 14485
diff changeset
   605
  fix n
paulson
parents: 14485
diff changeset
   606
  show "of_nat n = id n"  by (induct n, simp_all)
paulson
parents: 14485
diff changeset
   607
qed
paulson
parents: 14485
diff changeset
   608
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   609
21238
c46bc715bdfd generalized types of of_nat and of_int to work with non-commutative types
huffman
parents: 21210
diff changeset
   610
subsection{*Embedding of the Integers into any @{text ring_1}:
14740
c8e1937110c2 fixed latex problems
nipkow
parents: 14738
diff changeset
   611
@{term of_int}*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   612
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   613
constdefs
21238
c46bc715bdfd generalized types of of_nat and of_int to work with non-commutative types
huffman
parents: 21210
diff changeset
   614
   of_int :: "int => 'a::ring_1"
14532
paulson
parents: 14496
diff changeset
   615
   "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   616
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   617
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   618
lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
14496
paulson
parents: 14485
diff changeset
   619
proof -
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   620
  have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
14496
paulson
parents: 14485
diff changeset
   621
    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
paulson
parents: 14485
diff changeset
   622
            del: of_nat_add) 
paulson
parents: 14485
diff changeset
   623
  thus ?thesis
paulson
parents: 14485
diff changeset
   624
    by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
paulson
parents: 14485
diff changeset
   625
qed
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   626
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   627
lemma of_int_0 [simp]: "of_int 0 = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   628
by (simp add: of_int Zero_int_def int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   629
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   630
lemma of_int_1 [simp]: "of_int 1 = 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   631
by (simp add: of_int One_int_def int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   632
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   633
lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   634
by (cases w, cases z, simp add: compare_rls of_int add)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   635
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   636
lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   637
by (cases z, simp add: compare_rls of_int minus)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   638
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   639
lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   640
by (simp add: diff_minus)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   641
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   642
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   643
apply (cases w, cases z)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   644
apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   645
                 mult add_ac)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   646
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   647
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   648
lemma of_int_le_iff [simp]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   649
     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   650
apply (cases w)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   651
apply (cases z)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   652
apply (simp add: compare_rls of_int le diff_int_def add minus
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   653
                 of_nat_add [symmetric]   del: of_nat_add)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   654
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   655
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   656
text{*Special cases where either operand is zero*}
17085
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   657
lemmas of_int_0_le_iff = of_int_le_iff [of 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   658
lemmas of_int_le_0_iff = of_int_le_iff [of _ 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   659
declare of_int_0_le_iff [simp]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   660
declare of_int_le_0_iff [simp]
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   661
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   662
lemma of_int_less_iff [simp]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   663
     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   664
by (simp add: linorder_not_le [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   665
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   666
text{*Special cases where either operand is zero*}
17085
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   667
lemmas of_int_0_less_iff = of_int_less_iff [of 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   668
lemmas of_int_less_0_iff = of_int_less_iff [of _ 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   669
declare of_int_0_less_iff [simp]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   670
declare of_int_less_0_iff [simp]
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   671
21238
c46bc715bdfd generalized types of of_nat and of_int to work with non-commutative types
huffman
parents: 21210
diff changeset
   672
text{*The ordering on the @{text ring_1} is necessary.
14740
c8e1937110c2 fixed latex problems
nipkow
parents: 14738
diff changeset
   673
 See @{text of_nat_eq_iff} above.*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   674
lemma of_int_eq_iff [simp]:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14691
diff changeset
   675
     "(of_int w = (of_int z::'a::ordered_idom)) = (w = z)"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   676
by (simp add: order_eq_iff)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   677
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   678
text{*Special cases where either operand is zero*}
17085
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   679
lemmas of_int_0_eq_iff = of_int_eq_iff [of 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   680
lemmas of_int_eq_0_iff = of_int_eq_iff [of _ 0, simplified]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   681
declare of_int_0_eq_iff [simp]
5b57f995a179 more simprules now have names
paulson
parents: 16770
diff changeset
   682
declare of_int_eq_0_iff [simp]
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   683
14496
paulson
parents: 14485
diff changeset
   684
lemma of_int_eq_id [simp]: "of_int = (id :: int => int)"
paulson
parents: 14485
diff changeset
   685
proof
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   686
  fix z
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   687
  show "of_int z = id z"  
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   688
    by (cases z)
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   689
      (simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus)
14496
paulson
parents: 14485
diff changeset
   690
qed
paulson
parents: 14485
diff changeset
   691
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   692
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   693
subsection{*The Set of Integers*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   694
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   695
constdefs
21238
c46bc715bdfd generalized types of of_nat and of_int to work with non-commutative types
huffman
parents: 21210
diff changeset
   696
  Ints  :: "'a::ring_1 set"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   697
  "Ints == range of_int"
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   698
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 21191
diff changeset
   699
notation (xsymbols)
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19601
diff changeset
   700
  Ints  ("\<int>")
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   701
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   702
lemma Ints_0 [simp]: "0 \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   703
apply (simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   704
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   705
apply (rule of_int_0 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   706
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   707
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   708
lemma Ints_1 [simp]: "1 \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   709
apply (simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   710
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   711
apply (rule of_int_1 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   712
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   713
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   714
lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   715
apply (auto simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   716
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   717
apply (rule of_int_add [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   718
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   719
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   720
lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   721
apply (auto simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   722
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   723
apply (rule of_int_minus [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   724
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   725
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   726
lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   727
apply (auto simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   728
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   729
apply (rule of_int_diff [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   730
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   731
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   732
lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   733
apply (auto simp add: Ints_def)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   734
apply (rule range_eqI)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   735
apply (rule of_int_mult [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   736
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   737
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   738
text{*Collapse nested embeddings*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   739
lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   740
by (induct n, auto)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   741
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   742
lemma of_int_int_eq [simp]: "of_int (int n) = of_nat n"
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   743
by (simp add: int_eq_of_nat)
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   744
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   745
lemma Ints_cases [cases set: Ints]:
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   746
  assumes "q \<in> \<int>"
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   747
  obtains (of_int) z where "q = of_int z"
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   748
  unfolding Ints_def
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   749
proof -
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   750
  from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   751
  then obtain z where "q = of_int z" ..
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   752
  then show thesis ..
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   753
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   754
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   755
lemma Ints_induct [case_names of_int, induct set: Ints]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   756
  "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   757
  by (rule Ints_cases) auto
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   758
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   759
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   760
(* int (Suc n) = 1 + int n *)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   761
declare int_Suc [simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   762
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   763
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   764
subsection{*More Properties of @{term setsum} and  @{term setprod}*}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   765
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   766
text{*By Jeremy Avigad*}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   767
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   768
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   769
lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   770
  apply (cases "finite A")
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   771
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   772
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   773
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   774
lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   775
  apply (cases "finite A")
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   776
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   777
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   778
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   779
lemma int_setsum: "int (setsum f A) = (\<Sum>x\<in>A. int(f x))"
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   780
  by (simp add: int_eq_of_nat of_nat_setsum)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   781
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   782
lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   783
  apply (cases "finite A")
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   784
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   785
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   786
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   787
lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19348
diff changeset
   788
  apply (cases "finite A")
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   789
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   790
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   791
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   792
lemma int_setprod: "int (setprod f A) = (\<Prod>x\<in>A. int(f x))"
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15540
diff changeset
   793
  by (simp add: int_eq_of_nat of_nat_setprod)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   794
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   795
lemma setprod_nonzero_nat:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   796
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   797
  by (rule setprod_nonzero, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   798
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   799
lemma setprod_zero_eq_nat:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   800
    "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   801
  by (rule setprod_zero_eq, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   802
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   803
lemma setprod_nonzero_int:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   804
    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   805
  by (rule setprod_nonzero, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   806
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   807
lemma setprod_zero_eq_int:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   808
    "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   809
  by (rule setprod_zero_eq, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   810
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   811
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   812
text{*Now we replace the case analysis rule by a more conventional one:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   813
whether an integer is negative or not.*}
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   814
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   815
lemma zless_iff_Suc_zadd:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   816
    "(w < z) = (\<exists>n. z = w + int(Suc n))"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   817
apply (cases z, cases w)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   818
apply (auto simp add: le add int_def linorder_not_le [symmetric]) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   819
apply (rename_tac a b c d) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   820
apply (rule_tac x="a+d - Suc(c+b)" in exI) 
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   821
apply arith
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   822
done
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   823
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   824
lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   825
apply (cases x)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   826
apply (auto simp add: le minus Zero_int_def int_def order_less_le) 
14496
paulson
parents: 14485
diff changeset
   827
apply (rule_tac x="y - Suc x" in exI, arith)
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   828
done
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   829
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   830
theorem int_cases [cases type: int, case_names nonneg neg]:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   831
     "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   832
apply (case_tac "z < 0", blast dest!: negD)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   833
apply (simp add: linorder_not_less)
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   834
apply (blast dest: nat_0_le [THEN sym])
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   835
done
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   836
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   837
theorem int_induct [induct type: int, case_names nonneg neg]:
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   838
     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   839
  by (cases z) auto
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   840
15558
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   841
text{*Contributed by Brian Huffman*}
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   842
theorem int_diff_cases [case_names diff]:
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   843
assumes prem: "!!m n. z = int m - int n ==> P" shows "P"
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   844
 apply (rule_tac z=z in int_cases)
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   845
  apply (rule_tac m=n and n=0 in prem, simp)
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   846
 apply (rule_tac m=0 and n="Suc n" in prem, simp)
f5f4f89a3b84 new lemmas int_diff_cases
paulson
parents: 15554
diff changeset
   847
done
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
   848
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   849
lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   850
apply (cases z)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   851
apply (simp_all add: not_zle_0_negative del: int_Suc)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   852
done
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   853
22392
35f54980d4cc tuned code theorems
haftmann
parents: 21911
diff changeset
   854
lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq]
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   855
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   856
subsection {* Configuration of the code generator *}
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   857
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   858
(*FIXME: the IntInf.fromInt below hides a dependence on fixed-precision ints!*)
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   859
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   860
types_code
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   861
  "int" ("int")
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   862
attach (term_of) {*
21820
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   863
val term_of_int = HOLogic.mk_number HOLogic.intT o IntInf.fromInt;
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   864
*}
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   865
attach (test) {*
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   866
fun gen_int i = one_of [~1, 1] * random_range 0 i;
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16733
diff changeset
   867
*}
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   868
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   869
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21243
diff changeset
   870
  int_aux :: "int \<Rightarrow> nat \<Rightarrow> int" where
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   871
  "int_aux i n = (i + int n)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21243
diff changeset
   872
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21243
diff changeset
   873
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21243
diff changeset
   874
  nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat" where
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   875
  "nat_aux n i = (n + nat i)"
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   876
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   877
lemma [code]:
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   878
  "int_aux i 0 = i"
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   879
  "int_aux i (Suc n) = int_aux (i + 1) n" -- {* tail recursive *}
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   880
  "int n = int_aux 0 n"
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   881
  by (simp add: int_aux_def)+
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   882
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   883
lemma [code]: "nat_aux n i = (if i <= 0 then n else nat_aux (Suc n) (i - 1))"
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   884
  -- {* tail recursive *}
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   885
  by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   886
    dest: zless_imp_add1_zle)
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   887
lemma [code]: "nat i = nat_aux 0 i"
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   888
  by (simp add: nat_aux_def)
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   889
20105
454f4be984b7 adaptions in codegen
haftmann
parents: 19887
diff changeset
   890
lemma [code inline]:
22392
35f54980d4cc tuned code theorems
haftmann
parents: 21911
diff changeset
   891
  "neg k \<longleftrightarrow> k < 0"
19601
299d4cd2ef51 added codegen preprocessors for numerals
haftmann
parents: 19535
diff changeset
   892
  unfolding neg_def ..
299d4cd2ef51 added codegen preprocessors for numerals
haftmann
parents: 19535
diff changeset
   893
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   894
lemma [code func]:
22392
35f54980d4cc tuned code theorems
haftmann
parents: 21911
diff changeset
   895
  "\<bar>k\<Colon>int\<bar> = (if k < 0 then - k else k)"
35f54980d4cc tuned code theorems
haftmann
parents: 21911
diff changeset
   896
  unfolding zabs_def ..
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   897
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   898
consts_code
20713
823967ef47f1 renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents: 20595
diff changeset
   899
  "HOL.zero" :: "int"                ("0")
823967ef47f1 renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
haftmann
parents: 20595
diff changeset
   900
  "HOL.one" :: "int"                 ("1")
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19138
diff changeset
   901
  "HOL.uminus" :: "int => int"       ("~")
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19138
diff changeset
   902
  "HOL.plus" :: "int => int => int"  ("(_ +/ _)")
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19138
diff changeset
   903
  "HOL.times" :: "int => int => int" ("(_ */ _)")
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   904
  "Orderings.less" :: "int => int => bool" ("(_ </ _)")
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   905
  "Orderings.less_eq" :: "int => int => bool" ("(_ <=/ _)")
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19138
diff changeset
   906
  "neg"                              ("(_ < 0)")
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   907
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   908
instance int :: eq ..
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   909
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   910
code_type int
21113
5b76e541cc0a adapted to new serializer syntax
haftmann
parents: 21079
diff changeset
   911
  (SML "IntInf.int")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   912
  (OCaml "Big'_int.big'_int")
21113
5b76e541cc0a adapted to new serializer syntax
haftmann
parents: 21079
diff changeset
   913
  (Haskell "Integer")
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   914
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   915
code_instance int :: eq
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   916
  (Haskell -)
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   917
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21404
diff changeset
   918
code_const "op = \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
21113
5b76e541cc0a adapted to new serializer syntax
haftmann
parents: 21079
diff changeset
   919
  (SML "!((_ : IntInf.int) = _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   920
  (OCaml "Big'_int.eq'_big'_int")
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   921
  (Haskell infixl 4 "==")
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   922
21552
da4e5237dda2 small syntax tuning
haftmann
parents: 21454
diff changeset
   923
code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   924
  (SML "IntInf.<= (_, _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   925
  (OCaml "Big'_int.le'_big'_int")
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   926
  (Haskell infix 4 "<=")
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   927
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   928
code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   929
  (SML "IntInf.< (_, _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   930
  (OCaml "Big'_int.lt'_big'_int")
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   931
  (Haskell infix 4 "<")
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   932
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   933
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   934
  (SML "IntInf.+ (_, _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   935
  (OCaml "Big'_int.add'_big'_int")
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   936
  (Haskell infixl 6 "+")
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   937
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   938
code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   939
  (SML "IntInf.- (_, _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   940
  (OCaml "Big'_int.sub'_big'_int")
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   941
  (Haskell infixl 6 "-")
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   942
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   943
code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   944
  (SML "IntInf.* (_, _)")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   945
  (OCaml "Big'_int.mult'_big'_int")
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   946
  (Haskell infixl 7 "*")
18702
7dc7dcd63224 substantial improvements in code generator
haftmann
parents: 18518
diff changeset
   947
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20485
diff changeset
   948
code_const "uminus \<Colon> int \<Rightarrow> int"
21113
5b76e541cc0a adapted to new serializer syntax
haftmann
parents: 21079
diff changeset
   949
  (SML "IntInf.~")
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   950
  (OCaml "Big'_int.minus'_big'_int")
21113
5b76e541cc0a adapted to new serializer syntax
haftmann
parents: 21079
diff changeset
   951
  (Haskell "negate")
20453
855f07fabd76 final syntax for some Isar code generator keywords
haftmann
parents: 20432
diff changeset
   952
21079
747d716e98d0 added reserved words for Haskell
haftmann
parents: 21046
diff changeset
   953
code_reserved SML IntInf
21911
e29bcab0c81c added OCaml code generation (without dictionaries)
haftmann
parents: 21820
diff changeset
   954
code_reserved OCaml Big_int
21079
747d716e98d0 added reserved words for Haskell
haftmann
parents: 21046
diff changeset
   955
21191
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   956
code_modulename SML
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   957
  IntDef Integer
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   958
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   959
code_modulename Haskell
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   960
  IntDef Integer
c00161fbf990 code generator module naming improved
haftmann
parents: 21113
diff changeset
   961
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   962
ML {*
21820
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   963
fun number_of_codegen thy defs gr dep module b (Const ("Numeral.number_of", Type ("fun", [_, T])) $ t) =
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   964
      if T = HOLogic.intT then
17551
2a747fc49a8c Simplified code generator for numerals.
berghofe
parents: 17464
diff changeset
   965
        (SOME (fst (Codegen.invoke_tycodegen thy defs dep module false (gr, T)),
21820
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   966
          (Pretty.str o IntInf.toString o HOLogic.dest_numeral) t) handle TERM _ => NONE)
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   967
      else if T = HOLogic.natT then
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   968
        SOME (Codegen.invoke_codegen thy defs dep module b (gr,
18115
5d0243c9a302 (fix for accidental commit)
haftmann
parents: 18114
diff changeset
   969
          Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $
21820
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   970
            (Const ("Numeral.number_of", HOLogic.intT --> HOLogic.intT) $ t)))
2f2b6a965ccc introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents: 21552
diff changeset
   971
      else NONE
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   972
  | number_of_codegen _ _ _ _ _ _ _ = NONE;
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   973
*}
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   974
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18704
diff changeset
   975
setup {*
19601
299d4cd2ef51 added codegen preprocessors for numerals
haftmann
parents: 19535
diff changeset
   976
  Codegen.add_codegen "number_of_codegen" number_of_codegen
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18704
diff changeset
   977
*}
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   978
17464
a4090ccf14a8 added quickcheck_params (from Main.thy);
wenzelm
parents: 17085
diff changeset
   979
quickcheck_params [default_type = int]
a4090ccf14a8 added quickcheck_params (from Main.thy);
wenzelm
parents: 17085
diff changeset
   980
16642
849ec3962b55 Moved code generator setup from NatBin to IntDef.
berghofe
parents: 16413
diff changeset
   981
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   982
(*Legacy ML bindings, but no longer the structure Int.*)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   983
ML
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   984
{*
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   985
val zabs_def = thm "zabs_def"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   986
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   987
val int_0 = thm "int_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   988
val int_1 = thm "int_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   989
val int_Suc0_eq_1 = thm "int_Suc0_eq_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   990
val neg_eq_less_0 = thm "neg_eq_less_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   991
val not_neg_eq_ge_0 = thm "not_neg_eq_ge_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   992
val not_neg_0 = thm "not_neg_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   993
val not_neg_1 = thm "not_neg_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   994
val iszero_0 = thm "iszero_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   995
val not_iszero_1 = thm "not_iszero_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   996
val int_0_less_1 = thm "int_0_less_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   997
val int_0_neq_1 = thm "int_0_neq_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   998
val negative_zless = thm "negative_zless";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   999
val negative_zle = thm "negative_zle";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1000
val not_zle_0_negative = thm "not_zle_0_negative";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1001
val not_int_zless_negative = thm "not_int_zless_negative";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1002
val negative_eq_positive = thm "negative_eq_positive";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1003
val zle_iff_zadd = thm "zle_iff_zadd";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1004
val abs_int_eq = thm "abs_int_eq";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1005
val abs_split = thm"abs_split";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1006
val nat_int = thm "nat_int";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1007
val nat_zminus_int = thm "nat_zminus_int";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1008
val nat_zero = thm "nat_zero";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1009
val not_neg_nat = thm "not_neg_nat";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1010
val neg_nat = thm "neg_nat";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1011
val zless_nat_eq_int_zless = thm "zless_nat_eq_int_zless";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1012
val nat_0_le = thm "nat_0_le";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1013
val nat_le_0 = thm "nat_le_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1014
val zless_nat_conj = thm "zless_nat_conj";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1015
val int_cases = thm "int_cases";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1016
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1017
val int_def = thm "int_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1018
val Zero_int_def = thm "Zero_int_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1019
val One_int_def = thm "One_int_def";
14479
0eca4aabf371 streamlined treatment of quotients for the integers
paulson
parents: 14430
diff changeset
  1020
val diff_int_def = thm "diff_int_def";
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1021
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1022
val inj_int = thm "inj_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1023
val zminus_zminus = thm "zminus_zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1024
val zminus_0 = thm "zminus_0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1025
val zminus_zadd_distrib = thm "zminus_zadd_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1026
val zadd_commute = thm "zadd_commute";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1027
val zadd_assoc = thm "zadd_assoc";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1028
val zadd_left_commute = thm "zadd_left_commute";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1029
val zadd_ac = thms "zadd_ac";
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
  1030
val zmult_ac = thms "zmult_ac";
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1031
val zadd_int = thm "zadd_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1032
val zadd_int_left = thm "zadd_int_left";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1033
val int_Suc = thm "int_Suc";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1034
val zadd_0 = thm "zadd_0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1035
val zadd_0_right = thm "zadd_0_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1036
val zmult_zminus = thm "zmult_zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1037
val zmult_commute = thm "zmult_commute";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1038
val zmult_assoc = thm "zmult_assoc";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1039
val zadd_zmult_distrib = thm "zadd_zmult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1040
val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1041
val zdiff_zmult_distrib = thm "zdiff_zmult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1042
val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1043
val int_distrib = thms "int_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1044
val zmult_int = thm "zmult_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1045
val zmult_1 = thm "zmult_1";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1046
val zmult_1_right = thm "zmult_1_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1047
val int_int_eq = thm "int_int_eq";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1048
val int_eq_0_conv = thm "int_eq_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1049
val zless_int = thm "zless_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1050
val int_less_0_conv = thm "int_less_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1051
val zero_less_int_conv = thm "zero_less_int_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1052
val zle_int = thm "zle_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1053
val zero_zle_int = thm "zero_zle_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1054
val int_le_0_conv = thm "int_le_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1055
val zle_refl = thm "zle_refl";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1056
val zle_linear = thm "zle_linear";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1057
val zle_trans = thm "zle_trans";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1058
val zle_anti_sym = thm "zle_anti_sym";
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1059
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1060
val Ints_def = thm "Ints_def";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1061
val Nats_def = thm "Nats_def";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1062
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1063
val of_nat_0 = thm "of_nat_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1064
val of_nat_Suc = thm "of_nat_Suc";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1065
val of_nat_1 = thm "of_nat_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1066
val of_nat_add = thm "of_nat_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1067
val of_nat_mult = thm "of_nat_mult";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1068
val zero_le_imp_of_nat = thm "zero_le_imp_of_nat";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1069
val less_imp_of_nat_less = thm "less_imp_of_nat_less";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1070
val of_nat_less_imp_less = thm "of_nat_less_imp_less";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1071
val of_nat_less_iff = thm "of_nat_less_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1072
val of_nat_le_iff = thm "of_nat_le_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1073
val of_nat_eq_iff = thm "of_nat_eq_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1074
val Nats_0 = thm "Nats_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1075
val Nats_1 = thm "Nats_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1076
val Nats_add = thm "Nats_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1077
val Nats_mult = thm "Nats_mult";
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1078
val int_eq_of_nat = thm"int_eq_of_nat";
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1079
val of_int = thm "of_int";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1080
val of_int_0 = thm "of_int_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1081
val of_int_1 = thm "of_int_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1082
val of_int_add = thm "of_int_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1083
val of_int_minus = thm "of_int_minus";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1084
val of_int_diff = thm "of_int_diff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1085
val of_int_mult = thm "of_int_mult";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1086
val of_int_le_iff = thm "of_int_le_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1087
val of_int_less_iff = thm "of_int_less_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1088
val of_int_eq_iff = thm "of_int_eq_iff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1089
val Ints_0 = thm "Ints_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1090
val Ints_1 = thm "Ints_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1091
val Ints_add = thm "Ints_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1092
val Ints_minus = thm "Ints_minus";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1093
val Ints_diff = thm "Ints_diff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1094
val Ints_mult = thm "Ints_mult";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1095
val of_int_of_nat_eq = thm"of_int_of_nat_eq";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1096
val Ints_cases = thm "Ints_cases";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1097
val Ints_induct = thm "Ints_induct";
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1098
*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1099
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
  1100
end