src/HOL/Integ/IntDef.thy
author obua
Tue, 11 May 2004 20:11:08 +0200
changeset 14738 83f1a514dcb4
parent 14691 e1eedc8cad37
child 14740 c8e1937110c2
permissions -rw-r--r--
changes made due to new Ring_and_Field theory
     1 (*  Title:      IntDef.thy
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1996  University of Cambridge
     5 
     6 *)
     7 
     8 header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
     9 
    10 theory IntDef = Equiv + NatArith:
    11 
    12 constdefs
    13   intrel :: "((nat * nat) * (nat * nat)) set"
    14     --{*the equivalence relation underlying the integers*}
    15     "intrel == {((x,y),(u,v)) | x y u v. x+v = u+y}"
    16 
    17 typedef (Integ)
    18   int = "UNIV//intrel"
    19     by (auto simp add: quotient_def)
    20 
    21 instance int :: "{ord, zero, one, plus, times, minus}" ..
    22 
    23 constdefs
    24   int :: "nat => int"
    25   "int m == Abs_Integ(intrel `` {(m,0)})"
    26 
    27 
    28 defs (overloaded)
    29 
    30   Zero_int_def:  "0 == int 0"
    31   One_int_def:   "1 == int 1"
    32 
    33   minus_int_def:
    34     "- z == Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. intrel``{(y,x)})"
    35 
    36   add_int_def:
    37    "z + w ==
    38        Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w.
    39 		 intrel``{(x+u, y+v)})"
    40 
    41   diff_int_def:  "z - (w::int) == z + (-w)"
    42 
    43   mult_int_def:
    44    "z * w ==
    45        Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w.
    46 		  intrel``{(x*u + y*v, x*v + y*u)})"
    47 
    48   le_int_def:
    49    "z \<le> (w::int) == 
    50     \<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Integ z & (u,v) \<in> Rep_Integ w"
    51 
    52   less_int_def: "(z < (w::int)) == (z \<le> w & z \<noteq> w)"
    53 
    54 
    55 subsection{*Construction of the Integers*}
    56 
    57 subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
    58 
    59 lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)"
    60 by (simp add: intrel_def)
    61 
    62 lemma equiv_intrel: "equiv UNIV intrel"
    63 by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
    64 
    65 text{*Reduces equality of equivalence classes to the @{term intrel} relation:
    66   @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
    67 lemmas equiv_intrel_iff = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I]
    68 
    69 declare equiv_intrel_iff [simp]
    70 
    71 
    72 text{*All equivalence classes belong to set of representatives*}
    73 lemma [simp]: "intrel``{(x,y)} \<in> Integ"
    74 by (auto simp add: Integ_def intrel_def quotient_def)
    75 
    76 lemma inj_on_Abs_Integ: "inj_on Abs_Integ Integ"
    77 apply (rule inj_on_inverseI)
    78 apply (erule Abs_Integ_inverse)
    79 done
    80 
    81 text{*This theorem reduces equality on abstractions to equality on 
    82       representatives:
    83   @{term "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
    84 declare inj_on_Abs_Integ [THEN inj_on_iff, simp]
    85 
    86 declare Abs_Integ_inverse [simp]
    87 
    88 text{*Case analysis on the representation of an integer as an equivalence
    89       class of pairs of naturals.*}
    90 lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
    91      "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
    92 apply (rule Rep_Integ [of z, unfolded Integ_def, THEN quotientE])
    93 apply (drule arg_cong [where f=Abs_Integ])
    94 apply (auto simp add: Rep_Integ_inverse)
    95 done
    96 
    97 
    98 subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
    99 
   100 lemma inj_int: "inj int"
   101 by (simp add: inj_on_def int_def)
   102 
   103 lemma int_int_eq [iff]: "(int m = int n) = (m = n)"
   104 by (fast elim!: inj_int [THEN injD])
   105 
   106 
   107 subsubsection{*Integer Unary Negation*}
   108 
   109 lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
   110 proof -
   111   have "congruent intrel (\<lambda>(x,y). intrel``{(y,x)})"
   112     by (simp add: congruent_def) 
   113   thus ?thesis
   114     by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
   115 qed
   116 
   117 lemma zminus_zminus: "- (- z) = (z::int)"
   118 by (cases z, simp add: minus)
   119 
   120 lemma zminus_0: "- 0 = (0::int)"
   121 by (simp add: int_def Zero_int_def minus)
   122 
   123 
   124 subsection{*Integer Addition*}
   125 
   126 lemma add:
   127      "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
   128       Abs_Integ (intrel``{(x+u, y+v)})"
   129 proof -
   130   have "congruent2 intrel intrel
   131         (\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z)"
   132     by (simp add: congruent2_def)
   133   thus ?thesis
   134     by (simp add: add_int_def UN_UN_split_split_eq
   135                   UN_equiv_class2 [OF equiv_intrel equiv_intrel])
   136 qed
   137 
   138 lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)"
   139 by (cases z, cases w, simp add: minus add)
   140 
   141 lemma zadd_commute: "(z::int) + w = w + z"
   142 by (cases z, cases w, simp add: add_ac add)
   143 
   144 lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
   145 by (cases z1, cases z2, cases z3, simp add: add add_assoc)
   146 
   147 (*For AC rewriting*)
   148 lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
   149   apply (rule mk_left_commute [of "op +"])
   150   apply (rule zadd_assoc)
   151   apply (rule zadd_commute)
   152   done
   153 
   154 lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
   155 
   156 lemmas zmult_ac = OrderedGroup.mult_ac
   157 
   158 lemma zadd_int: "(int m) + (int n) = int (m + n)"
   159 by (simp add: int_def add)
   160 
   161 lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
   162 by (simp add: zadd_int zadd_assoc [symmetric])
   163 
   164 lemma int_Suc: "int (Suc m) = 1 + (int m)"
   165 by (simp add: One_int_def zadd_int)
   166 
   167 (*also for the instance declaration int :: comm_monoid_add*)
   168 lemma zadd_0: "(0::int) + z = z"
   169 apply (simp add: Zero_int_def int_def)
   170 apply (cases z, simp add: add)
   171 done
   172 
   173 lemma zadd_0_right: "z + (0::int) = z"
   174 by (rule trans [OF zadd_commute zadd_0])
   175 
   176 lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
   177 by (cases z, simp add: int_def Zero_int_def minus add)
   178 
   179 
   180 subsection{*Integer Multiplication*}
   181 
   182 text{*Congruence property for multiplication*}
   183 lemma mult_congruent2:
   184      "congruent2 intrel intrel
   185         (%p1 p2. (%(x,y). (%(u,v).
   186                     intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)"
   187 apply (rule equiv_intrel [THEN congruent2_commuteI])
   188  apply (force simp add: mult_ac, clarify) 
   189 apply (simp add: congruent_def mult_ac)  
   190 apply (rename_tac u v w x y z)
   191 apply (subgoal_tac "u*y + x*y = w*y + v*y  &  u*z + x*z = w*z + v*z")
   192 apply (simp add: mult_ac, arith)
   193 apply (simp add: add_mult_distrib [symmetric])
   194 done
   195 
   196 
   197 lemma mult:
   198      "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
   199       Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
   200 by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
   201               UN_equiv_class2 [OF equiv_intrel equiv_intrel])
   202 
   203 lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
   204 by (cases z, cases w, simp add: minus mult add_ac)
   205 
   206 lemma zmult_commute: "(z::int) * w = w * z"
   207 by (cases z, cases w, simp add: mult add_ac mult_ac)
   208 
   209 lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
   210 by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac)
   211 
   212 lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
   213 by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac)
   214 
   215 lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
   216 by (simp add: zmult_commute [of w] zadd_zmult_distrib)
   217 
   218 lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
   219 by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
   220 
   221 lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
   222 by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
   223 
   224 lemmas int_distrib =
   225   zadd_zmult_distrib zadd_zmult_distrib2
   226   zdiff_zmult_distrib zdiff_zmult_distrib2
   227 
   228 lemma zmult_int: "(int m) * (int n) = int (m * n)"
   229 by (simp add: int_def mult)
   230 
   231 lemma zmult_1: "(1::int) * z = z"
   232 by (cases z, simp add: One_int_def int_def mult)
   233 
   234 lemma zmult_1_right: "z * (1::int) = z"
   235 by (rule trans [OF zmult_commute zmult_1])
   236 
   237 
   238 text{*The Integers Form A comm_ring_1*}
   239 instance int :: comm_ring_1
   240 proof
   241   fix i j k :: int
   242   show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
   243   show "i + j = j + i" by (simp add: zadd_commute)
   244   show "0 + i = i" by (rule zadd_0)
   245   show "- i + i = 0" by (rule zadd_zminus_inverse2)
   246   show "i - j = i + (-j)" by (simp add: diff_int_def)
   247   show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
   248   show "i * j = j * i" by (rule zmult_commute)
   249   show "1 * i = i" by (rule zmult_1) 
   250   show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
   251   show "0 \<noteq> (1::int)"
   252     by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
   253 qed
   254 
   255 
   256 subsection{*The @{text "\<le>"} Ordering*}
   257 
   258 lemma le:
   259   "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
   260 by (force simp add: le_int_def)
   261 
   262 lemma zle_refl: "w \<le> (w::int)"
   263 by (cases w, simp add: le)
   264 
   265 lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
   266 by (cases i, cases j, cases k, simp add: le)
   267 
   268 lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
   269 by (cases w, cases z, simp add: le)
   270 
   271 (* Axiom 'order_less_le' of class 'order': *)
   272 lemma zless_le: "((w::int) < z) = (w \<le> z & w \<noteq> z)"
   273 by (simp add: less_int_def)
   274 
   275 instance int :: order
   276   by intro_classes
   277     (assumption |
   278       rule zle_refl zle_trans zle_anti_sym zless_le)+
   279 
   280 (* Axiom 'linorder_linear' of class 'linorder': *)
   281 lemma zle_linear: "(z::int) \<le> w | w \<le> z"
   282 by (cases z, cases w) (simp add: le linorder_linear)
   283 
   284 instance int :: linorder
   285   by intro_classes (rule zle_linear)
   286 
   287 
   288 lemmas zless_linear = linorder_less_linear [where 'a = int]
   289 
   290 
   291 lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)"
   292 by (simp add: Zero_int_def)
   293 
   294 lemma zless_int [simp]: "(int m < int n) = (m<n)"
   295 by (simp add: le add int_def linorder_not_le [symmetric]) 
   296 
   297 lemma int_less_0_conv [simp]: "~ (int k < 0)"
   298 by (simp add: Zero_int_def)
   299 
   300 lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
   301 by (simp add: Zero_int_def)
   302 
   303 lemma int_0_less_1: "0 < (1::int)"
   304 by (simp only: Zero_int_def One_int_def One_nat_def zless_int)
   305 
   306 lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)"
   307 by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
   308 
   309 lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)"
   310 by (simp add: linorder_not_less [symmetric])
   311 
   312 lemma zero_zle_int [simp]: "(0 \<le> int n)"
   313 by (simp add: Zero_int_def)
   314 
   315 lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)"
   316 by (simp add: Zero_int_def)
   317 
   318 lemma int_0 [simp]: "int 0 = (0::int)"
   319 by (simp add: Zero_int_def)
   320 
   321 lemma int_1 [simp]: "int 1 = 1"
   322 by (simp add: One_int_def)
   323 
   324 lemma int_Suc0_eq_1: "int (Suc 0) = 1"
   325 by (simp add: One_int_def One_nat_def)
   326 
   327 
   328 subsection{*Monotonicity results*}
   329 
   330 lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)"
   331 by (cases i, cases j, cases k, simp add: le add)
   332 
   333 lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)"
   334 apply (cases i, cases j, cases k)
   335 apply (simp add: linorder_not_le [where 'a = int, symmetric]
   336                  linorder_not_le [where 'a = nat]  le add)
   337 done
   338 
   339 lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
   340 by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono])
   341 
   342 
   343 subsection{*Strict Monotonicity of Multiplication*}
   344 
   345 text{*strict, in 1st argument; proof is by induction on k>0*}
   346 lemma zmult_zless_mono2_lemma [rule_format]:
   347      "i<j ==> 0<k --> int k * i < int k * j"
   348 apply (induct_tac "k", simp)
   349 apply (simp add: int_Suc)
   350 apply (case_tac "n=0")
   351 apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less)
   352 apply (simp add: zadd_zless_mono int_Suc0_eq_1 order_le_less)
   353 done
   354 
   355 lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n"
   356 apply (cases k)
   357 apply (auto simp add: le add int_def Zero_int_def)
   358 apply (rule_tac x="x-y" in exI, simp)
   359 done
   360 
   361 lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
   362 apply (frule order_less_imp_le [THEN zero_le_imp_eq_int])
   363 apply (auto simp add: zmult_zless_mono2_lemma)
   364 done
   365 
   366 
   367 defs (overloaded)
   368     zabs_def:  "abs(i::int) == if i < 0 then -i else i"
   369 
   370 
   371 text{*The Integers Form an Ordered comm_ring_1*}
   372 instance int :: ordered_idom
   373 proof
   374   fix i j k :: int
   375   show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono)
   376   show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2)
   377   show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def)
   378 qed
   379 
   380 
   381 lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
   382 apply (cases w, cases z) 
   383 apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def)
   384 done
   385 
   386 subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
   387 
   388 constdefs
   389    nat  :: "int => nat"
   390     "nat z == contents (\<Union>(x,y) \<in> Rep_Integ z. { x-y })"
   391 
   392 lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
   393 proof -
   394   have "congruent intrel (\<lambda>(x,y). {x-y})"
   395     by (simp add: congruent_def, arith) 
   396   thus ?thesis
   397     by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
   398 qed
   399 
   400 lemma nat_int [simp]: "nat(int n) = n"
   401 by (simp add: nat int_def) 
   402 
   403 lemma nat_zero [simp]: "nat 0 = 0"
   404 by (simp only: Zero_int_def nat_int)
   405 
   406 lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
   407 by (cases z, simp add: nat le int_def Zero_int_def)
   408 
   409 corollary nat_0_le: "0 \<le> z ==> int (nat z) = z"
   410 apply simp 
   411 done
   412 
   413 lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
   414 by (cases z, simp add: nat le int_def Zero_int_def)
   415 
   416 lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
   417 apply (cases w, cases z) 
   418 apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith)
   419 done
   420 
   421 text{*An alternative condition is @{term "0 \<le> w"} *}
   422 corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
   423 by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
   424 
   425 corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
   426 by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
   427 
   428 lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
   429 apply (cases w, cases z) 
   430 apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
   431 done
   432 
   433 lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
   434 by (blast dest: nat_0_le sym)
   435 
   436 lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
   437 by (cases w, simp add: nat le int_def Zero_int_def, arith)
   438 
   439 corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
   440 by (simp only: eq_commute [of m] nat_eq_iff) 
   441 
   442 lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
   443 apply (cases w)
   444 apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
   445 done
   446 
   447 lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
   448 by (auto simp add: nat_eq_iff2)
   449 
   450 lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
   451 by (insert zless_nat_conj [of 0], auto)
   452 
   453 
   454 lemma nat_add_distrib:
   455      "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
   456 by (cases z, cases z', simp add: nat add le int_def Zero_int_def)
   457 
   458 lemma nat_diff_distrib:
   459      "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
   460 by (cases z, cases z', 
   461     simp add: nat add minus diff_minus le int_def Zero_int_def)
   462 
   463 
   464 lemma nat_zminus_int [simp]: "nat (- (int n)) = 0"
   465 by (simp add: int_def minus nat Zero_int_def) 
   466 
   467 lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
   468 by (cases z, simp add: nat le int_def  linorder_not_le [symmetric], arith)
   469 
   470 
   471 subsection{*Lemmas about the Function @{term int} and Orderings*}
   472 
   473 lemma negative_zless_0: "- (int (Suc n)) < 0"
   474 by (simp add: order_less_le)
   475 
   476 lemma negative_zless [iff]: "- (int (Suc n)) < int m"
   477 by (rule negative_zless_0 [THEN order_less_le_trans], simp)
   478 
   479 lemma negative_zle_0: "- int n \<le> 0"
   480 by (simp add: minus_le_iff)
   481 
   482 lemma negative_zle [iff]: "- int n \<le> int m"
   483 by (rule order_trans [OF negative_zle_0 zero_zle_int])
   484 
   485 lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))"
   486 by (subst le_minus_iff, simp)
   487 
   488 lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
   489 by (simp add: int_def le minus Zero_int_def) 
   490 
   491 lemma not_int_zless_negative [simp]: "~ (int n < - int m)"
   492 by (simp add: linorder_not_less)
   493 
   494 lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)"
   495 by (force simp add: order_eq_iff [of "- int n"] int_zle_neg)
   496 
   497 lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
   498 apply (cases w, cases z)
   499 apply (auto simp add: le add int_def) 
   500 apply (rename_tac a b c d) 
   501 apply (rule_tac x="c+b - (a+d)" in exI) 
   502 apply arith
   503 done
   504 
   505 lemma abs_int_eq [simp]: "abs (int m) = int m"
   506 by (simp add: zabs_def)
   507 
   508 text{*This version is proved for all ordered rings, not just integers!
   509       It is proved here because attribute @{text arith_split} is not available
   510       in theory @{text Ring_and_Field}.
   511       But is it really better than just rewriting with @{text abs_if}?*}
   512 lemma abs_split [arith_split]:
   513      "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
   514 by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
   515 
   516 
   517 
   518 subsection{*The Constants @{term neg} and @{term iszero}*}
   519 
   520 constdefs
   521 
   522   neg   :: "'a::ordered_idom => bool"
   523   "neg(Z) == Z < 0"
   524 
   525   (*For simplifying equalities*)
   526   iszero :: "'a::comm_semiring_1_cancel => bool"
   527   "iszero z == z = (0)"
   528 
   529 
   530 lemma not_neg_int [simp]: "~ neg(int n)"
   531 by (simp add: neg_def)
   532 
   533 lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))"
   534 by (simp add: neg_def neg_less_0_iff_less)
   535 
   536 lemmas neg_eq_less_0 = neg_def
   537 
   538 lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
   539 by (simp add: neg_def linorder_not_less)
   540 
   541 
   542 subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
   543 
   544 lemma not_neg_0: "~ neg 0"
   545 by (simp add: One_int_def neg_def)
   546 
   547 lemma not_neg_1: "~ neg 1"
   548 by (simp add: neg_def linorder_not_less zero_le_one)
   549 
   550 lemma iszero_0: "iszero 0"
   551 by (simp add: iszero_def)
   552 
   553 lemma not_iszero_1: "~ iszero 1"
   554 by (simp add: iszero_def eq_commute)
   555 
   556 lemma neg_nat: "neg z ==> nat z = 0"
   557 by (simp add: neg_def order_less_imp_le) 
   558 
   559 lemma not_neg_nat: "~ neg z ==> int (nat z) = z"
   560 by (simp add: linorder_not_less neg_def)
   561 
   562 
   563 subsection{*Embedding of the Naturals into any comm_semiring_1_cancel: @{term of_nat}*}
   564 
   565 consts of_nat :: "nat => 'a::comm_semiring_1_cancel"
   566 
   567 primrec
   568   of_nat_0:   "of_nat 0 = 0"
   569   of_nat_Suc: "of_nat (Suc m) = of_nat m + 1"
   570 
   571 lemma of_nat_1 [simp]: "of_nat 1 = 1"
   572 by simp
   573 
   574 lemma of_nat_add [simp]: "of_nat (m+n) = of_nat m + of_nat n"
   575 apply (induct m)
   576 apply (simp_all add: add_ac)
   577 done
   578 
   579 lemma of_nat_mult [simp]: "of_nat (m*n) = of_nat m * of_nat n"
   580 apply (induct m)
   581 apply (simp_all add: mult_ac add_ac right_distrib)
   582 done
   583 
   584 lemma zero_le_imp_of_nat: "0 \<le> (of_nat m::'a::ordered_semidom)"
   585 apply (induct m, simp_all)
   586 apply (erule order_trans)
   587 apply (rule less_add_one [THEN order_less_imp_le])
   588 done
   589 
   590 lemma less_imp_of_nat_less:
   591      "m < n ==> of_nat m < (of_nat n::'a::ordered_semidom)"
   592 apply (induct m n rule: diff_induct, simp_all)
   593 apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force)
   594 done
   595 
   596 lemma of_nat_less_imp_less:
   597      "of_nat m < (of_nat n::'a::ordered_semidom) ==> m < n"
   598 apply (induct m n rule: diff_induct, simp_all)
   599 apply (insert zero_le_imp_of_nat)
   600 apply (force simp add: linorder_not_less [symmetric])
   601 done
   602 
   603 lemma of_nat_less_iff [simp]:
   604      "(of_nat m < (of_nat n::'a::ordered_semidom)) = (m<n)"
   605 by (blast intro: of_nat_less_imp_less less_imp_of_nat_less)
   606 
   607 text{*Special cases where either operand is zero*}
   608 declare of_nat_less_iff [of 0, simplified, simp]
   609 declare of_nat_less_iff [of _ 0, simplified, simp]
   610 
   611 lemma of_nat_le_iff [simp]:
   612      "(of_nat m \<le> (of_nat n::'a::ordered_semidom)) = (m \<le> n)"
   613 by (simp add: linorder_not_less [symmetric])
   614 
   615 text{*Special cases where either operand is zero*}
   616 declare of_nat_le_iff [of 0, simplified, simp]
   617 declare of_nat_le_iff [of _ 0, simplified, simp]
   618 
   619 text{*The ordering on the comm_semiring_1_cancel is necessary to exclude the possibility of
   620 a finite field, which indeed wraps back to zero.*}
   621 lemma of_nat_eq_iff [simp]:
   622      "(of_nat m = (of_nat n::'a::ordered_semidom)) = (m = n)"
   623 by (simp add: order_eq_iff)
   624 
   625 text{*Special cases where either operand is zero*}
   626 declare of_nat_eq_iff [of 0, simplified, simp]
   627 declare of_nat_eq_iff [of _ 0, simplified, simp]
   628 
   629 lemma of_nat_diff [simp]:
   630      "n \<le> m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::comm_ring_1)"
   631 by (simp del: of_nat_add
   632 	 add: compare_rls of_nat_add [symmetric] split add: nat_diff_split)
   633 
   634 
   635 subsection{*The Set of Natural Numbers*}
   636 
   637 constdefs
   638    Nats  :: "'a::comm_semiring_1_cancel set"
   639     "Nats == range of_nat"
   640 
   641 syntax (xsymbols)    Nats :: "'a set"   ("\<nat>")
   642 
   643 lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
   644 by (simp add: Nats_def)
   645 
   646 lemma Nats_0 [simp]: "0 \<in> Nats"
   647 apply (simp add: Nats_def)
   648 apply (rule range_eqI)
   649 apply (rule of_nat_0 [symmetric])
   650 done
   651 
   652 lemma Nats_1 [simp]: "1 \<in> Nats"
   653 apply (simp add: Nats_def)
   654 apply (rule range_eqI)
   655 apply (rule of_nat_1 [symmetric])
   656 done
   657 
   658 lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
   659 apply (auto simp add: Nats_def)
   660 apply (rule range_eqI)
   661 apply (rule of_nat_add [symmetric])
   662 done
   663 
   664 lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
   665 apply (auto simp add: Nats_def)
   666 apply (rule range_eqI)
   667 apply (rule of_nat_mult [symmetric])
   668 done
   669 
   670 text{*Agreement with the specific embedding for the integers*}
   671 lemma int_eq_of_nat: "int = (of_nat :: nat => int)"
   672 proof
   673   fix n
   674   show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac)
   675 qed
   676 
   677 lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)"
   678 proof
   679   fix n
   680   show "of_nat n = id n"  by (induct n, simp_all)
   681 qed
   682 
   683 
   684 subsection{*Embedding of the Integers into any comm_ring_1: @{term of_int}*}
   685 
   686 constdefs
   687    of_int :: "int => 'a::comm_ring_1"
   688    "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
   689 
   690 
   691 lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
   692 proof -
   693   have "congruent intrel (\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) })"
   694     by (simp add: congruent_def compare_rls of_nat_add [symmetric]
   695             del: of_nat_add) 
   696   thus ?thesis
   697     by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
   698 qed
   699 
   700 lemma of_int_0 [simp]: "of_int 0 = 0"
   701 by (simp add: of_int Zero_int_def int_def)
   702 
   703 lemma of_int_1 [simp]: "of_int 1 = 1"
   704 by (simp add: of_int One_int_def int_def)
   705 
   706 lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
   707 by (cases w, cases z, simp add: compare_rls of_int add)
   708 
   709 lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
   710 by (cases z, simp add: compare_rls of_int minus)
   711 
   712 lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
   713 by (simp add: diff_minus)
   714 
   715 lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
   716 apply (cases w, cases z)
   717 apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
   718                  mult add_ac)
   719 done
   720 
   721 lemma of_int_le_iff [simp]:
   722      "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
   723 apply (cases w)
   724 apply (cases z)
   725 apply (simp add: compare_rls of_int le diff_int_def add minus
   726                  of_nat_add [symmetric]   del: of_nat_add)
   727 done
   728 
   729 text{*Special cases where either operand is zero*}
   730 declare of_int_le_iff [of 0, simplified, simp]
   731 declare of_int_le_iff [of _ 0, simplified, simp]
   732 
   733 lemma of_int_less_iff [simp]:
   734      "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
   735 by (simp add: linorder_not_le [symmetric])
   736 
   737 text{*Special cases where either operand is zero*}
   738 declare of_int_less_iff [of 0, simplified, simp]
   739 declare of_int_less_iff [of _ 0, simplified, simp]
   740 
   741 text{*The ordering on the comm_ring_1 is necessary. See @{text of_nat_eq_iff} above.*}
   742 lemma of_int_eq_iff [simp]:
   743      "(of_int w = (of_int z::'a::ordered_idom)) = (w = z)"
   744 by (simp add: order_eq_iff)
   745 
   746 text{*Special cases where either operand is zero*}
   747 declare of_int_eq_iff [of 0, simplified, simp]
   748 declare of_int_eq_iff [of _ 0, simplified, simp]
   749 
   750 lemma of_int_eq_id [simp]: "of_int = (id :: int => int)"
   751 proof
   752  fix z
   753  show "of_int z = id z"  
   754   by (cases z,
   755       simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus)
   756 qed
   757 
   758 
   759 subsection{*The Set of Integers*}
   760 
   761 constdefs
   762    Ints  :: "'a::comm_ring_1 set"
   763     "Ints == range of_int"
   764 
   765 
   766 syntax (xsymbols)
   767   Ints      :: "'a set"                   ("\<int>")
   768 
   769 lemma Ints_0 [simp]: "0 \<in> Ints"
   770 apply (simp add: Ints_def)
   771 apply (rule range_eqI)
   772 apply (rule of_int_0 [symmetric])
   773 done
   774 
   775 lemma Ints_1 [simp]: "1 \<in> Ints"
   776 apply (simp add: Ints_def)
   777 apply (rule range_eqI)
   778 apply (rule of_int_1 [symmetric])
   779 done
   780 
   781 lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
   782 apply (auto simp add: Ints_def)
   783 apply (rule range_eqI)
   784 apply (rule of_int_add [symmetric])
   785 done
   786 
   787 lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
   788 apply (auto simp add: Ints_def)
   789 apply (rule range_eqI)
   790 apply (rule of_int_minus [symmetric])
   791 done
   792 
   793 lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
   794 apply (auto simp add: Ints_def)
   795 apply (rule range_eqI)
   796 apply (rule of_int_diff [symmetric])
   797 done
   798 
   799 lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
   800 apply (auto simp add: Ints_def)
   801 apply (rule range_eqI)
   802 apply (rule of_int_mult [symmetric])
   803 done
   804 
   805 text{*Collapse nested embeddings*}
   806 lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
   807 by (induct n, auto)
   808 
   809 lemma of_int_int_eq [simp]: "of_int (int n) = int n"
   810 by (simp add: int_eq_of_nat)
   811 
   812 
   813 lemma Ints_cases [case_names of_int, cases set: Ints]:
   814   "q \<in> \<int> ==> (!!z. q = of_int z ==> C) ==> C"
   815 proof (simp add: Ints_def)
   816   assume "!!z. q = of_int z ==> C"
   817   assume "q \<in> range of_int" thus C ..
   818 qed
   819 
   820 lemma Ints_induct [case_names of_int, induct set: Ints]:
   821   "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q"
   822   by (rule Ints_cases) auto
   823 
   824 
   825 (* int (Suc n) = 1 + int n *)
   826 declare int_Suc [simp]
   827 
   828 text{*Simplification of @{term "x-y < 0"}, etc.*}
   829 declare less_iff_diff_less_0 [symmetric, simp]
   830 declare eq_iff_diff_eq_0 [symmetric, simp]
   831 declare le_iff_diff_le_0 [symmetric, simp]
   832 
   833 
   834 subsection{*More Properties of @{term setsum} and  @{term setprod}*}
   835 
   836 text{*By Jeremy Avigad*}
   837 
   838 
   839 lemma setsum_of_nat: "of_nat (setsum f A) = setsum (of_nat \<circ> f) A"
   840   apply (case_tac "finite A")
   841   apply (erule finite_induct, auto)
   842   apply (simp add: setsum_def)
   843   done
   844 
   845 lemma setsum_of_int: "of_int (setsum f A) = setsum (of_int \<circ> f) A"
   846   apply (case_tac "finite A")
   847   apply (erule finite_induct, auto)
   848   apply (simp add: setsum_def)
   849   done
   850 
   851 lemma int_setsum: "int (setsum f A) = setsum (int \<circ> f) A"
   852   by (subst int_eq_of_nat, rule setsum_of_nat)
   853 
   854 lemma setprod_of_nat: "of_nat (setprod f A) = setprod (of_nat \<circ> f) A"
   855   apply (case_tac "finite A")
   856   apply (erule finite_induct, auto)
   857   apply (simp add: setprod_def)
   858   done
   859 
   860 lemma setprod_of_int: "of_int (setprod f A) = setprod (of_int \<circ> f) A"
   861   apply (case_tac "finite A")
   862   apply (erule finite_induct, auto)
   863   apply (simp add: setprod_def)
   864   done
   865 
   866 lemma int_setprod: "int (setprod f A) = setprod (int \<circ> f) A"
   867   by (subst int_eq_of_nat, rule setprod_of_nat)
   868 
   869 lemma setsum_constant: "finite A ==> (\<Sum>x \<in> A. y) = of_nat(card A) * y"
   870   apply (erule finite_induct)
   871   apply (auto simp add: ring_distrib add_ac)
   872   done
   873 
   874 lemma setprod_nonzero_nat:
   875     "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
   876   by (rule setprod_nonzero, auto)
   877 
   878 lemma setprod_zero_eq_nat:
   879     "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
   880   by (rule setprod_zero_eq, auto)
   881 
   882 lemma setprod_nonzero_int:
   883     "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
   884   by (rule setprod_nonzero, auto)
   885 
   886 lemma setprod_zero_eq_int:
   887     "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
   888   by (rule setprod_zero_eq, auto)
   889 
   890 
   891 text{*Now we replace the case analysis rule by a more conventional one:
   892 whether an integer is negative or not.*}
   893 
   894 lemma zless_iff_Suc_zadd:
   895     "(w < z) = (\<exists>n. z = w + int(Suc n))"
   896 apply (cases z, cases w)
   897 apply (auto simp add: le add int_def linorder_not_le [symmetric]) 
   898 apply (rename_tac a b c d) 
   899 apply (rule_tac x="a+d - Suc(c+b)" in exI) 
   900 apply arith
   901 done
   902 
   903 lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
   904 apply (cases x)
   905 apply (auto simp add: le minus Zero_int_def int_def order_less_le) 
   906 apply (rule_tac x="y - Suc x" in exI, arith)
   907 done
   908 
   909 theorem int_cases [cases type: int, case_names nonneg neg]:
   910      "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
   911 apply (case_tac "z < 0", blast dest!: negD)
   912 apply (simp add: linorder_not_less)
   913 apply (blast dest: nat_0_le [THEN sym])
   914 done
   915 
   916 theorem int_induct [induct type: int, case_names nonneg neg]:
   917      "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
   918   by (cases z) auto
   919 
   920 
   921 (*Legacy ML bindings, but no longer the structure Int.*)
   922 ML
   923 {*
   924 val zabs_def = thm "zabs_def"
   925 
   926 val int_0 = thm "int_0";
   927 val int_1 = thm "int_1";
   928 val int_Suc0_eq_1 = thm "int_Suc0_eq_1";
   929 val neg_eq_less_0 = thm "neg_eq_less_0";
   930 val not_neg_eq_ge_0 = thm "not_neg_eq_ge_0";
   931 val not_neg_0 = thm "not_neg_0";
   932 val not_neg_1 = thm "not_neg_1";
   933 val iszero_0 = thm "iszero_0";
   934 val not_iszero_1 = thm "not_iszero_1";
   935 val int_0_less_1 = thm "int_0_less_1";
   936 val int_0_neq_1 = thm "int_0_neq_1";
   937 val negative_zless = thm "negative_zless";
   938 val negative_zle = thm "negative_zle";
   939 val not_zle_0_negative = thm "not_zle_0_negative";
   940 val not_int_zless_negative = thm "not_int_zless_negative";
   941 val negative_eq_positive = thm "negative_eq_positive";
   942 val zle_iff_zadd = thm "zle_iff_zadd";
   943 val abs_int_eq = thm "abs_int_eq";
   944 val abs_split = thm"abs_split";
   945 val nat_int = thm "nat_int";
   946 val nat_zminus_int = thm "nat_zminus_int";
   947 val nat_zero = thm "nat_zero";
   948 val not_neg_nat = thm "not_neg_nat";
   949 val neg_nat = thm "neg_nat";
   950 val zless_nat_eq_int_zless = thm "zless_nat_eq_int_zless";
   951 val nat_0_le = thm "nat_0_le";
   952 val nat_le_0 = thm "nat_le_0";
   953 val zless_nat_conj = thm "zless_nat_conj";
   954 val int_cases = thm "int_cases";
   955 
   956 val int_def = thm "int_def";
   957 val Zero_int_def = thm "Zero_int_def";
   958 val One_int_def = thm "One_int_def";
   959 val diff_int_def = thm "diff_int_def";
   960 
   961 val inj_int = thm "inj_int";
   962 val zminus_zminus = thm "zminus_zminus";
   963 val zminus_0 = thm "zminus_0";
   964 val zminus_zadd_distrib = thm "zminus_zadd_distrib";
   965 val zadd_commute = thm "zadd_commute";
   966 val zadd_assoc = thm "zadd_assoc";
   967 val zadd_left_commute = thm "zadd_left_commute";
   968 val zadd_ac = thms "zadd_ac";
   969 val zmult_ac = thms "zmult_ac";
   970 val zadd_int = thm "zadd_int";
   971 val zadd_int_left = thm "zadd_int_left";
   972 val int_Suc = thm "int_Suc";
   973 val zadd_0 = thm "zadd_0";
   974 val zadd_0_right = thm "zadd_0_right";
   975 val zmult_zminus = thm "zmult_zminus";
   976 val zmult_commute = thm "zmult_commute";
   977 val zmult_assoc = thm "zmult_assoc";
   978 val zadd_zmult_distrib = thm "zadd_zmult_distrib";
   979 val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2";
   980 val zdiff_zmult_distrib = thm "zdiff_zmult_distrib";
   981 val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2";
   982 val int_distrib = thms "int_distrib";
   983 val zmult_int = thm "zmult_int";
   984 val zmult_1 = thm "zmult_1";
   985 val zmult_1_right = thm "zmult_1_right";
   986 val int_int_eq = thm "int_int_eq";
   987 val int_eq_0_conv = thm "int_eq_0_conv";
   988 val zless_int = thm "zless_int";
   989 val int_less_0_conv = thm "int_less_0_conv";
   990 val zero_less_int_conv = thm "zero_less_int_conv";
   991 val zle_int = thm "zle_int";
   992 val zero_zle_int = thm "zero_zle_int";
   993 val int_le_0_conv = thm "int_le_0_conv";
   994 val zle_refl = thm "zle_refl";
   995 val zle_linear = thm "zle_linear";
   996 val zle_trans = thm "zle_trans";
   997 val zle_anti_sym = thm "zle_anti_sym";
   998 
   999 val Ints_def = thm "Ints_def";
  1000 val Nats_def = thm "Nats_def";
  1001 
  1002 val of_nat_0 = thm "of_nat_0";
  1003 val of_nat_Suc = thm "of_nat_Suc";
  1004 val of_nat_1 = thm "of_nat_1";
  1005 val of_nat_add = thm "of_nat_add";
  1006 val of_nat_mult = thm "of_nat_mult";
  1007 val zero_le_imp_of_nat = thm "zero_le_imp_of_nat";
  1008 val less_imp_of_nat_less = thm "less_imp_of_nat_less";
  1009 val of_nat_less_imp_less = thm "of_nat_less_imp_less";
  1010 val of_nat_less_iff = thm "of_nat_less_iff";
  1011 val of_nat_le_iff = thm "of_nat_le_iff";
  1012 val of_nat_eq_iff = thm "of_nat_eq_iff";
  1013 val Nats_0 = thm "Nats_0";
  1014 val Nats_1 = thm "Nats_1";
  1015 val Nats_add = thm "Nats_add";
  1016 val Nats_mult = thm "Nats_mult";
  1017 val int_eq_of_nat = thm"int_eq_of_nat";
  1018 val of_int = thm "of_int";
  1019 val of_int_0 = thm "of_int_0";
  1020 val of_int_1 = thm "of_int_1";
  1021 val of_int_add = thm "of_int_add";
  1022 val of_int_minus = thm "of_int_minus";
  1023 val of_int_diff = thm "of_int_diff";
  1024 val of_int_mult = thm "of_int_mult";
  1025 val of_int_le_iff = thm "of_int_le_iff";
  1026 val of_int_less_iff = thm "of_int_less_iff";
  1027 val of_int_eq_iff = thm "of_int_eq_iff";
  1028 val Ints_0 = thm "Ints_0";
  1029 val Ints_1 = thm "Ints_1";
  1030 val Ints_add = thm "Ints_add";
  1031 val Ints_minus = thm "Ints_minus";
  1032 val Ints_diff = thm "Ints_diff";
  1033 val Ints_mult = thm "Ints_mult";
  1034 val of_int_of_nat_eq = thm"of_int_of_nat_eq";
  1035 val Ints_cases = thm "Ints_cases";
  1036 val Ints_induct = thm "Ints_induct";
  1037 *}
  1038 
  1039 end