src/HOL/Integ/IntDef.thy
author paulson
Thu, 04 Mar 2004 12:06:07 +0100
changeset 14430 5cb24165a2e1
parent 14421 ee97b6463cb4
child 14479 0eca4aabf371
permissions -rw-r--r--
new material from Avigad, and simplified treatment of division by 0
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
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
     8
header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
     9
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    10
theory IntDef = Equiv + NatArith:
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    11
constdefs
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    12
  intrel :: "((nat * nat) * (nat * nat)) set"
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    13
    "intrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    14
691c70898518 new files in Integ
paulson
parents:
diff changeset
    15
typedef (Integ)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    16
  int = "UNIV//intrel"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    17
    by (auto simp add: quotient_def) 
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    18
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    19
instance int :: ord ..
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    20
instance int :: zero ..
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    21
instance int :: one ..
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    22
instance int :: plus ..
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    23
instance int :: times ..
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    24
instance int :: minus ..
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    25
691c70898518 new files in Integ
paulson
parents:
diff changeset
    26
constdefs
691c70898518 new files in Integ
paulson
parents:
diff changeset
    27
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    28
  int :: "nat => int"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    29
  "int m == Abs_Integ(intrel `` {(m,0)})"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    30
  
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14259
diff changeset
    31
defs (overloaded)
9392
c8e6529cc082 changed / to // for quotienting
paulson
parents: 8937
diff changeset
    32
  
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    33
  zminus_def:    "- Z == Abs_Integ(\<Union>(x,y) \<in> Rep_Integ(Z). intrel``{(y,x)})"
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    34
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    35
  Zero_int_def:  "0 == int 0"
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    36
  One_int_def:   "1 == int 1"
8937
7328d7c8d201 defining 0::int to be (int 0)
paulson
parents: 7375
diff changeset
    37
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    38
  zadd_def:
5594
e4439230af67 much tidying
paulson
parents: 5582
diff changeset
    39
   "z + w == 
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    40
       Abs_Integ(\<Union>(x1,y1) \<in> Rep_Integ(z). \<Union>(x2,y2) \<in> Rep_Integ(w).   
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    41
		 intrel``{(x1+x2, y1+y2)})"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    42
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    43
  zdiff_def:  "z - (w::int) == z + (-w)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    44
  zmult_def:
5594
e4439230af67 much tidying
paulson
parents: 5582
diff changeset
    45
   "z * w == 
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    46
       Abs_Integ(\<Union>(x1,y1) \<in> Rep_Integ(z). \<Union>(x2,y2) \<in> Rep_Integ(w).   
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    47
		 intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)})"
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
    48
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    49
  zless_def: "(z < (w::int)) == (z \<le> w & z \<noteq> w)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    50
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    51
  zle_def:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    52
  "z \<le> (w::int) == \<exists>x1 y1 x2 y2. x1 + y2 \<le> x2 + y1 &
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    53
                            (x1,y1) \<in> Rep_Integ z & (x2,y2) \<in> Rep_Integ w"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
    54
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
    55
lemma intrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in>  intrel) = (x1+y2 = x2+y1)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    56
by (unfold intrel_def, blast)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    57
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    58
lemma equiv_intrel: "equiv UNIV intrel"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    59
by (unfold intrel_def equiv_def refl_def sym_def trans_def, auto)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    60
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    61
lemmas equiv_intrel_iff =
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    62
       eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I, simp]
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 intrel_in_integ [simp]: "intrel``{(x,y)}:Integ"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    65
by (unfold Integ_def intrel_def quotient_def, fast)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    66
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    67
lemma inj_on_Abs_Integ: "inj_on Abs_Integ Integ"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    68
apply (rule inj_on_inverseI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    69
apply (erule Abs_Integ_inverse)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    70
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    71
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    72
declare inj_on_Abs_Integ [THEN inj_on_iff, simp] 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    73
        Abs_Integ_inverse [simp]
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    74
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    75
lemma inj_Rep_Integ: "inj(Rep_Integ)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    76
apply (rule inj_on_inverseI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    77
apply (rule Rep_Integ_inverse)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    78
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    79
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    80
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    81
(** int: the injection from "nat" to "int" **)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    82
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    83
lemma inj_int: "inj int"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    84
apply (rule inj_onI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    85
apply (unfold int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    86
apply (drule inj_on_Abs_Integ [THEN inj_onD])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    87
apply (rule intrel_in_integ)+
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    88
apply (drule eq_equiv_class)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    89
apply (rule equiv_intrel, fast)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    90
apply (simp add: intrel_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    91
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    92
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    93
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
    94
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
    95
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
    96
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    97
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    98
subsection{*zminus: unary negation on Integ*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
    99
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   100
lemma zminus_congruent: "congruent intrel (%(x,y). intrel``{(y,x)})"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   101
apply (unfold congruent_def intrel_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   102
apply (auto simp add: add_ac)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   103
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   104
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   105
lemma zminus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   106
by (simp add: zminus_def equiv_intrel [THEN UN_equiv_class] zminus_congruent)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   107
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   108
(*Every integer can be written in the form Abs_Integ(...) *)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   109
lemma eq_Abs_Integ: 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   110
     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   111
apply (rule_tac x1=z in Rep_Integ [unfolded Integ_def, THEN quotientE]) 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   112
apply (drule_tac f = Abs_Integ in arg_cong)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   113
apply (rule_tac p = x in PairE)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   114
apply (simp add: Rep_Integ_inverse)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   115
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   116
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   117
lemma zminus_zminus [simp]: "- (- z) = (z::int)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   118
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   119
apply (simp add: zminus)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   120
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   121
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   122
lemma inj_zminus: "inj(%z::int. -z)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   123
apply (rule inj_onI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   124
apply (drule_tac f = uminus in arg_cong, simp)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   125
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   126
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   127
lemma zminus_0 [simp]: "- 0 = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   128
by (simp add: int_def Zero_int_def zminus)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   129
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   130
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   131
subsection{*zadd: addition on Integ*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   132
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   133
lemma zadd: 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   134
  "Abs_Integ(intrel``{(x1,y1)}) + Abs_Integ(intrel``{(x2,y2)}) =  
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   135
   Abs_Integ(intrel``{(x1+x2, y1+y2)})"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   136
apply (simp add: zadd_def UN_UN_split_split_eq)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   137
apply (subst equiv_intrel [THEN UN_equiv_class2])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   138
apply (auto simp add: congruent2_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   139
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   140
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   141
lemma zminus_zadd_distrib [simp]: "- (z + w) = (- z) + (- w::int)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   142
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   143
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   144
apply (simp add: zminus zadd)
14259
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
lemma zadd_commute: "(z::int) + w = w + z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   148
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   149
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   150
apply (simp add: add_ac zadd)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   151
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   152
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   153
lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   154
apply (rule eq_Abs_Integ [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   155
apply (rule eq_Abs_Integ [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   156
apply (rule eq_Abs_Integ [of z3])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   157
apply (simp add: zadd add_assoc)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   158
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   159
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   160
(*For AC rewriting*)
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   161
lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   162
  apply (rule mk_left_commute [of "op +"])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   163
  apply (rule zadd_assoc)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   164
  apply (rule zadd_commute)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   165
  done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   166
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   167
(*Integer addition is an AC operator*)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   168
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   169
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   170
lemmas zmult_ac = Ring_and_Field.mult_ac
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   171
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   172
lemma zadd_int: "(int m) + (int n) = int (m + n)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   173
by (simp add: int_def zadd)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   174
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   175
lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   176
by (simp add: zadd_int zadd_assoc [symmetric])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   177
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   178
lemma int_Suc: "int (Suc m) = 1 + (int m)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   179
by (simp add: One_int_def zadd_int)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   180
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   181
(*also for the instance declaration int :: plus_ac0*)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   182
lemma zadd_0 [simp]: "(0::int) + z = z"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   183
apply (unfold Zero_int_def int_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   184
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   185
apply (simp add: zadd)
14259
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
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   188
lemma zadd_0_right [simp]: "z + (0::int) = z"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   189
by (rule trans [OF zadd_commute zadd_0])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   190
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   191
lemma zadd_zminus_inverse [simp]: "z + (- z) = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   192
apply (unfold int_def Zero_int_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   193
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   194
apply (simp add: zminus zadd add_commute)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   195
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   196
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   197
lemma zadd_zminus_inverse2 [simp]: "(- z) + z = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   198
apply (rule zadd_commute [THEN trans])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   199
apply (rule zadd_zminus_inverse)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   200
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   201
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   202
lemma zadd_zminus_cancel [simp]: "z + (- z + w) = (w::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   203
by (simp add: zadd_assoc [symmetric] zadd_0)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   204
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   205
lemma zminus_zadd_cancel [simp]: "(-z) + (z + w) = (w::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   206
by (simp add: zadd_assoc [symmetric] zadd_0)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   207
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   208
lemma zdiff0 [simp]: "(0::int) - x = -x"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   209
by (simp add: zdiff_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   210
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   211
lemma zdiff0_right [simp]: "x - (0::int) = x"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   212
by (simp add: zdiff_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   213
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   214
lemma zdiff_self [simp]: "x - x = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   215
by (simp add: zdiff_def Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   216
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   217
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   218
(** Lemmas **)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   219
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   220
lemma zadd_assoc_cong: "(z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   221
by (simp add: zadd_assoc [symmetric])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   222
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   223
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   224
subsection{*zmult: multiplication on Integ*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   225
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   226
text{*Congruence property for multiplication*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   227
lemma zmult_congruent2: "congruent2 intrel  
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   228
        (%p1 p2. (%(x1,y1). (%(x2,y2).    
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   229
                    intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   230
apply (rule equiv_intrel [THEN congruent2_commuteI])
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   231
 apply (force simp add: add_ac mult_ac) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   232
apply (clarify, simp del: equiv_intrel_iff add: add_ac mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   233
apply (rename_tac x1 x2 y1 y2 z1 z2)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   234
apply (rule equiv_class_eq [OF equiv_intrel intrel_iff [THEN iffD2]])
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   235
apply (subgoal_tac "x1*z1 + y2*z1 = y1*z1 + x2*z1 & x1*z2 + y2*z2 = y1*z2 + x2*z2")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   236
apply (simp add: mult_ac, arith) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   237
apply (simp add: add_mult_distrib [symmetric])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   238
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   239
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   240
lemma zmult: 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   241
   "Abs_Integ((intrel``{(x1,y1)})) * Abs_Integ((intrel``{(x2,y2)})) =    
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   242
    Abs_Integ(intrel `` {(x1*x2 + y1*y2, x1*y2 + y1*x2)})"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   243
by (simp add: zmult_def UN_UN_split_split_eq zmult_congruent2 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   244
              equiv_intrel [THEN UN_equiv_class2])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   245
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   246
lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   247
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   248
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   249
apply (simp add: zminus zmult add_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   250
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   251
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   252
lemma zmult_commute: "(z::int) * w = w * z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   253
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   254
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   255
apply (simp add: zmult add_ac mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   256
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   257
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   258
lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   259
apply (rule eq_Abs_Integ [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   260
apply (rule eq_Abs_Integ [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   261
apply (rule eq_Abs_Integ [of z3])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   262
apply (simp add: add_mult_distrib2 zmult add_ac mult_ac)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   263
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   264
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   265
lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   266
apply (rule eq_Abs_Integ [of z1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   267
apply (rule eq_Abs_Integ [of z2])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   268
apply (rule eq_Abs_Integ [of w])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   269
apply (simp add: add_mult_distrib2 zadd zmult add_ac mult_ac)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   270
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   271
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   272
lemma zmult_zminus_right: "w * (- z) = - (w * (z::int))"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   273
by (simp add: zmult_commute [of w] zmult_zminus)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   274
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   275
lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   276
by (simp add: zmult_commute [of w] zadd_zmult_distrib)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   277
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   278
lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   279
apply (unfold zdiff_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   280
apply (subst zadd_zmult_distrib)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   281
apply (simp add: zmult_zminus)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   282
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   283
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   284
lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   285
by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   286
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   287
lemmas int_distrib =
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   288
  zadd_zmult_distrib zadd_zmult_distrib2 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   289
  zdiff_zmult_distrib zdiff_zmult_distrib2
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 zmult_int: "(int m) * (int n) = int (m * n)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   292
by (simp add: int_def zmult)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   293
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   294
lemma zmult_0 [simp]: "0 * z = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   295
apply (unfold Zero_int_def int_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   296
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   297
apply (simp add: zmult)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   298
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   299
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   300
lemma zmult_1 [simp]: "(1::int) * z = z"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   301
apply (unfold One_int_def int_def)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   302
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   303
apply (simp add: zmult)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   304
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   305
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   306
lemma zmult_0_right [simp]: "z * 0 = (0::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   307
by (rule trans [OF zmult_commute zmult_0])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   308
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   309
lemma zmult_1_right [simp]: "z * (1::int) = z"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   310
by (rule trans [OF zmult_commute zmult_1])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   311
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   312
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   313
text{*The Integers Form A Ring*}
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   314
instance int :: ring
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   315
proof
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   316
  fix i j k :: int
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   317
  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
   318
  show "i + j = j + i" by (simp add: zadd_commute)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   319
  show "0 + i = i" by simp
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   320
  show "- i + i = 0" by simp
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   321
  show "i - j = i + (-j)" by (simp add: zdiff_def)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   322
  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
   323
  show "i * j = j * i" by (rule zmult_commute)
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   324
  show "1 * i = i" by simp
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   325
  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   326
  show "0 \<noteq> (1::int)" 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   327
    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
   328
qed
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   329
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   330
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   331
subsection{*The @{text "\<le>"} Ordering*}
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
lemma zle: 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   334
  "(Abs_Integ(intrel``{(x1,y1)}) \<le> Abs_Integ(intrel``{(x2,y2)})) =  
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   335
   (x1 + y2 \<le> x2 + y1)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   336
by (force simp add: zle_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   337
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   338
lemma zle_refl: "w \<le> (w::int)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   339
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   340
apply (force simp add: zle)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   341
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   342
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   343
lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   344
apply (rule eq_Abs_Integ [of i]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   345
apply (rule eq_Abs_Integ [of j]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   346
apply (rule eq_Abs_Integ [of k]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   347
apply (simp add: zle) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   348
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   349
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   350
lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   351
apply (rule eq_Abs_Integ [of w]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   352
apply (rule eq_Abs_Integ [of z]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   353
apply (simp add: zle) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   354
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   355
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   356
(* Axiom 'order_less_le' of class 'order': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   357
lemma zless_le: "((w::int) < z) = (w \<le> z & w \<noteq> z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   358
by (simp add: zless_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   359
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   360
instance int :: order
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   361
proof qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   362
 (assumption |
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   363
  rule zle_refl zle_trans zle_anti_sym zless_le)+
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   364
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   365
(* Axiom 'linorder_linear' of class 'linorder': *)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   366
lemma zle_linear: "(z::int) \<le> w | w \<le> z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   367
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   368
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   369
apply (simp add: zle linorder_linear) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   370
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   371
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   372
instance int :: linorder
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   373
proof qed (rule zle_linear)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   374
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   375
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   376
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
   377
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
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
   380
by (simp add: Zero_int_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   381
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   382
(*This lemma allows direct proofs of other <-properties*)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   383
lemma zless_iff_Suc_zadd: 
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   384
    "(w < z) = (\<exists>n. z = w + int(Suc n))"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   385
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   386
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   387
apply (simp add: linorder_not_le [where 'a = int, symmetric] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   388
                 linorder_not_le [where 'a = nat] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   389
                 zle int_def zdiff_def zadd zminus) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   390
apply (safe dest!: less_imp_Suc_add)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   391
apply (rule_tac x = k in exI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   392
apply (simp_all add: add_ac)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   393
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   394
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   395
lemma zless_int [simp]: "(int m < int n) = (m<n)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   396
by (simp add: less_iff_Suc_add zless_iff_Suc_zadd zadd_int)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   397
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   398
lemma int_less_0_conv [simp]: "~ (int k < 0)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   399
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   400
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   401
lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   402
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   403
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   404
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
   405
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
   406
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   407
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
   408
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
   409
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   410
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
   411
by (simp add: linorder_not_less [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   412
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   413
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
   414
by (simp add: Zero_int_def)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   415
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   416
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
   417
by (simp add: Zero_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   418
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   419
lemma int_0 [simp]: "int 0 = (0::int)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   420
by (simp add: Zero_int_def)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   421
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   422
lemma int_1 [simp]: "int 1 = 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   423
by (simp add: One_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   424
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   425
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
   426
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
   427
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   428
subsection{*Monotonicity results*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   429
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   430
lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   431
apply (rule eq_Abs_Integ [of i]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   432
apply (rule eq_Abs_Integ [of j]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   433
apply (rule eq_Abs_Integ [of k]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   434
apply (simp add: zle zadd) 
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
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   437
lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   438
apply (rule eq_Abs_Integ [of i]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   439
apply (rule eq_Abs_Integ [of j]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   440
apply (rule eq_Abs_Integ [of k]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   441
apply (simp add: linorder_not_le [where 'a = int, symmetric] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   442
                 linorder_not_le [where 'a = nat]  zle zadd)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   443
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   444
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   445
lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   446
by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   447
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   448
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   449
subsection{*Strict Monotonicity of Multiplication*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   450
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   451
text{*strict, in 1st argument; proof is by induction on k>0*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   452
lemma zmult_zless_mono2_lemma [rule_format]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   453
     "i<j ==> 0<k --> int k * i < int k * j"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   454
apply (induct_tac "k", simp) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   455
apply (simp add: int_Suc)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   456
apply (case_tac "n=0")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   457
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
   458
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
   459
done
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   460
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   461
lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   462
apply (rule eq_Abs_Integ [of k]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   463
apply (auto simp add: zle zadd int_def Zero_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   464
apply (rule_tac x="x-y" in exI, simp) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   465
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   466
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   467
lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   468
apply (frule order_less_imp_le [THEN zero_le_imp_eq_int]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   469
apply (auto simp add: zmult_zless_mono2_lemma) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   470
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   471
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   472
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   473
defs (overloaded)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   474
    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
   475
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
text{*The Integers Form an Ordered Ring*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   478
instance int :: ordered_ring
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   479
proof
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   480
  fix i j k :: int
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   481
  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
   482
  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
   483
  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
   484
qed
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
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   487
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
   488
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   489
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   490
   nat  :: "int => nat"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   491
    "nat(Z) == if Z<0 then 0 else (THE m. Z = int m)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   492
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   493
lemma nat_int [simp]: "nat(int n) = n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   494
by (unfold nat_def, auto)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   495
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   496
lemma nat_zero [simp]: "nat 0 = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   497
apply (unfold Zero_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   498
apply (rule nat_int)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   499
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   500
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   501
lemma nat_0_le [simp]: "0 \<le> z ==> int (nat z) = z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   502
apply (rule eq_Abs_Integ [of z]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   503
apply (simp add: nat_def linorder_not_le [symmetric] zle int_def Zero_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   504
apply (subgoal_tac "(THE m. x = m + y) = x-y")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   505
apply (auto simp add: the_equality) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   506
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   507
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   508
lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   509
by (simp add: nat_def  order_less_le eq_commute [of 0])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   510
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   511
text{*An alternative condition is @{term "0 \<le> w"} *}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   512
lemma nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   513
apply (subst zless_int [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   514
apply (simp add: order_le_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   515
apply (case_tac "w < 0")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   516
 apply (simp add: order_less_imp_le)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   517
apply (simp add: linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   518
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   519
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   520
lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   521
apply (case_tac "0 < z")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   522
apply (auto simp add: nat_mono_iff linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   523
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   524
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   525
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   526
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
   527
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   528
lemma negative_zless_0: "- (int (Suc n)) < 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   529
by (simp add: zless_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   530
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   531
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
   532
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
   533
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   534
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
   535
by (simp add: minus_le_iff)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   536
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   537
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
   538
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
   539
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   540
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
   541
by (subst le_minus_iff, simp)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   542
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   543
lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   544
apply safe 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   545
apply (drule_tac [2] le_minus_iff [THEN iffD1])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   546
apply (auto dest: zle_trans [OF _ negative_zle_0]) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   547
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   548
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   549
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
   550
by (simp add: linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   551
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   552
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
   553
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
   554
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   555
lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   556
by (force intro: exI [of _ "0::nat"] 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   557
            intro!: not_sym [THEN not0_implies_Suc]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   558
            simp add: zless_iff_Suc_zadd order_le_less)
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
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   561
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
   562
      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
   563
      in theory @{text Ring_and_Field}.
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   564
      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
   565
lemma abs_split [arith_split]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   566
     "P(abs(a::'a::ordered_ring)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   567
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
   568
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   569
lemma abs_int_eq [simp]: "abs (int m) = int m"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   570
by (simp add: zabs_def)
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
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   573
subsection{*Misc Results*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   574
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   575
lemma nat_zminus_int [simp]: "nat(- (int n)) = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   576
by (auto simp add: nat_def zero_reorient minus_less_iff)
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 zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   579
apply (case_tac "0 \<le> z")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   580
apply (erule nat_0_le [THEN subst], simp) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   581
apply (simp add: linorder_not_le)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   582
apply (auto dest: order_less_trans simp add: order_less_imp_le)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   583
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   584
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   585
text{*A case theorem distinguishing non-negative and negative int*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   586
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   587
lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   588
by (auto simp add: zless_iff_Suc_zadd 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   589
                   diff_eq_eq [symmetric] zdiff_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   590
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   591
lemma int_cases [cases type: int, case_names nonneg neg]: 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   592
     "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   593
apply (case_tac "z < 0", blast dest!: negD)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   594
apply (simp add: linorder_not_less)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   595
apply (blast dest: nat_0_le [THEN sym])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   596
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   597
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   598
lemma int_induct [induct type: int, case_names nonneg neg]: 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   599
     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   600
  by (cases z) auto
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   601
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   602
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   603
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
   604
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   605
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   606
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   607
  neg   :: "'a::ordered_ring => bool"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   608
  "neg(Z) == Z < 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   609
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   610
  (*For simplifying equalities*)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   611
  iszero :: "'a::semiring => bool"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   612
  "iszero z == z = (0)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   613
  
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   614
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   615
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
   616
by (simp add: neg_def)
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 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
   619
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
   620
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   621
lemmas neg_eq_less_0 = neg_def
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   622
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   623
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
   624
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
   625
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   626
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
   627
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   628
lemma not_neg_0: "~ neg 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   629
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
   630
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   631
lemma not_neg_1: "~ neg 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   632
by (simp add: neg_def linorder_not_less zero_le_one) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   633
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   634
lemma iszero_0: "iszero 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   635
by (simp add: iszero_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   636
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   637
lemma not_iszero_1: "~ iszero 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   638
by (simp add: iszero_def eq_commute) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   639
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   640
lemma neg_nat: "neg z ==> nat z = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   641
by (simp add: nat_def neg_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   642
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   643
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
   644
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
   645
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   646
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   647
subsection{*Embedding of the Naturals into any Semiring: @{term of_nat}*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   648
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   649
consts of_nat :: "nat => 'a::semiring"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   650
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   651
primrec
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   652
  of_nat_0:   "of_nat 0 = 0"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   653
  of_nat_Suc: "of_nat (Suc m) = of_nat m + 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   654
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   655
lemma of_nat_1 [simp]: "of_nat 1 = 1"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   656
by simp
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   657
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   658
lemma of_nat_add [simp]: "of_nat (m+n) = of_nat m + of_nat n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   659
apply (induct m)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   660
apply (simp_all add: add_ac) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   661
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   662
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   663
lemma of_nat_mult [simp]: "of_nat (m*n) = of_nat m * of_nat n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   664
apply (induct m) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   665
apply (simp_all add: mult_ac add_ac right_distrib) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   666
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   667
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   668
lemma zero_le_imp_of_nat: "0 \<le> (of_nat m::'a::ordered_semiring)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   669
apply (induct m, simp_all) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   670
apply (erule order_trans) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   671
apply (rule less_add_one [THEN order_less_imp_le]) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   672
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   673
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   674
lemma less_imp_of_nat_less:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   675
     "m < n ==> of_nat m < (of_nat n::'a::ordered_semiring)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   676
apply (induct m n rule: diff_induct, simp_all) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   677
apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   678
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   679
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   680
lemma of_nat_less_imp_less:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   681
     "of_nat m < (of_nat n::'a::ordered_semiring) ==> m < n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   682
apply (induct m n rule: diff_induct, simp_all) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   683
apply (insert zero_le_imp_of_nat) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   684
apply (force simp add: linorder_not_less [symmetric]) 
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   685
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   686
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   687
lemma of_nat_less_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   688
     "(of_nat m < (of_nat n::'a::ordered_semiring)) = (m<n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   689
by (blast intro: of_nat_less_imp_less less_imp_of_nat_less ) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   690
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   691
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   692
declare of_nat_less_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   693
declare of_nat_less_iff [of _ 0, simplified, simp]
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
lemma of_nat_le_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   696
     "(of_nat m \<le> (of_nat n::'a::ordered_semiring)) = (m \<le> n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   697
by (simp add: linorder_not_less [symmetric]) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   698
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   699
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   700
declare of_nat_le_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   701
declare of_nat_le_iff [of _ 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   702
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   703
text{*The ordering on the semiring is necessary to exclude the possibility of
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   704
a finite field, which indeed wraps back to zero.*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   705
lemma of_nat_eq_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   706
     "(of_nat m = (of_nat n::'a::ordered_semiring)) = (m = n)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   707
by (simp add: order_eq_iff) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   708
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   709
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   710
declare of_nat_eq_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   711
declare of_nat_eq_iff [of _ 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   712
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   713
lemma of_nat_diff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   714
     "n \<le> m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::ring)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   715
by (simp del: of_nat_add
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   716
	 add: compare_rls of_nat_add [symmetric] split add: nat_diff_split) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   717
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   718
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   719
subsection{*The Set of Natural Numbers*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   720
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   721
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   722
   Nats  :: "'a::semiring set"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   723
    "Nats == range of_nat"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   724
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   725
syntax (xsymbols)    Nats :: "'a set"   ("\<nat>")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   726
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   727
lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   728
by (simp add: Nats_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   729
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   730
lemma Nats_0 [simp]: "0 \<in> Nats"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   731
apply (simp add: Nats_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   732
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   733
apply (rule of_nat_0 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   734
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   735
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   736
lemma Nats_1 [simp]: "1 \<in> Nats"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   737
apply (simp add: Nats_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   738
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   739
apply (rule of_nat_1 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   740
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   741
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   742
lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   743
apply (auto simp add: Nats_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   744
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   745
apply (rule of_nat_add [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   746
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   747
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   748
lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   749
apply (auto simp add: Nats_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   750
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   751
apply (rule of_nat_mult [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   752
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   753
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   754
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
   755
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
   756
proof
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   757
  fix n
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   758
  show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   759
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   760
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   761
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   762
subsection{*Embedding of the Integers into any Ring: @{term of_int}*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   763
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   764
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   765
   of_int :: "int => 'a::ring"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   766
   "of_int z ==
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   767
      (THE a. \<exists>i j. (i,j) \<in> Rep_Integ z & a = (of_nat i) - (of_nat j))"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   768
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   769
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   770
lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   771
apply (simp add: of_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   772
apply (rule the_equality, auto) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   773
apply (simp add: compare_rls add_ac of_nat_add [symmetric]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   774
            del: of_nat_add) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   775
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   776
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   777
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
   778
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
   779
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   780
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
   781
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
   782
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   783
lemma of_int_add [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
   784
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   785
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   786
apply (simp add: compare_rls of_int zadd) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   787
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   788
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   789
lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   790
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   791
apply (simp add: compare_rls of_int zminus) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   792
done
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   793
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   794
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
   795
by (simp add: diff_minus)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   796
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   797
lemma of_int_mult [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
   798
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   799
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   800
apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   801
                 zmult add_ac) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   802
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   803
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   804
lemma of_int_le_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   805
     "(of_int w \<le> (of_int z::'a::ordered_ring)) = (w \<le> z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   806
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   807
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   808
apply (simp add: compare_rls of_int zle zdiff_def zadd zminus 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   809
                 of_nat_add [symmetric]   del: of_nat_add) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   810
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   811
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   812
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   813
declare of_int_le_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   814
declare of_int_le_iff [of _ 0, simplified, simp]
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   815
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   816
lemma of_int_less_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   817
     "(of_int w < (of_int z::'a::ordered_ring)) = (w < z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   818
by (simp add: linorder_not_le [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   819
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   820
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   821
declare of_int_less_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   822
declare of_int_less_iff [of _ 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   823
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   824
text{*The ordering on the ring is necessary. 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
   825
lemma of_int_eq_iff [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   826
     "(of_int w = (of_int z::'a::ordered_ring)) = (w = z)"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   827
by (simp add: order_eq_iff) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   828
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   829
text{*Special cases where either operand is zero*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   830
declare of_int_eq_iff [of 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   831
declare of_int_eq_iff [of _ 0, simplified, simp]
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   832
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   833
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   834
subsection{*The Set of Integers*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   835
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   836
constdefs
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   837
   Ints  :: "'a::ring set"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   838
    "Ints == range of_int"
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14269
diff changeset
   839
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   840
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   841
syntax (xsymbols)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   842
  Ints      :: "'a set"                   ("\<int>")
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   843
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   844
lemma Ints_0 [simp]: "0 \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   845
apply (simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   846
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   847
apply (rule of_int_0 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   848
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   849
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   850
lemma Ints_1 [simp]: "1 \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   851
apply (simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   852
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   853
apply (rule of_int_1 [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   854
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   855
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   856
lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   857
apply (auto simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   858
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   859
apply (rule of_int_add [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   860
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   861
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   862
lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   863
apply (auto simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   864
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   865
apply (rule of_int_minus [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   866
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   867
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   868
lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   869
apply (auto simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   870
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   871
apply (rule of_int_diff [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   872
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   873
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   874
lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   875
apply (auto simp add: Ints_def) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   876
apply (rule range_eqI) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   877
apply (rule of_int_mult [symmetric])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   878
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   879
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   880
text{*Collapse nested embeddings*}
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   881
lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   882
by (induct n, auto) 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   883
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   884
lemma of_int_int_eq [simp]: "of_int (int n) = int n"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   885
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
   886
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14271
diff changeset
   887
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   888
lemma Ints_cases [case_names of_int, cases set: Ints]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   889
  "q \<in> \<int> ==> (!!z. q = of_int z ==> C) ==> C"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   890
proof (unfold Ints_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   891
  assume "!!z. q = of_int z ==> C"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   892
  assume "q \<in> range of_int" thus C ..
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   893
qed
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   894
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   895
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
   896
  "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
   897
  by (rule Ints_cases) auto
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   898
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   899
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   900
(* int (Suc n) = 1 + int n *)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   901
declare int_Suc [simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   902
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   903
text{*Simplification of @{term "x-y < 0"}, etc.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   904
declare less_iff_diff_less_0 [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   905
declare eq_iff_diff_eq_0 [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   906
declare le_iff_diff_le_0 [symmetric, simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   907
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   908
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   909
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
   910
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   911
text{*By Jeremy Avigad*}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   912
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   913
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   914
lemma setsum_of_nat: "of_nat (setsum f A) = setsum (of_nat \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   915
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   916
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   917
  apply (simp add: setsum_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   918
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   919
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   920
lemma setsum_of_int: "of_int (setsum f A) = setsum (of_int \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   921
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   922
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   923
  apply (simp add: setsum_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   924
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   925
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   926
lemma int_setsum: "int (setsum f A) = setsum (int \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   927
  by (subst int_eq_of_nat, rule setsum_of_nat)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   928
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   929
lemma setprod_of_nat: "of_nat (setprod f A) = setprod (of_nat \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   930
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   931
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   932
  apply (simp add: setprod_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   933
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   934
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   935
lemma setprod_of_int: "of_int (setprod f A) = setprod (of_int \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   936
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   937
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   938
  apply (simp add: setprod_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   939
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   940
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   941
lemma int_setprod: "int (setprod f A) = setprod (int \<circ> f) A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   942
  by (subst int_eq_of_nat, rule setprod_of_nat)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   943
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   944
lemma setsum_constant: "finite A ==> (\<Sum>x \<in> A. y) = of_nat(card A) * y"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   945
  apply (erule finite_induct)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   946
  apply (auto simp add: ring_distrib add_ac)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   947
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   948
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   949
lemma setprod_nonzero_nat:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   950
    "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
   951
  by (rule setprod_nonzero, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   952
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   953
lemma setprod_zero_eq_nat:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   954
    "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
   955
  by (rule setprod_zero_eq, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   956
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   957
lemma setprod_nonzero_int:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   958
    "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
   959
  by (rule setprod_nonzero, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   960
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   961
lemma setprod_zero_eq_int:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   962
    "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
   963
  by (rule setprod_zero_eq, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   964
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   965
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   966
(*Legacy ML bindings, but no longer the structure Int.*)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   967
ML
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
   968
{*
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   969
val zabs_def = thm "zabs_def"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   970
val nat_def  = thm "nat_def"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   971
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   972
val int_0 = thm "int_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   973
val int_1 = thm "int_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   974
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
   975
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
   976
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
   977
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
   978
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
   979
val iszero_0 = thm "iszero_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   980
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
   981
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
   982
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
   983
val negative_zless = thm "negative_zless";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   984
val negative_zle = thm "negative_zle";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   985
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
   986
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
   987
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
   988
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
   989
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
   990
val abs_split = thm"abs_split";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   991
val nat_int = thm "nat_int";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   992
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
   993
val nat_zero = thm "nat_zero";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   994
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
   995
val neg_nat = thm "neg_nat";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
   996
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
   997
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
   998
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
   999
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
  1000
val int_cases = thm "int_cases";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1001
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1002
val int_def = thm "int_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1003
val Zero_int_def = thm "Zero_int_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1004
val One_int_def = thm "One_int_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1005
val zadd_def = thm "zadd_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1006
val zdiff_def = thm "zdiff_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1007
val zless_def = thm "zless_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1008
val zle_def = thm "zle_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1009
val zmult_def = thm "zmult_def";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1010
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1011
val intrel_iff = thm "intrel_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1012
val equiv_intrel = thm "equiv_intrel";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1013
val equiv_intrel_iff = thm "equiv_intrel_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1014
val intrel_in_integ = thm "intrel_in_integ";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1015
val inj_on_Abs_Integ = thm "inj_on_Abs_Integ";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1016
val inj_Rep_Integ = thm "inj_Rep_Integ";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1017
val inj_int = thm "inj_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1018
val zminus_congruent = thm "zminus_congruent";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1019
val zminus = thm "zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1020
val eq_Abs_Integ = thm "eq_Abs_Integ";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1021
val zminus_zminus = thm "zminus_zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1022
val inj_zminus = thm "inj_zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1023
val zminus_0 = thm "zminus_0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1024
val zadd = thm "zadd";
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 zadd_zminus_inverse = thm "zadd_zminus_inverse";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1037
val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1038
val zadd_zminus_cancel = thm "zadd_zminus_cancel";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1039
val zminus_zadd_cancel = thm "zminus_zadd_cancel";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1040
val zdiff0 = thm "zdiff0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1041
val zdiff0_right = thm "zdiff0_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1042
val zdiff_self = thm "zdiff_self";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1043
val zmult_congruent2 = thm "zmult_congruent2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1044
val zmult = thm "zmult";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1045
val zmult_zminus = thm "zmult_zminus";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1046
val zmult_commute = thm "zmult_commute";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1047
val zmult_assoc = thm "zmult_assoc";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1048
val zadd_zmult_distrib = thm "zadd_zmult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1049
val zmult_zminus_right = thm "zmult_zminus_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1050
val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1051
val zdiff_zmult_distrib = thm "zdiff_zmult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1052
val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1053
val int_distrib = thms "int_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1054
val zmult_int = thm "zmult_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1055
val zmult_0 = thm "zmult_0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1056
val zmult_1 = thm "zmult_1";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1057
val zmult_0_right = thm "zmult_0_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1058
val zmult_1_right = thm "zmult_1_right";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1059
val zless_iff_Suc_zadd = thm "zless_iff_Suc_zadd";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1060
val int_int_eq = thm "int_int_eq";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1061
val int_eq_0_conv = thm "int_eq_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1062
val zless_int = thm "zless_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1063
val int_less_0_conv = thm "int_less_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1064
val zero_less_int_conv = thm "zero_less_int_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1065
val zle_int = thm "zle_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1066
val zero_zle_int = thm "zero_zle_int";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1067
val int_le_0_conv = thm "int_le_0_conv";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1068
val zle_refl = thm "zle_refl";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1069
val zle_linear = thm "zle_linear";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1070
val zle_trans = thm "zle_trans";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1071
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
  1072
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1073
val Ints_def = thm "Ints_def";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1074
val Nats_def = thm "Nats_def";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1075
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1076
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
  1077
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
  1078
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
  1079
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
  1080
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
  1081
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
  1082
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
  1083
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
  1084
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
  1085
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
  1086
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
  1087
val Nats_0 = thm "Nats_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1088
val Nats_1 = thm "Nats_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1089
val Nats_add = thm "Nats_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1090
val Nats_mult = thm "Nats_mult";
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1091
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
  1092
val of_int = thm "of_int";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1093
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
  1094
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
  1095
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
  1096
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
  1097
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
  1098
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
  1099
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
  1100
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
  1101
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
  1102
val Ints_0 = thm "Ints_0";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1103
val Ints_1 = thm "Ints_1";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1104
val Ints_add = thm "Ints_add";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1105
val Ints_minus = thm "Ints_minus";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1106
val Ints_diff = thm "Ints_diff";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1107
val Ints_mult = thm "Ints_mult";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1108
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
  1109
val Ints_cases = thm "Ints_cases";
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14348
diff changeset
  1110
val Ints_induct = thm "Ints_induct";
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1111
*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13487
diff changeset
  1112
5508
691c70898518 new files in Integ
paulson
parents:
diff changeset
  1113
end