two target language numeral types: integer and natural, as replacement for code_numeral;
authorhaftmann
Fri, 15 Feb 2013 08:31:31 +0100
changeset 51143 0a2371e7ced3
parent 51142 ac9e909fe55d
child 51144 0ede9e2266a8
two target language numeral types: integer and natural, as replacement for code_numeral; former theory HOL/Library/Code_Numeral_Types replaces HOL/Code_Numeral; refined stack of theories implementing int and/or nat by target language numerals; reduced number of target language numeral types to exactly one
NEWS
src/Doc/Classes/Classes.thy
src/Doc/Classes/Setup.thy
src/Doc/Codegen/Adaptation.thy
src/Doc/Codegen/Foundations.thy
src/HOL/Code_Evaluation.thy
src/HOL/Code_Numeral.thy
src/HOL/Codegenerator_Test/Candidates_Pretty.thy
src/HOL/Codegenerator_Test/Generate_Pretty.thy
src/HOL/Decision_Procs/Approximation.thy
src/HOL/Decision_Procs/Cooper.thy
src/HOL/Decision_Procs/Ferrack.thy
src/HOL/Decision_Procs/MIR.thy
src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
src/HOL/Imperative_HOL/Array.thy
src/HOL/Imperative_HOL/Heap_Monad.thy
src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy
src/HOL/Imperative_HOL/ex/Linked_Lists.thy
src/HOL/Int.thy
src/HOL/Lazy_Sequence.thy
src/HOL/Library/Cardinality.thy
src/HOL/Library/Code_Binary_Nat.thy
src/HOL/Library/Code_Char_chr.thy
src/HOL/Library/Code_Integer.thy
src/HOL/Library/Code_Natural.thy
src/HOL/Library/Code_Numeral_Types.thy
src/HOL/Library/Code_Real_Approx_By_Float.thy
src/HOL/Library/Code_Target_Int.thy
src/HOL/Library/Code_Target_Nat.thy
src/HOL/Library/DAList.thy
src/HOL/Library/Efficient_Nat.thy
src/HOL/Library/IArray.thy
src/HOL/Library/Multiset.thy
src/HOL/Library/Predicate_Compile_Alternative_Defs.thy
src/HOL/Library/Predicate_Compile_Quickcheck.thy
src/HOL/Limited_Sequence.thy
src/HOL/Num.thy
src/HOL/Predicate.thy
src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy
src/HOL/Proofs/Extraction/Euclid.thy
src/HOL/Proofs/Extraction/Pigeonhole.thy
src/HOL/Proofs/Lambda/WeakNorm.thy
src/HOL/Quickcheck_Benchmark/Needham_Schroeder_Base.thy
src/HOL/Quickcheck_Examples/Completeness.thy
src/HOL/Quickcheck_Examples/Hotel_Example.thy
src/HOL/Quickcheck_Exhaustive.thy
src/HOL/Quickcheck_Narrowing.thy
src/HOL/Quickcheck_Random.thy
src/HOL/ROOT
src/HOL/Random.thy
src/HOL/Random_Pred.thy
src/HOL/Random_Sequence.thy
src/HOL/Rat.thy
src/HOL/Tools/Nitpick/nitpick_hol.ML
src/HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML
src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML
src/HOL/Tools/Qelim/cooper.ML
src/HOL/Tools/Qelim/cooper_procedure.ML
src/HOL/Tools/Quickcheck/exhaustive_generators.ML
src/HOL/Tools/Quickcheck/narrowing_generators.ML
src/HOL/Tools/Quickcheck/quickcheck_common.ML
src/HOL/Tools/Quickcheck/random_generators.ML
src/HOL/Tools/hologic.ML
src/HOL/Tools/record.ML
src/HOL/Word/Word.thy
src/HOL/ex/Code_Binary_Nat_examples.thy
src/HOL/ex/IArray_Examples.thy
src/Tools/Code/code_haskell.ML
src/Tools/Code/code_ml.ML
src/Tools/Code/code_printer.ML
src/Tools/Code/code_scala.ML
--- a/NEWS	Fri Feb 15 08:31:30 2013 +0100
+++ b/NEWS	Fri Feb 15 08:31:31 2013 +0100
@@ -170,6 +170,18 @@
 
 *** HOL ***
 
+* Numeric types mapped by default to target language numerals:
+natural (replaces former code_numeral) and integer (replaces
+former code_int).  Conversions are available as integer_of_natural /
+natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and
+Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in ML).
+INCOMPATIBILITY.
+
+* Discontinued theories Code_Integer and Efficient_Nat by a more
+fine-grain stack of theories Code_Target_Int, Code_Binary_Nat,
+Code_Target_Nat and Code_Target_Numeral.  See the tutorial on
+code generation for details.  INCOMPATIBILITY.
+
 * Sledgehammer:
 
   - Added MaSh relevance filter based on machine-learning; see the
--- a/src/Doc/Classes/Classes.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Doc/Classes/Classes.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -600,7 +600,6 @@
 text {*
   \noindent This maps to Haskell as follows:
 *}
-(*<*)code_include %invisible Haskell "Natural" -(*>*)
 text %quotetypewriter {*
   @{code_stmts example (Haskell)}
 *}
@@ -616,7 +615,6 @@
 text {*
   \noindent In Scala, implicts are used as dictionaries:
 *}
-(*<*)code_include %invisible Scala "Natural" -(*>*)
 text %quotetypewriter {*
   @{code_stmts example (Scala)}
 *}
@@ -640,3 +638,4 @@
 *}
 
 end
+
--- a/src/Doc/Classes/Setup.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Doc/Classes/Setup.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1,5 +1,5 @@
 theory Setup
-imports Main "~~/src/HOL/Library/Code_Integer"
+imports Main
 begin
 
 ML_file "../antiquote_setup.ML"
@@ -37,4 +37,4 @@
   end
 *}
 
-end
\ No newline at end of file
+end
--- a/src/Doc/Codegen/Adaptation.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Doc/Codegen/Adaptation.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -125,11 +125,30 @@
 
   \begin{description}
 
-    \item[@{text "Code_Integer"}] represents @{text HOL} integers by
-       big integer literals in target languages.
+    \item[@{theory "Code_Numeral"}] provides additional numeric
+       types @{typ integer} and @{typ natural} isomorphic to types
+       @{typ int} and @{typ nat} respectively.  Type @{typ integer}
+       is mapped to target-language built-in integers; @{typ natural}
+       is implemented as abstract type over @{typ integer}.
+       Useful for code setups which involve e.g.~indexing
+       of target-language arrays.  Part of @{text "HOL-Main"}.
+
+    \item[@{text "Code_Target_Int"}] implements type @{typ int}
+       by @{typ integer} and thus by target-language built-in integers.
 
-    \item[@{text "Code_Char"}] represents @{text HOL} characters by
-       character literals in target languages.
+    \item[@{text "Code_Binary_Nat"}] \label{eff_nat} implements type
+       @{typ nat} using a binary rather than a linear representation,
+       which yields a considerable speedup for computations.
+       Pattern matching with @{term "0\<Colon>nat"} / @{const "Suc"} is eliminated
+       by a preprocessor.
+
+    \item[@{text "Code_Target_Nat"}] implements type @{typ int}
+       by @{typ integer} and thus by target-language built-in integers;
+       contains @{text "Code_Binary_Nat"} as a prerequisite.
+
+    \item[@{text "Code_Target_Numeral"}] is a convenience node
+       containing both @{text "Code_Target_Nat"} and
+       @{text "Code_Target_Int"}.
 
     \item[@{text "Code_Char_chr"}] like @{text "Code_Char"}, but
        also offers treatment of character codes; includes @{text
@@ -141,10 +160,8 @@
        @{const "Suc"} is eliminated; includes @{text "Code_Integer"}
        and @{text "Code_Numeral"}.
 
-    \item[@{theory "Code_Numeral"}] provides an additional datatype
-       @{typ index} which is mapped to target-language built-in
-       integers.  Useful for code setups which involve e.g.~indexing
-       of target-language arrays.  Part of @{text "HOL-Main"}.
+    \item[@{text "Code_Char"}] represents @{text HOL} characters by
+       character literals in target languages.
 
     \item[@{theory "String"}] provides an additional datatype @{typ
        String.literal} which is isomorphic to strings; @{typ
--- a/src/Doc/Codegen/Foundations.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Doc/Codegen/Foundations.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -117,7 +117,7 @@
   interface, transforming a list of function theorems to another list
   of function theorems, provided that neither the heading constant nor
   its type change.  The @{term "0\<Colon>nat"} / @{const Suc} pattern
-  elimination implemented in theory @{text Efficient_Nat} (see
+  elimination implemented in theory @{text Code_Binary_Nat} (see
   \secref{eff_nat}) uses this interface.
 
   \noindent The current setup of the preprocessor may be inspected
--- a/src/HOL/Code_Evaluation.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Code_Evaluation.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -158,10 +158,16 @@
     else termify (numeral :: num \<Rightarrow> nat) <\<cdot>> term_of_num_semiring (2::nat) n)"
   by (simp only: term_of_anything)
 
-lemma (in term_syntax) term_of_code_numeral_code [code]:
-  "term_of (k::code_numeral) = (
-    if k = 0 then termify (0 :: code_numeral)
-    else termify (numeral :: num \<Rightarrow> code_numeral) <\<cdot>> term_of_num_semiring (2::code_numeral) k)"
+lemma (in term_syntax) term_of_natural_code [code]:
+  "term_of (k::natural) = (
+    if k = 0 then termify (0 :: natural)
+    else termify (numeral :: num \<Rightarrow> natural) <\<cdot>> term_of_num_semiring (2::natural) k)"
+  by (simp only: term_of_anything)
+
+lemma (in term_syntax) term_of_integer_code [code]:
+  "term_of (k::integer) = (if k = 0 then termify (0 :: integer)
+    else if k < 0 then termify (neg_numeral :: num \<Rightarrow> integer) <\<cdot>> term_of_num_semiring (2::integer) (- k)
+    else termify (numeral :: num \<Rightarrow> integer) <\<cdot>> term_of_num_semiring (2::integer) k)"
   by (simp only: term_of_anything)
 
 lemma (in term_syntax) term_of_int_code [code]:
@@ -199,3 +205,4 @@
 hide_const (open) Const App Abs Free termify valtermify term_of term_of_num_semiring tracing
 
 end
+
--- a/src/HOL/Code_Numeral.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Code_Numeral.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1,337 +1,614 @@
-(* Author: Florian Haftmann, TU Muenchen *)
+(*  Title:      HOL/Code_Numeral.thy
+    Author:     Florian Haftmann, TU Muenchen
+*)
 
-header {* Type of target language numerals *}
+header {* Numeric types for code generation onto target language numerals only *}
 
 theory Code_Numeral
-imports Nat_Transfer Divides
+imports Nat_Transfer Divides Lifting
+begin
+
+subsection {* Type of target language integers *}
+
+typedef integer = "UNIV \<Colon> int set"
+  morphisms int_of_integer integer_of_int ..
+
+setup_lifting (no_code) type_definition_integer
+
+lemma integer_eq_iff:
+  "k = l \<longleftrightarrow> int_of_integer k = int_of_integer l"
+  by transfer rule
+
+lemma integer_eqI:
+  "int_of_integer k = int_of_integer l \<Longrightarrow> k = l"
+  using integer_eq_iff [of k l] by simp
+
+lemma int_of_integer_integer_of_int [simp]:
+  "int_of_integer (integer_of_int k) = k"
+  by transfer rule
+
+lemma integer_of_int_int_of_integer [simp]:
+  "integer_of_int (int_of_integer k) = k"
+  by transfer rule
+
+instantiation integer :: ring_1
 begin
 
-text {*
-  Code numerals are isomorphic to HOL @{typ nat} but
-  mapped to target-language builtin numerals.
-*}
+lift_definition zero_integer :: integer
+  is "0 :: int"
+  .
+
+declare zero_integer.rep_eq [simp]
 
-subsection {* Datatype of target language numerals *}
+lift_definition one_integer :: integer
+  is "1 :: int"
+  .
+
+declare one_integer.rep_eq [simp]
 
-typedef code_numeral = "UNIV \<Colon> nat set"
-  morphisms nat_of of_nat ..
+lift_definition plus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+  is "plus :: int \<Rightarrow> int \<Rightarrow> int"
+  .
+
+declare plus_integer.rep_eq [simp]
 
-lemma of_nat_nat_of [simp]:
-  "of_nat (nat_of k) = k"
-  by (rule nat_of_inverse)
+lift_definition uminus_integer :: "integer \<Rightarrow> integer"
+  is "uminus :: int \<Rightarrow> int"
+  .
+
+declare uminus_integer.rep_eq [simp]
 
-lemma nat_of_of_nat [simp]:
-  "nat_of (of_nat n) = n"
-  by (rule of_nat_inverse) (rule UNIV_I)
+lift_definition minus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+  is "minus :: int \<Rightarrow> int \<Rightarrow> int"
+  .
+
+declare minus_integer.rep_eq [simp]
 
-lemma [measure_function]:
-  "is_measure nat_of" by (rule is_measure_trivial)
+lift_definition times_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+  is "times :: int \<Rightarrow> int \<Rightarrow> int"
+  .
+
+declare times_integer.rep_eq [simp]
 
-lemma code_numeral:
-  "(\<And>n\<Colon>code_numeral. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"
-proof
-  fix n :: nat
-  assume "\<And>n\<Colon>code_numeral. PROP P n"
-  then show "PROP P (of_nat n)" .
-next
-  fix n :: code_numeral
-  assume "\<And>n\<Colon>nat. PROP P (of_nat n)"
-  then have "PROP P (of_nat (nat_of n))" .
-  then show "PROP P n" by simp
+instance proof
+qed (transfer, simp add: algebra_simps)+
+
+end
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
+  by (unfold of_nat_def [abs_def])  transfer_prover
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
+proof -
+  have "fun_rel HOL.eq cr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)"
+    by (unfold of_int_of_nat [abs_def]) transfer_prover
+  then show ?thesis by (simp add: id_def)
 qed
 
-lemma code_numeral_case:
-  assumes "\<And>n. k = of_nat n \<Longrightarrow> P"
-  shows P
-  by (rule assms [of "nat_of k"]) simp
-
-lemma code_numeral_induct_raw:
-  assumes "\<And>n. P (of_nat n)"
-  shows "P k"
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer (numeral :: num \<Rightarrow> int) (numeral :: num \<Rightarrow> integer)"
 proof -
-  from assms have "P (of_nat (nat_of k))" .
+  have "fun_rel HOL.eq cr_integer (numeral :: num \<Rightarrow> int) (\<lambda>n. of_int (numeral n))"
+    by transfer_prover
   then show ?thesis by simp
 qed
 
-lemma nat_of_inject [simp]:
-  "nat_of k = nat_of l \<longleftrightarrow> k = l"
-  by (rule nat_of_inject)
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer (neg_numeral :: num \<Rightarrow> int) (neg_numeral :: num \<Rightarrow> integer)"
+  by (unfold neg_numeral_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq (fun_rel HOL.eq cr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold Num.sub_def [abs_def]) transfer_prover
+
+lemma int_of_integer_of_nat [simp]:
+  "int_of_integer (of_nat n) = of_nat n"
+  by transfer rule
+
+lift_definition integer_of_nat :: "nat \<Rightarrow> integer"
+  is "of_nat :: nat \<Rightarrow> int"
+  .
+
+lemma integer_of_nat_eq_of_nat [code]:
+  "integer_of_nat = of_nat"
+  by transfer rule
+
+lemma int_of_integer_integer_of_nat [simp]:
+  "int_of_integer (integer_of_nat n) = of_nat n"
+  by transfer rule
+
+lift_definition nat_of_integer :: "integer \<Rightarrow> nat"
+  is Int.nat
+  .
 
-lemma of_nat_inject [simp]:
-  "of_nat n = of_nat m \<longleftrightarrow> n = m"
-  by (rule of_nat_inject) (rule UNIV_I)+
+lemma nat_of_integer_of_nat [simp]:
+  "nat_of_integer (of_nat n) = n"
+  by transfer simp
+
+lemma int_of_integer_of_int [simp]:
+  "int_of_integer (of_int k) = k"
+  by transfer simp
+
+lemma nat_of_integer_integer_of_nat [simp]:
+  "nat_of_integer (integer_of_nat n) = n"
+  by transfer simp
+
+lemma integer_of_int_eq_of_int [simp, code_abbrev]:
+  "integer_of_int = of_int"
+  by transfer (simp add: fun_eq_iff)
 
-instantiation code_numeral :: zero
+lemma of_int_integer_of [simp]:
+  "of_int (int_of_integer k) = (k :: integer)"
+  by transfer rule
+
+lemma int_of_integer_numeral [simp]:
+  "int_of_integer (numeral k) = numeral k"
+  by transfer rule
+
+lemma int_of_integer_neg_numeral [simp]:
+  "int_of_integer (neg_numeral k) = neg_numeral k"
+  by transfer rule
+
+lemma int_of_integer_sub [simp]:
+  "int_of_integer (Num.sub k l) = Num.sub k l"
+  by transfer rule
+
+instantiation integer :: "{ring_div, equal, linordered_idom}"
 begin
 
-definition [simp, code del]:
-  "0 = of_nat 0"
+lift_definition div_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+  is "Divides.div :: int \<Rightarrow> int \<Rightarrow> int"
+  .
+
+declare div_integer.rep_eq [simp]
+
+lift_definition mod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
+  is "Divides.mod :: int \<Rightarrow> int \<Rightarrow> int"
+  .
+
+declare mod_integer.rep_eq [simp]
+
+lift_definition abs_integer :: "integer \<Rightarrow> integer"
+  is "abs :: int \<Rightarrow> int"
+  .
+
+declare abs_integer.rep_eq [simp]
 
-instance ..
+lift_definition sgn_integer :: "integer \<Rightarrow> integer"
+  is "sgn :: int \<Rightarrow> int"
+  .
+
+declare sgn_integer.rep_eq [simp]
+
+lift_definition less_eq_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
+  is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool"
+  .
+
+lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
+  is "less :: int \<Rightarrow> int \<Rightarrow> bool"
+  .
+
+lift_definition equal_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
+  is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool"
+  .
+
+instance proof
+qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
 
 end
 
-definition Suc where [simp]:
-  "Suc k = of_nat (Nat.Suc (nat_of k))"
+lemma [transfer_rule]:
+  "fun_rel cr_integer (fun_rel cr_integer cr_integer) (min :: _ \<Rightarrow> _ \<Rightarrow> int) (min :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold min_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "fun_rel cr_integer (fun_rel cr_integer cr_integer) (max :: _ \<Rightarrow> _ \<Rightarrow> int) (max :: _ \<Rightarrow> _ \<Rightarrow> integer)"
+  by (unfold max_def [abs_def]) transfer_prover
+
+lemma int_of_integer_min [simp]:
+  "int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)"
+  by transfer rule
+
+lemma int_of_integer_max [simp]:
+  "int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)"
+  by transfer rule
 
-rep_datatype "0 \<Colon> code_numeral" Suc
-proof -
-  fix P :: "code_numeral \<Rightarrow> bool"
-  fix k :: code_numeral
-  assume "P 0" then have init: "P (of_nat 0)" by simp
-  assume "\<And>k. P k \<Longrightarrow> P (Suc k)"
-    then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc (of_nat n))" .
-    then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Nat.Suc n))" by simp
-  from init step have "P (of_nat (nat_of k))"
-    by (induct ("nat_of k")) simp_all
-  then show "P k" by simp
-qed simp_all
+lemma nat_of_integer_non_positive [simp]:
+  "k \<le> 0 \<Longrightarrow> nat_of_integer k = 0"
+  by transfer simp
+
+lemma of_nat_of_integer [simp]:
+  "of_nat (nat_of_integer k) = max 0 k"
+  by transfer auto
+
 
-declare code_numeral_case [case_names nat, cases type: code_numeral]
-declare code_numeral.induct [case_names nat, induct type: code_numeral]
+subsection {* Code theorems for target language integers *}
+
+text {* Constructors *}
 
-lemma code_numeral_decr [termination_simp]:
-  "k \<noteq> of_nat 0 \<Longrightarrow> nat_of k - Nat.Suc 0 < nat_of k"
-  by (cases k) simp
+definition Pos :: "num \<Rightarrow> integer"
+where
+  [simp, code_abbrev]: "Pos = numeral"
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer numeral Pos"
+  by simp transfer_prover
 
-lemma [simp, code]:
-  "code_numeral_size = nat_of"
-proof (rule ext)
-  fix k
-  have "code_numeral_size k = nat_size (nat_of k)"
-    by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
-  also have "nat_size (nat_of k) = nat_of k" by (induct ("nat_of k")) simp_all
-  finally show "code_numeral_size k = nat_of k" .
-qed
+definition Neg :: "num \<Rightarrow> integer"
+where
+  [simp, code_abbrev]: "Neg = neg_numeral"
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_integer neg_numeral Neg"
+  by simp transfer_prover
+
+code_datatype "0::integer" Pos Neg
+
+
+text {* Auxiliary operations *}
+
+lift_definition dup :: "integer \<Rightarrow> integer"
+  is "\<lambda>k::int. k + k"
+  .
 
-lemma [simp, code]:
-  "size = nat_of"
-proof (rule ext)
-  fix k
-  show "size k = nat_of k"
-  by (induct k) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
-qed
+lemma dup_code [code]:
+  "dup 0 = 0"
+  "dup (Pos n) = Pos (Num.Bit0 n)"
+  "dup (Neg n) = Neg (Num.Bit0 n)"
+  by (transfer, simp only: neg_numeral_def numeral_Bit0 minus_add_distrib)+
+
+lift_definition sub :: "num \<Rightarrow> num \<Rightarrow> integer"
+  is "\<lambda>m n. numeral m - numeral n :: int"
+  .
 
-lemmas [code del] = code_numeral.recs code_numeral.cases
-
-lemma [code]:
-  "HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)"
-  by (cases k, cases l) (simp add: equal)
-
-lemma [code nbe]:
-  "HOL.equal (k::code_numeral) k \<longleftrightarrow> True"
-  by (rule equal_refl)
+lemma sub_code [code]:
+  "sub Num.One Num.One = 0"
+  "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"
+  "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"
+  "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"
+  "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"
+  "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)"
+  "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
+  "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
+  "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
+  by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+
 
 
-subsection {* Basic arithmetic *}
+text {* Implementations *}
 
-instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}"
-begin
-
-definition [simp, code del]:
-  "(1\<Colon>code_numeral) = of_nat 1"
+lemma one_integer_code [code, code_unfold]:
+  "1 = Pos Num.One"
+  by simp
 
-definition [simp, code del]:
-  "n + m = of_nat (nat_of n + nat_of m)"
-
-definition [simp, code del]:
-  "n - m = of_nat (nat_of n - nat_of m)"
-
-definition [simp, code del]:
-  "n * m = of_nat (nat_of n * nat_of m)"
-
-definition [simp, code del]:
-  "n div m = of_nat (nat_of n div nat_of m)"
+lemma plus_integer_code [code]:
+  "k + 0 = (k::integer)"
+  "0 + l = (l::integer)"
+  "Pos m + Pos n = Pos (m + n)"
+  "Pos m + Neg n = sub m n"
+  "Neg m + Pos n = sub n m"
+  "Neg m + Neg n = Neg (m + n)"
+  by (transfer, simp)+
 
-definition [simp, code del]:
-  "n mod m = of_nat (nat_of n mod nat_of m)"
-
-definition [simp, code del]:
-  "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"
-
-definition [simp, code del]:
-  "n < m \<longleftrightarrow> nat_of n < nat_of m"
-
-instance proof
-qed (auto simp add: code_numeral distrib_right intro: mult_commute)
+lemma uminus_integer_code [code]:
+  "uminus 0 = (0::integer)"
+  "uminus (Pos m) = Neg m"
+  "uminus (Neg m) = Pos m"
+  by simp_all
 
-end
-
-lemma nat_of_numeral [simp]: "nat_of (numeral k) = numeral k"
-  by (induct k rule: num_induct) (simp_all add: numeral_inc)
+lemma minus_integer_code [code]:
+  "k - 0 = (k::integer)"
+  "0 - l = uminus (l::integer)"
+  "Pos m - Pos n = sub m n"
+  "Pos m - Neg n = Pos (m + n)"
+  "Neg m - Pos n = Neg (m + n)"
+  "Neg m - Neg n = sub n m"
+  by (transfer, simp)+
 
-definition Num :: "num \<Rightarrow> code_numeral"
-  where [simp, code_abbrev]: "Num = numeral"
+lemma abs_integer_code [code]:
+  "\<bar>k\<bar> = (if (k::integer) < 0 then - k else k)"
+  by simp
 
-code_datatype "0::code_numeral" Num
-
-lemma one_code_numeral_code [code]:
-  "(1\<Colon>code_numeral) = Numeral1"
+lemma sgn_integer_code [code]:
+  "sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)"
   by simp
 
-lemma [code_abbrev]: "Numeral1 = (1\<Colon>code_numeral)"
-  using one_code_numeral_code ..
+lemma times_integer_code [code]:
+  "k * 0 = (0::integer)"
+  "0 * l = (0::integer)"
+  "Pos m * Pos n = Pos (m * n)"
+  "Pos m * Neg n = Neg (m * n)"
+  "Neg m * Pos n = Neg (m * n)"
+  "Neg m * Neg n = Pos (m * n)"
+  by simp_all
+
+definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
+where
+  "divmod_integer k l = (k div l, k mod l)"
+
+lemma fst_divmod [simp]:
+  "fst (divmod_integer k l) = k div l"
+  by (simp add: divmod_integer_def)
+
+lemma snd_divmod [simp]:
+  "snd (divmod_integer k l) = k mod l"
+  by (simp add: divmod_integer_def)
+
+definition divmod_abs :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
+where
+  "divmod_abs k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
+
+lemma fst_divmod_abs [simp]:
+  "fst (divmod_abs k l) = \<bar>k\<bar> div \<bar>l\<bar>"
+  by (simp add: divmod_abs_def)
+
+lemma snd_divmod_abs [simp]:
+  "snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"
+  by (simp add: divmod_abs_def)
 
-lemma plus_code_numeral_code [code nbe]:
-  "of_nat n + of_nat m = of_nat (n + m)"
-  by simp
+lemma divmod_abs_terminate_code [code]:
+  "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)"
+  "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)"
+  "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)"
+  "divmod_abs j 0 = (0, \<bar>j\<bar>)"
+  "divmod_abs 0 j = (0, 0)"
+  by (simp_all add: prod_eq_iff)
+
+lemma divmod_abs_rec_code [code]:
+  "divmod_abs (Pos k) (Pos l) =
+    (let j = sub k l in
+       if j < 0 then (0, Pos k)
+       else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))"
+  apply (simp add: prod_eq_iff Let_def prod_case_beta)
+  apply transfer
+  apply (simp add: sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq)
+  done
 
-lemma minus_code_numeral_code [code nbe]:
-  "of_nat n - of_nat m = of_nat (n - m)"
+lemma divmod_integer_code [code]:
+  "divmod_integer k l =
+    (if k = 0 then (0, 0) else if l = 0 then (0, k) else
+    (apsnd \<circ> times \<circ> sgn) l (if sgn k = sgn l
+      then divmod_abs k l
+      else (let (r, s) = divmod_abs k l in
+        if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
+proof -
+  have aux1: "\<And>k l::int. sgn k = sgn l \<longleftrightarrow> k = 0 \<and> l = 0 \<or> 0 < l \<and> 0 < k \<or> l < 0 \<and> k < 0"
+    by (auto simp add: sgn_if)
+  have aux2: "\<And>q::int. - int_of_integer k = int_of_integer l * q \<longleftrightarrow> int_of_integer k = int_of_integer l * - q" by auto
+  show ?thesis
+    by (simp add: prod_eq_iff integer_eq_iff prod_case_beta aux1)
+      (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2)
+qed
+
+lemma div_integer_code [code]:
+  "k div l = fst (divmod_integer k l)"
   by simp
 
-lemma times_code_numeral_code [code nbe]:
-  "of_nat n * of_nat m = of_nat (n * m)"
-  by simp
-
-lemma less_eq_code_numeral_code [code nbe]:
-  "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
-  by simp
-
-lemma less_code_numeral_code [code nbe]:
-  "of_nat n < of_nat m \<longleftrightarrow> n < m"
+lemma mod_integer_code [code]:
+  "k mod l = snd (divmod_integer k l)"
   by simp
 
-lemma code_numeral_zero_minus_one:
-  "(0::code_numeral) - 1 = 0"
-  by simp
+lemma equal_integer_code [code]:
+  "HOL.equal 0 (0::integer) \<longleftrightarrow> True"
+  "HOL.equal 0 (Pos l) \<longleftrightarrow> False"
+  "HOL.equal 0 (Neg l) \<longleftrightarrow> False"
+  "HOL.equal (Pos k) 0 \<longleftrightarrow> False"
+  "HOL.equal (Pos k) (Pos l) \<longleftrightarrow> HOL.equal k l"
+  "HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"
+  "HOL.equal (Neg k) 0 \<longleftrightarrow> False"
+  "HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"
+  "HOL.equal (Neg k) (Neg l) \<longleftrightarrow> HOL.equal k l"
+  by (simp_all add: equal)
+
+lemma equal_integer_refl [code nbe]:
+  "HOL.equal (k::integer) k \<longleftrightarrow> True"
+  by (fact equal_refl)
 
-lemma Suc_code_numeral_minus_one:
-  "Suc n - 1 = n"
-  by simp
+lemma less_eq_integer_code [code]:
+  "0 \<le> (0::integer) \<longleftrightarrow> True"
+  "0 \<le> Pos l \<longleftrightarrow> True"
+  "0 \<le> Neg l \<longleftrightarrow> False"
+  "Pos k \<le> 0 \<longleftrightarrow> False"
+  "Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"
+  "Pos k \<le> Neg l \<longleftrightarrow> False"
+  "Neg k \<le> 0 \<longleftrightarrow> True"
+  "Neg k \<le> Pos l \<longleftrightarrow> True"
+  "Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"
+  by simp_all
+
+lemma less_integer_code [code]:
+  "0 < (0::integer) \<longleftrightarrow> False"
+  "0 < Pos l \<longleftrightarrow> True"
+  "0 < Neg l \<longleftrightarrow> False"
+  "Pos k < 0 \<longleftrightarrow> False"
+  "Pos k < Pos l \<longleftrightarrow> k < l"
+  "Pos k < Neg l \<longleftrightarrow> False"
+  "Neg k < 0 \<longleftrightarrow> True"
+  "Neg k < Pos l \<longleftrightarrow> True"
+  "Neg k < Neg l \<longleftrightarrow> l < k"
+  by simp_all
 
-lemma of_nat_code [code]:
-  "of_nat = Nat.of_nat"
-proof
-  fix n :: nat
-  have "Nat.of_nat n = of_nat n"
-    by (induct n) simp_all
-  then show "of_nat n = Nat.of_nat n"
-    by (rule sym)
+lift_definition integer_of_num :: "num \<Rightarrow> integer"
+  is "numeral :: num \<Rightarrow> int"
+  .
+
+lemma integer_of_num [code]:
+  "integer_of_num num.One = 1"
+  "integer_of_num (num.Bit0 n) = (let k = integer_of_num n in k + k)"
+  "integer_of_num (num.Bit1 n) = (let k = integer_of_num n in k + k + 1)"
+  by (transfer, simp only: numeral.simps Let_def)+
+
+lift_definition num_of_integer :: "integer \<Rightarrow> num"
+  is "num_of_nat \<circ> nat"
+  .
+
+lemma num_of_integer_code [code]:
+  "num_of_integer k = (if k \<le> 1 then Num.One
+     else let
+       (l, j) = divmod_integer k 2;
+       l' = num_of_integer l;
+       l'' = l' + l'
+     in if j = 0 then l'' else l'' + Num.One)"
+proof -
+  {
+    assume "int_of_integer k mod 2 = 1"
+    then have "nat (int_of_integer k mod 2) = nat 1" by simp
+    moreover assume *: "1 < int_of_integer k"
+    ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib)
+    have "num_of_nat (nat (int_of_integer k)) =
+      num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)"
+      by simp
+    then have "num_of_nat (nat (int_of_integer k)) =
+      num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)"
+      by (simp add: mult_2)
+    with ** have "num_of_nat (nat (int_of_integer k)) =
+      num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)"
+      by simp
+  }
+  note aux = this
+  show ?thesis
+    by (auto simp add: num_of_integer_def nat_of_integer_def Let_def prod_case_beta
+      not_le integer_eq_iff less_eq_integer_def
+      nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib
+       mult_2 [where 'a=nat] aux add_One)
 qed
 
-lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
-  by (cases i) auto
-
-definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where
-  "nat_of_aux i n = nat_of i + n"
-
-lemma nat_of_aux_code [code]:
-  "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Nat.Suc n))"
-  by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero)
-
-lemma nat_of_code [code]:
-  "nat_of i = nat_of_aux i 0"
-  by (simp add: nat_of_aux_def)
-
-definition div_mod :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
-  [code del]: "div_mod n m = (n div m, n mod m)"
-
-lemma [code]:
-  "div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))"
-  unfolding div_mod_def by auto
-
-lemma [code]:
-  "n div m = fst (div_mod n m)"
-  unfolding div_mod_def by simp
-
-lemma [code]:
-  "n mod m = snd (div_mod n m)"
-  unfolding div_mod_def by simp
-
-definition int_of :: "code_numeral \<Rightarrow> int" where
-  "int_of = Nat.of_nat o nat_of"
-
-lemma int_of_code [code]:
-  "int_of k = (if k = 0 then 0
-    else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
+lemma nat_of_integer_code [code]:
+  "nat_of_integer k = (if k \<le> 0 then 0
+     else let
+       (l, j) = divmod_integer k 2;
+       l' = nat_of_integer l;
+       l'' = l' + l'
+     in if j = 0 then l'' else l'' + 1)"
 proof -
-  have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k" 
-    by (rule mod_div_equality)
-  then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)" 
-    by simp
-  then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" 
-    unfolding of_nat_mult of_nat_add by simp
-  then show ?thesis by (auto simp add: int_of_def mult_ac)
+  obtain j where "k = integer_of_int j"
+  proof
+    show "k = integer_of_int (int_of_integer k)" by simp
+  qed
+  moreover have "2 * (j div 2) = j - j mod 2"
+    by (simp add: zmult_div_cancel mult_commute)
+  ultimately show ?thesis
+    by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le
+      nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1)
+      (auto simp add: mult_2 [symmetric])
 qed
 
+lemma int_of_integer_code [code]:
+  "int_of_integer k = (if k < 0 then - (int_of_integer (- k))
+     else if k = 0 then 0
+     else let
+       (l, j) = divmod_integer k 2;
+       l' = 2 * int_of_integer l
+     in if j = 0 then l' else l' + 1)"
+  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
 
-hide_const (open) of_nat nat_of Suc int_of
+lemma integer_of_int_code [code]:
+  "integer_of_int k = (if k < 0 then - (integer_of_int (- k))
+     else if k = 0 then 0
+     else let
+       (l, j) = divmod_int k 2;
+       l' = 2 * integer_of_int l
+     in if j = 0 then l' else l' + 1)"
+  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
+
+hide_const (open) Pos Neg sub dup divmod_abs
 
 
-subsection {* Code generator setup *}
+subsection {* Serializer setup for target language integers *}
 
-text {* Implementation of code numerals by bounded integers *}
+code_reserved Eval int Integer abs
 
-code_type code_numeral
-  (SML "int")
+code_type integer
+  (SML "IntInf.int")
   (OCaml "Big'_int.big'_int")
   (Haskell "Integer")
   (Scala "BigInt")
+  (Eval "int")
 
-code_instance code_numeral :: equal
+code_instance integer :: equal
   (Haskell -)
 
-setup {*
-  Numeral.add_code @{const_name Num}
-    false Code_Printer.literal_naive_numeral "SML"
-  #> fold (Numeral.add_code @{const_name Num}
-    false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"]
-*}
-
-code_reserved SML Int int
-code_reserved Eval Integer
-
-code_const "0::code_numeral"
+code_const "0::integer"
   (SML "0")
   (OCaml "Big'_int.zero'_big'_int")
   (Haskell "0")
   (Scala "BigInt(0)")
 
-code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (SML "Int.+/ ((_),/ (_))")
+setup {*
+  fold (Numeral.add_code @{const_name Code_Numeral.Pos}
+    false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
+*}
+
+setup {*
+  fold (Numeral.add_code @{const_name Code_Numeral.Neg}
+    true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
+*}
+
+code_const "plus :: integer \<Rightarrow> _ \<Rightarrow> _"
+  (SML "IntInf.+ ((_), (_))")
   (OCaml "Big'_int.add'_big'_int")
   (Haskell infixl 6 "+")
   (Scala infixl 7 "+")
   (Eval infixl 8 "+")
 
-code_const "minus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (SML "Int.max/ (0 : int,/ Int.-/ ((_),/ (_)))")
-  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)")
-  (Haskell "Prelude.max/ (0 :: Integer)/ (_/ -/ _)")
-  (Scala "!(_/ -/ _).max(0)")
-  (Eval "Integer.max/ 0/ (_/ -/ _)")
+code_const "uminus :: integer \<Rightarrow> _"
+  (SML "IntInf.~")
+  (OCaml "Big'_int.minus'_big'_int")
+  (Haskell "negate")
+  (Scala "!(- _)")
+  (Eval "~/ _")
+
+code_const "minus :: integer \<Rightarrow> _"
+  (SML "IntInf.- ((_), (_))")
+  (OCaml "Big'_int.sub'_big'_int")
+  (Haskell infixl 6 "-")
+  (Scala infixl 7 "-")
+  (Eval infixl 8 "-")
 
-code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (SML "Int.*/ ((_),/ (_))")
+code_const Code_Numeral.dup
+  (SML "IntInf.*/ (2,/ (_))")
+  (OCaml "Big'_int.mult'_big'_int/ (Big'_int.big'_int'_of'_int/ 2)")
+  (Haskell "!(2 * _)")
+  (Scala "!(2 * _)")
+  (Eval "!(2 * _)")
+
+code_const Code_Numeral.sub
+  (SML "!(raise/ Fail/ \"sub\")")
+  (OCaml "failwith/ \"sub\"")
+  (Haskell "error/ \"sub\"")
+  (Scala "!sys.error(\"sub\")")
+
+code_const "times :: integer \<Rightarrow> _ \<Rightarrow> _"
+  (SML "IntInf.* ((_), (_))")
   (OCaml "Big'_int.mult'_big'_int")
   (Haskell infixl 7 "*")
   (Scala infixl 8 "*")
-  (Eval infixl 8 "*")
+  (Eval infixl 9 "*")
 
-code_const Code_Numeral.div_mod
-  (SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Int.div (n, m), Int.mod (n, m)))")
+code_const Code_Numeral.divmod_abs
+  (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
   (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
-  (Haskell "divMod")
+  (Haskell "divMod/ (abs _)/ (abs _)")
   (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
-  (Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Integer.div'_mod n m))")
+  (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
 
-code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (SML "!((_ : Int.int) = _)")
+code_const "HOL.equal :: integer \<Rightarrow> _ \<Rightarrow> bool"
+  (SML "!((_ : IntInf.int) = _)")
   (OCaml "Big'_int.eq'_big'_int")
   (Haskell infix 4 "==")
   (Scala infixl 5 "==")
-  (Eval "!((_ : int) = _)")
+  (Eval infixl 6 "=")
 
-code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (SML "Int.<=/ ((_),/ (_))")
+code_const "less_eq :: integer \<Rightarrow> _ \<Rightarrow> bool"
+  (SML "IntInf.<= ((_), (_))")
   (OCaml "Big'_int.le'_big'_int")
   (Haskell infix 4 "<=")
   (Scala infixl 4 "<=")
   (Eval infixl 6 "<=")
 
-code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (SML "Int.</ ((_),/ (_))")
+code_const "less :: integer \<Rightarrow> _ \<Rightarrow> bool"
+  (SML "IntInf.< ((_), (_))")
   (OCaml "Big'_int.lt'_big'_int")
   (Haskell infix 4 "<")
   (Scala infixl 4 "<")
@@ -346,5 +623,321 @@
 code_modulename Haskell
   Code_Numeral Arith
 
+
+subsection {* Type of target language naturals *}
+
+typedef natural = "UNIV \<Colon> nat set"
+  morphisms nat_of_natural natural_of_nat ..
+
+setup_lifting (no_code) type_definition_natural
+
+lemma natural_eq_iff [termination_simp]:
+  "m = n \<longleftrightarrow> nat_of_natural m = nat_of_natural n"
+  by transfer rule
+
+lemma natural_eqI:
+  "nat_of_natural m = nat_of_natural n \<Longrightarrow> m = n"
+  using natural_eq_iff [of m n] by simp
+
+lemma nat_of_natural_of_nat_inverse [simp]:
+  "nat_of_natural (natural_of_nat n) = n"
+  by transfer rule
+
+lemma natural_of_nat_of_natural_inverse [simp]:
+  "natural_of_nat (nat_of_natural n) = n"
+  by transfer rule
+
+instantiation natural :: "{comm_monoid_diff, semiring_1}"
+begin
+
+lift_definition zero_natural :: natural
+  is "0 :: nat"
+  .
+
+declare zero_natural.rep_eq [simp]
+
+lift_definition one_natural :: natural
+  is "1 :: nat"
+  .
+
+declare one_natural.rep_eq [simp]
+
+lift_definition plus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+  is "plus :: nat \<Rightarrow> nat \<Rightarrow> nat"
+  .
+
+declare plus_natural.rep_eq [simp]
+
+lift_definition minus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+  is "minus :: nat \<Rightarrow> nat \<Rightarrow> nat"
+  .
+
+declare minus_natural.rep_eq [simp]
+
+lift_definition times_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+  is "times :: nat \<Rightarrow> nat \<Rightarrow> nat"
+  .
+
+declare times_natural.rep_eq [simp]
+
+instance proof
+qed (transfer, simp add: algebra_simps)+
+
+end
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)"
+proof -
+  have "fun_rel HOL.eq cr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"
+    by (unfold of_nat_def [abs_def]) transfer_prover
+  then show ?thesis by (simp add: id_def)
+qed
+
+lemma [transfer_rule]:
+  "fun_rel HOL.eq cr_natural (numeral :: num \<Rightarrow> nat) (numeral :: num \<Rightarrow> natural)"
+proof -
+  have "fun_rel HOL.eq cr_natural (numeral :: num \<Rightarrow> nat) (\<lambda>n. of_nat (numeral n))"
+    by transfer_prover
+  then show ?thesis by simp
+qed
+
+lemma nat_of_natural_of_nat [simp]:
+  "nat_of_natural (of_nat n) = n"
+  by transfer rule
+
+lemma natural_of_nat_of_nat [simp, code_abbrev]:
+  "natural_of_nat = of_nat"
+  by transfer rule
+
+lemma of_nat_of_natural [simp]:
+  "of_nat (nat_of_natural n) = n"
+  by transfer rule
+
+lemma nat_of_natural_numeral [simp]:
+  "nat_of_natural (numeral k) = numeral k"
+  by transfer rule
+
+instantiation natural :: "{semiring_div, equal, linordered_semiring}"
+begin
+
+lift_definition div_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+  is "Divides.div :: nat \<Rightarrow> nat \<Rightarrow> nat"
+  .
+
+declare div_natural.rep_eq [simp]
+
+lift_definition mod_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
+  is "Divides.mod :: nat \<Rightarrow> nat \<Rightarrow> nat"
+  .
+
+declare mod_natural.rep_eq [simp]
+
+lift_definition less_eq_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
+  is "less_eq :: nat \<Rightarrow> nat \<Rightarrow> bool"
+  .
+
+declare less_eq_natural.rep_eq [termination_simp]
+
+lift_definition less_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
+  is "less :: nat \<Rightarrow> nat \<Rightarrow> bool"
+  .
+
+declare less_natural.rep_eq [termination_simp]
+
+lift_definition equal_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
+  is "HOL.equal :: nat \<Rightarrow> nat \<Rightarrow> bool"
+  .
+
+instance proof
+qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] linear)+
+
 end
 
+lemma [transfer_rule]:
+  "fun_rel cr_natural (fun_rel cr_natural cr_natural) (min :: _ \<Rightarrow> _ \<Rightarrow> nat) (min :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold min_def [abs_def]) transfer_prover
+
+lemma [transfer_rule]:
+  "fun_rel cr_natural (fun_rel cr_natural cr_natural) (max :: _ \<Rightarrow> _ \<Rightarrow> nat) (max :: _ \<Rightarrow> _ \<Rightarrow> natural)"
+  by (unfold max_def [abs_def]) transfer_prover
+
+lemma nat_of_natural_min [simp]:
+  "nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)"
+  by transfer rule
+
+lemma nat_of_natural_max [simp]:
+  "nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"
+  by transfer rule
+
+lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
+  is "nat :: int \<Rightarrow> nat"
+  .
+
+lift_definition integer_of_natural :: "natural \<Rightarrow> integer"
+  is "of_nat :: nat \<Rightarrow> int"
+  .
+
+lemma natural_of_integer_of_natural [simp]:
+  "natural_of_integer (integer_of_natural n) = n"
+  by transfer simp
+
+lemma integer_of_natural_of_integer [simp]:
+  "integer_of_natural (natural_of_integer k) = max 0 k"
+  by transfer auto
+
+lemma int_of_integer_of_natural [simp]:
+  "int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)"
+  by transfer rule
+
+lemma integer_of_natural_of_nat [simp]:
+  "integer_of_natural (of_nat n) = of_nat n"
+  by transfer rule
+
+lemma [measure_function]:
+  "is_measure nat_of_natural"
+  by (rule is_measure_trivial)
+
+
+subsection {* Inductive represenation of target language naturals *}
+
+lift_definition Suc :: "natural \<Rightarrow> natural"
+  is Nat.Suc
+  .
+
+declare Suc.rep_eq [simp]
+
+rep_datatype "0::natural" Suc
+  by (transfer, fact nat.induct nat.inject nat.distinct)+
+
+lemma natural_case [case_names nat, cases type: natural]:
+  fixes m :: natural
+  assumes "\<And>n. m = of_nat n \<Longrightarrow> P"
+  shows P
+  using assms by transfer blast
+
+lemma [simp, code]:
+  "natural_size = nat_of_natural"
+proof (rule ext)
+  fix n
+  show "natural_size n = nat_of_natural n"
+    by (induct n) simp_all
+qed
+
+lemma [simp, code]:
+  "size = nat_of_natural"
+proof (rule ext)
+  fix n
+  show "size n = nat_of_natural n"
+    by (induct n) simp_all
+qed
+
+lemma natural_decr [termination_simp]:
+  "n \<noteq> 0 \<Longrightarrow> nat_of_natural n - Nat.Suc 0 < nat_of_natural n"
+  by transfer simp
+
+lemma natural_zero_minus_one:
+  "(0::natural) - 1 = 0"
+  by simp
+
+lemma Suc_natural_minus_one:
+  "Suc n - 1 = n"
+  by transfer simp
+
+hide_const (open) Suc
+
+
+subsection {* Code refinement for target language naturals *}
+
+lift_definition Nat :: "integer \<Rightarrow> natural"
+  is nat
+  .
+
+lemma [code_post]:
+  "Nat 0 = 0"
+  "Nat 1 = 1"
+  "Nat (numeral k) = numeral k"
+  by (transfer, simp)+
+
+lemma [code abstype]:
+  "Nat (integer_of_natural n) = n"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural (natural_of_nat n) = of_nat n"
+  by simp
+
+lemma [code abstract]:
+  "integer_of_natural (natural_of_integer k) = max 0 k"
+  by simp
+
+lemma [code_abbrev]:
+  "natural_of_integer (Code_Numeral.Pos k) = numeral k"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural 0 = 0"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural 1 = 1"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural (Code_Numeral.Suc n) = integer_of_natural n + 1"
+  by transfer simp
+
+lemma [code]:
+  "nat_of_natural = nat_of_integer \<circ> integer_of_natural"
+  by transfer (simp add: fun_eq_iff)
+
+lemma [code, code_unfold]:
+  "natural_case f g n = (if n = 0 then f else g (n - 1))"
+  by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def)
+
+declare natural.recs [code del]
+
+lemma [code abstract]:
+  "integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)"
+  by transfer simp
+
+lemma [code abstract]:
+  "integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"
+  by transfer (simp add: of_nat_mult)
+
+lemma [code abstract]:
+  "integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"
+  by transfer (simp add: zdiv_int)
+
+lemma [code abstract]:
+  "integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n"
+  by transfer (simp add: zmod_int)
+
+lemma [code]:
+  "HOL.equal m n \<longleftrightarrow> HOL.equal (integer_of_natural m) (integer_of_natural n)"
+  by transfer (simp add: equal)
+
+lemma [code nbe]:
+  "HOL.equal n (n::natural) \<longleftrightarrow> True"
+  by (simp add: equal)
+
+lemma [code]:
+  "m \<le> n \<longleftrightarrow> integer_of_natural m \<le> integer_of_natural n"
+  by transfer simp
+
+lemma [code]:
+  "m < n \<longleftrightarrow> integer_of_natural m < integer_of_natural n"
+  by transfer simp
+
+hide_const (open) Nat
+
+
+code_reflect Code_Numeral
+  datatypes natural = _
+  functions integer_of_natural natural_of_integer
+
+end
+
--- a/src/HOL/Codegenerator_Test/Candidates_Pretty.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Codegenerator_Test/Candidates_Pretty.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -4,7 +4,7 @@
 header {* Generating code using pretty literals and natural number literals  *}
 
 theory Candidates_Pretty
-imports Candidates Code_Char_ord Efficient_Nat
+imports Candidates Code_Char_ord Code_Target_Numeral
 begin
 
 end
--- a/src/HOL/Codegenerator_Test/Generate_Pretty.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Codegenerator_Test/Generate_Pretty.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -10,9 +10,6 @@
 lemma [code, code del]: "nat_of_char = nat_of_char" ..
 lemma [code, code del]: "char_of_nat = char_of_nat" ..
 
-declare Quickcheck_Narrowing.one_code_int_code [code del]
-declare Quickcheck_Narrowing.int_of_code [code del]
-
 subsection {* Check whether generated code compiles *}
 
 text {*
--- a/src/HOL/Decision_Procs/Approximation.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -9,7 +9,7 @@
   "~~/src/HOL/Library/Float"
   "~~/src/HOL/Library/Reflection"
   "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
-  "~~/src/HOL/Library/Efficient_Nat"
+  "~~/src/HOL/Library/Code_Target_Numeral"
 begin
 
 declare powr_numeral[simp]
@@ -3329,8 +3329,11 @@
   fun term_of_bool true = @{term True}
     | term_of_bool false = @{term False};
 
+  val mk_int = HOLogic.mk_number @{typ int} o @{code integer_of_int};
+  val dest_int = @{code int_of_integer} o snd o HOLogic.dest_number;
+
   fun term_of_float (@{code Float} (k, l)) =
-    @{term Float} $ HOLogic.mk_number @{typ int} k $ HOLogic.mk_number @{typ int} l;
+    @{term Float} $ mk_int k $ mk_int l;
 
   fun term_of_float_float_option NONE = @{term "None :: (float \<times> float) option"}
     | term_of_float_float_option (SOME ff) = @{term "Some :: float \<times> float \<Rightarrow> _"}
@@ -3339,10 +3342,11 @@
   val term_of_float_float_option_list =
     HOLogic.mk_list @{typ "(float \<times> float) option"} o map term_of_float_float_option;
 
-  fun nat_of_term t = HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t);
+  fun nat_of_term t = @{code nat_of_integer}
+    (HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t));
 
   fun float_of_term (@{term Float} $ k $ l) =
-        @{code Float} (snd (HOLogic.dest_number k), snd (HOLogic.dest_number l))
+        @{code Float} (dest_int k, dest_int l)
     | float_of_term t = bad t;
 
   fun floatarith_of_term (@{term Add} $ a $ b) = @{code Add} (floatarith_of_term a, floatarith_of_term b)
--- a/src/HOL/Decision_Procs/Cooper.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -3,7 +3,7 @@
 *)
 
 theory Cooper
-imports Complex_Main "~~/src/HOL/Library/Efficient_Nat" "~~/src/HOL/Library/Old_Recdef"
+imports Complex_Main "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
 (* Periodicity of dvd *)
@@ -1996,21 +1996,23 @@
 
 ML {* @{code cooper_test} () *}
 
-(* code_reflect Cooper_Procedure
+(*code_reflect Cooper_Procedure
   functions pa
-  file "~~/src/HOL/Tools/Qelim/cooper_procedure.ML" *)
+  file "~~/src/HOL/Tools/Qelim/cooper_procedure.ML"*)
 
 oracle linzqe_oracle = {*
 let
 
 fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t
      of NONE => error "Variable not found in the list!"
-      | SOME n => @{code Bound} n)
-  | num_of_term vs @{term "0::int"} = @{code C} 0
-  | num_of_term vs @{term "1::int"} = @{code C} 1
-  | num_of_term vs (@{term "numeral :: _ \<Rightarrow> int"} $ t) = @{code C} (HOLogic.dest_num t)
-  | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t) = @{code C} (~(HOLogic.dest_num t))
-  | num_of_term vs (Bound i) = @{code Bound} i
+      | SOME n => @{code Bound} (@{code nat_of_integer} n))
+  | num_of_term vs @{term "0::int"} = @{code C} (@{code int_of_integer} 0)
+  | num_of_term vs @{term "1::int"} = @{code C} (@{code int_of_integer} 1)
+  | num_of_term vs (@{term "numeral :: _ \<Rightarrow> int"} $ t) =
+      @{code C} (@{code int_of_integer} (HOLogic.dest_num t))
+  | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t) =
+      @{code C} (@{code int_of_integer} (~(HOLogic.dest_num t)))
+  | num_of_term vs (Bound i) = @{code Bound} (@{code nat_of_integer} i)
   | num_of_term vs (@{term "uminus :: int \<Rightarrow> int"} $ t') = @{code Neg} (num_of_term vs t')
   | num_of_term vs (@{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
       @{code Add} (num_of_term vs t1, num_of_term vs t2)
@@ -2018,9 +2020,9 @@
       @{code Sub} (num_of_term vs t1, num_of_term vs t2)
   | num_of_term vs (@{term "op * :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
       (case try HOLogic.dest_number t1
-       of SOME (_, i) => @{code Mul} (i, num_of_term vs t2)
+       of SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t2)
         | NONE => (case try HOLogic.dest_number t2
-                of SOME (_, i) => @{code Mul} (i, num_of_term vs t1)
+                of SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t1)
                  | NONE => error "num_of_term: unsupported multiplication"))
   | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
 
@@ -2034,7 +2036,7 @@
       @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
   | fm_of_term ps vs (@{term "op dvd :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
       (case try HOLogic.dest_number t1
-       of SOME (_, i) => @{code Dvd} (i, num_of_term vs t2)
+       of SOME (_, i) => @{code Dvd} (@{code int_of_integer} i, num_of_term vs t2)
         | NONE => error "num_of_term: unsupported dvd")
   | fm_of_term ps vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
@@ -2058,8 +2060,11 @@
       in @{code A} (fm_of_term ps vs' p) end
   | fm_of_term ps vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 
-fun term_of_num vs (@{code C} i) = HOLogic.mk_number HOLogic.intT i
-  | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs))
+fun term_of_num vs (@{code C} i) = HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
+  | term_of_num vs (@{code Bound} n) =
+      let
+        val q = @{code integer_of_nat} n
+      in fst (the (find_first (fn (_, m) => q = m) vs)) end
   | term_of_num vs (@{code Neg} t') = @{term "uminus :: int \<Rightarrow> int"} $ term_of_num vs t'
   | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $
       term_of_num vs t1 $ term_of_num vs t2
@@ -2097,7 +2102,10 @@
       HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
   | term_of_fm ps vs (@{code Iff} (t1, t2)) =
       @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
-  | term_of_fm ps vs (@{code Closed} n) = (fst o the) (find_first (fn (_, m) => m = n) ps)
+  | term_of_fm ps vs (@{code Closed} n) =
+      let
+        val q = @{code integer_of_nat} n
+      in (fst o the) (find_first (fn (_, m) => m = q) ps) end
   | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code NOT} (@{code Closed} n));
 
 fun term_bools acc t =
--- a/src/HOL/Decision_Procs/Ferrack.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -4,7 +4,7 @@
 
 theory Ferrack
 imports Complex_Main Dense_Linear_Order DP_Library
-  "~~/src/HOL/Library/Efficient_Nat" "~~/src/HOL/Library/Old_Recdef"
+  "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
 section {* Quantifier elimination for @{text "\<real> (0, 1, +, <)"} *}
@@ -1818,7 +1818,7 @@
   with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast
 qed
 
-lemma ferrack: 
+lemma ferrack:
   assumes qf: "qfree p"
   shows "qfree (ferrack p) \<and> ((Ifm bs (ferrack p)) = (\<exists> x. Ifm (x#bs) p))"
   (is "_ \<and> (?rhs = ?lhs)")
@@ -1922,12 +1922,15 @@
 oracle linr_oracle = {*
 let
 
-fun num_of_term vs (Free vT) = @{code Bound} (find_index (fn vT' => vT = vT') vs)
-  | num_of_term vs @{term "real (0::int)"} = @{code C} 0
-  | num_of_term vs @{term "real (1::int)"} = @{code C} 1
-  | num_of_term vs @{term "0::real"} = @{code C} 0
-  | num_of_term vs @{term "1::real"} = @{code C} 1
-  | num_of_term vs (Bound i) = @{code Bound} i
+val mk_C = @{code C} o @{code int_of_integer};
+val mk_Bound = @{code Bound} o @{code nat_of_integer};
+
+fun num_of_term vs (Free vT) = mk_Bound (find_index (fn vT' => vT = vT') vs)
+  | num_of_term vs @{term "real (0::int)"} = mk_C 0
+  | num_of_term vs @{term "real (1::int)"} = mk_C 1
+  | num_of_term vs @{term "0::real"} = mk_C 0
+  | num_of_term vs @{term "1::real"} = mk_C 1
+  | num_of_term vs (Bound i) = mk_Bound i
   | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
   | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
      @{code Add} (num_of_term vs t1, num_of_term vs t2)
@@ -1937,10 +1940,10 @@
      of @{code C} i => @{code Mul} (i, num_of_term vs t2)
       | _ => error "num_of_term: unsupported multiplication")
   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ t') =
-     (@{code C} (snd (HOLogic.dest_number t'))
+     (mk_C (snd (HOLogic.dest_number t'))
        handle TERM _ => error ("num_of_term: unknown term"))
   | num_of_term vs t' =
-     (@{code C} (snd (HOLogic.dest_number t'))
+     (mk_C (snd (HOLogic.dest_number t'))
        handle TERM _ => error ("num_of_term: unknown term"));
 
 fun fm_of_term vs @{term True} = @{code T}
@@ -1963,8 +1966,9 @@
       @{code A} (fm_of_term (("", dummyT) ::  vs) p)
   | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 
-fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $ HOLogic.mk_number HOLogic.intT i
-  | term_of_num vs (@{code Bound} n) = Free (nth vs n)
+fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $
+      HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
+  | term_of_num vs (@{code Bound} n) = Free (nth vs (@{code integer_of_nat} n))
   | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t'
   | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $
       term_of_num vs t1 $ term_of_num vs t2
@@ -2026,3 +2030,4 @@
   by rferrack
 
 end
+
--- a/src/HOL/Decision_Procs/MIR.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -4,7 +4,7 @@
 
 theory MIR
 imports Complex_Main Dense_Linear_Order DP_Library
-  "~~/src/HOL/Library/Efficient_Nat" "~~/src/HOL/Library/Old_Recdef"
+  "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
 section {* Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"} *}
@@ -5521,14 +5521,18 @@
 oracle mirfr_oracle = {* fn (proofs, ct) =>
 let
 
+val mk_C = @{code C} o @{code int_of_integer};
+val mk_Dvd = @{code Dvd} o apfst @{code int_of_integer};
+val mk_Bound = @{code Bound} o @{code nat_of_integer};
+
 fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t
      of NONE => error "Variable not found in the list!"
-      | SOME n => @{code Bound} n)
-  | num_of_term vs @{term "real (0::int)"} = @{code C} 0
-  | num_of_term vs @{term "real (1::int)"} = @{code C} 1
-  | num_of_term vs @{term "0::real"} = @{code C} 0
-  | num_of_term vs @{term "1::real"} = @{code C} 1
-  | num_of_term vs (Bound i) = @{code Bound} i
+      | SOME n => mk_Bound n)
+  | num_of_term vs @{term "real (0::int)"} = mk_C 0
+  | num_of_term vs @{term "real (1::int)"} = mk_C 1
+  | num_of_term vs @{term "0::real"} = mk_C 0
+  | num_of_term vs @{term "1::real"} = mk_C 1
+  | num_of_term vs (Bound i) = mk_Bound i
   | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
   | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
       @{code Add} (num_of_term vs t1, num_of_term vs t2)
@@ -5539,17 +5543,17 @@
        of @{code C} i => @{code Mul} (i, num_of_term vs t2)
         | _ => error "num_of_term: unsupported Multiplication")
   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
-      @{code C} (HOLogic.dest_num t')
+      mk_C (HOLogic.dest_num t')
   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t')) =
-      @{code C} (~ (HOLogic.dest_num t'))
+      mk_C (~ (HOLogic.dest_num t'))
   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
       @{code Floor} (num_of_term vs t')
   | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) =
       @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t')))
   | num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
-      @{code C} (HOLogic.dest_num t')
+      mk_C (HOLogic.dest_num t')
   | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> real"} $ t') =
-      @{code C} (~ (HOLogic.dest_num t'))
+      mk_C (~ (HOLogic.dest_num t'))
   | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
 
 fun fm_of_term vs @{term True} = @{code T}
@@ -5561,9 +5565,9 @@
   | fm_of_term vs (@{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) 
   | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
-      @{code Dvd} (HOLogic.dest_num t1, num_of_term vs t2)
+      mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2)
   | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
-      @{code Dvd} (~ (HOLogic.dest_num t1), num_of_term vs t2)
+      mk_Dvd (~ (HOLogic.dest_num t1), num_of_term vs t2)
   | fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
       @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
   | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
@@ -5580,8 +5584,12 @@
       @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
   | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 
-fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $ HOLogic.mk_number HOLogic.intT i
-  | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs))
+fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $
+      HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
+  | term_of_num vs (@{code Bound} n) =
+      let
+        val m = @{code integer_of_nat} n;
+      in fst (the (find_first (fn (_, q) => m = q) vs)) end
   | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) =
       @{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ term_of_num vs t')
   | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t'
@@ -5660,3 +5668,4 @@
   by mir
 
 end
+
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -6,7 +6,7 @@
 
 theory Parametric_Ferrante_Rackoff
 imports Reflected_Multivariate_Polynomial Dense_Linear_Order DP_Library
-  "~~/src/HOL/Library/Efficient_Nat" "~~/src/HOL/Library/Old_Recdef"
+  "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
 begin
 
 subsection {* Terms *}
@@ -2795,6 +2795,10 @@
 fun binopT T = T --> T --> T;
 fun relT T = T --> T --> @{typ bool};
 
+val mk_C = @{code C} o pairself @{code int_of_integer};
+val mk_poly_Bound = @{code poly.Bound} o @{code nat_of_integer};
+val mk_Bound = @{code Bound} o @{code nat_of_integer};
+
 val dest_num = snd o HOLogic.dest_number;
 
 fun try_dest_num t = SOME ((snd o HOLogic.dest_number) t)
@@ -2812,19 +2816,19 @@
   | num_of_term ps (Const (@{const_name Groups.plus}, _) $ a $ b) = @{code poly.Add} (num_of_term ps a, num_of_term ps b)
   | num_of_term ps (Const (@{const_name Groups.minus}, _) $ a $ b) = @{code poly.Sub} (num_of_term ps a, num_of_term ps b)
   | num_of_term ps (Const (@{const_name Groups.times}, _) $ a $ b) = @{code poly.Mul} (num_of_term ps a, num_of_term ps b)
-  | num_of_term ps (Const (@{const_name Power.power}, _) $ a $ n) = @{code poly.Pw} (num_of_term ps a, dest_nat n)
-  | num_of_term ps (Const (@{const_name Fields.divide}, _) $ a $ b) = @{code poly.C} (dest_num a, dest_num b)
+  | num_of_term ps (Const (@{const_name Power.power}, _) $ a $ n) = @{code poly.Pw} (num_of_term ps a, @{code nat_of_integer} (dest_nat n))
+  | num_of_term ps (Const (@{const_name Fields.divide}, _) $ a $ b) = mk_C (dest_num a, dest_num b)
   | num_of_term ps t = (case try_dest_num t
-     of SOME k => @{code poly.C} (k, 1)
-      | NONE => @{code poly.Bound} (the_index ps t));
+     of SOME k => mk_C (k, 1)
+      | NONE => mk_poly_Bound (the_index ps t));
 
 fun tm_of_term fs ps (Const(@{const_name Groups.uminus}, _) $ t) = @{code Neg} (tm_of_term fs ps t)
   | tm_of_term fs ps (Const(@{const_name Groups.plus}, _) $ a $ b) = @{code Add} (tm_of_term fs ps a, tm_of_term fs ps b)
   | tm_of_term fs ps (Const(@{const_name Groups.minus}, _) $ a $ b) = @{code Sub} (tm_of_term fs ps a, tm_of_term fs ps b)
   | tm_of_term fs ps (Const(@{const_name Groups.times}, _) $ a $ b) = @{code Mul} (num_of_term ps a, tm_of_term fs ps b)
   | tm_of_term fs ps t = (@{code CP} (num_of_term ps t) 
-      handle TERM _ => @{code Bound} (the_index fs t)
-           | General.Subscript => @{code Bound} (the_index fs t));
+      handle TERM _ => mk_Bound (the_index fs t)
+           | General.Subscript => mk_Bound (the_index fs t));
 
 fun fm_of_term fs ps @{term True} = @{code T}
   | fm_of_term fs ps @{term False} = @{code F}
@@ -2850,21 +2854,25 @@
   | fm_of_term fs ps _ = error "fm_of_term";
 
 fun term_of_num T ps (@{code poly.C} (a, b)) = 
-    (if b = 1 then HOLogic.mk_number T a
-     else if b = 0 then Const (@{const_name Groups.zero}, T)
-     else Const (@{const_name Fields.divide}, binopT T) $ HOLogic.mk_number T a $ HOLogic.mk_number T b)
-  | term_of_num T ps (@{code poly.Bound} i) = nth ps i
+      let
+        val (c, d) = pairself (@{code integer_of_int}) (a, b)
+      in
+        (if d = 1 then HOLogic.mk_number T c
+        else if d = 0 then Const (@{const_name Groups.zero}, T)
+        else Const (@{const_name Fields.divide}, binopT T) $ HOLogic.mk_number T c $ HOLogic.mk_number T d)
+      end
+  | term_of_num T ps (@{code poly.Bound} i) = nth ps (@{code integer_of_nat} i)
   | term_of_num T ps (@{code poly.Add} (a, b)) = Const (@{const_name Groups.plus}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
   | term_of_num T ps (@{code poly.Mul} (a, b)) = Const (@{const_name Groups.times}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
   | term_of_num T ps (@{code poly.Sub} (a, b)) = Const (@{const_name Groups.minus}, binopT T) $ term_of_num T ps a $ term_of_num T ps b
   | term_of_num T ps (@{code poly.Neg} a) = Const (@{const_name Groups.uminus}, T --> T) $ term_of_num T ps a
-  | term_of_num T ps (@{code poly.Pw} (a, n)) =
-      Const (@{const_name Power.power}, T --> @{typ nat} --> T) $ term_of_num T ps a $ HOLogic.mk_number HOLogic.natT n
+  | term_of_num T ps (@{code poly.Pw} (a, n)) = Const (@{const_name Power.power},
+      T --> @{typ nat} --> T) $ term_of_num T ps a $ HOLogic.mk_number HOLogic.natT (@{code integer_of_nat} n)
   | term_of_num T ps (@{code poly.CN} (c, n, p)) =
       term_of_num T ps (@{code poly.Add} (c, @{code poly.Mul} (@{code poly.Bound} n, p)));
 
 fun term_of_tm T fs ps (@{code CP} p) = term_of_num T ps p
-  | term_of_tm T fs ps (@{code Bound} i) = nth fs i
+  | term_of_tm T fs ps (@{code Bound} i) = nth fs (@{code integer_of_nat} i)
   | term_of_tm T fs ps (@{code Add} (a, b)) = Const (@{const_name Groups.plus}, binopT T) $ term_of_tm T fs ps a $ term_of_tm T fs ps b
   | term_of_tm T fs ps (@{code Mul} (a, b)) = Const (@{const_name Groups.times}, binopT T) $ term_of_num T ps a $ term_of_tm T fs ps b
   | term_of_tm T fs ps (@{code Sub} (a, b)) = Const (@{const_name Groups.minus}, binopT T) $ term_of_tm T fs ps a $ term_of_tm T fs ps b
@@ -2993,3 +3001,4 @@
 *)
 end
 
+
--- a/src/HOL/Imperative_HOL/Array.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Imperative_HOL/Array.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -361,38 +361,38 @@
 subsubsection {* Logical intermediate layer *}
 
 definition new' where
-  [code del]: "new' = Array.new o Code_Numeral.nat_of"
+  [code del]: "new' = Array.new o nat_of_integer"
 
 lemma [code]:
-  "Array.new = new' o Code_Numeral.of_nat"
+  "Array.new = new' o of_nat"
   by (simp add: new'_def o_def)
 
 definition make' where
-  [code del]: "make' i f = Array.make (Code_Numeral.nat_of i) (f o Code_Numeral.of_nat)"
+  [code del]: "make' i f = Array.make (nat_of_integer i) (f o of_nat)"
 
 lemma [code]:
-  "Array.make n f = make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)"
+  "Array.make n f = make' (of_nat n) (f o nat_of_integer)"
   by (simp add: make'_def o_def)
 
 definition len' where
-  [code del]: "len' a = Array.len a \<guillemotright>= (\<lambda>n. return (Code_Numeral.of_nat n))"
+  [code del]: "len' a = Array.len a \<guillemotright>= (\<lambda>n. return (of_nat n))"
 
 lemma [code]:
-  "Array.len a = len' a \<guillemotright>= (\<lambda>i. return (Code_Numeral.nat_of i))"
+  "Array.len a = len' a \<guillemotright>= (\<lambda>i. return (nat_of_integer i))"
   by (simp add: len'_def)
 
 definition nth' where
-  [code del]: "nth' a = Array.nth a o Code_Numeral.nat_of"
+  [code del]: "nth' a = Array.nth a o nat_of_integer"
 
 lemma [code]:
-  "Array.nth a n = nth' a (Code_Numeral.of_nat n)"
+  "Array.nth a n = nth' a (of_nat n)"
   by (simp add: nth'_def)
 
 definition upd' where
-  [code del]: "upd' a i x = Array.upd (Code_Numeral.nat_of i) x a \<guillemotright> return ()"
+  [code del]: "upd' a i x = Array.upd (nat_of_integer i) x a \<guillemotright> return ()"
 
 lemma [code]:
-  "Array.upd i x a = upd' a (Code_Numeral.of_nat i) x \<guillemotright> return a"
+  "Array.upd i x a = upd' a (of_nat i) x \<guillemotright> return a"
   by (simp add: upd'_def upd_return)
 
 lemma [code]:
@@ -501,3 +501,4 @@
 code_const "HOL.equal :: 'a array \<Rightarrow> 'a array \<Rightarrow> bool" (Scala infixl 5 "==")
 
 end
+
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -8,7 +8,6 @@
 imports
   Heap
   "~~/src/HOL/Library/Monad_Syntax"
-  "~~/src/HOL/Library/Code_Natural"
 begin
 
 subsection {* The monad *}
@@ -529,33 +528,31 @@
 import qualified Data.STRef;
 import qualified Data.Array.ST;
 
-import Natural;
-
 type RealWorld = Control.Monad.ST.RealWorld;
 type ST s a = Control.Monad.ST.ST s a;
 type STRef s a = Data.STRef.STRef s a;
-type STArray s a = Data.Array.ST.STArray s Natural a;
+type STArray s a = Data.Array.ST.STArray s Integer a;
 
 newSTRef = Data.STRef.newSTRef;
 readSTRef = Data.STRef.readSTRef;
 writeSTRef = Data.STRef.writeSTRef;
 
-newArray :: Natural -> a -> ST s (STArray s a);
+newArray :: Integer -> a -> ST s (STArray s a);
 newArray k = Data.Array.ST.newArray (0, k);
 
 newListArray :: [a] -> ST s (STArray s a);
 newListArray xs = Data.Array.ST.newListArray (0, (fromInteger . toInteger . length) xs) xs;
 
-newFunArray :: Natural -> (Natural -> a) -> ST s (STArray s a);
+newFunArray :: Integer -> (Integer -> a) -> ST s (STArray s a);
 newFunArray k f = Data.Array.ST.newListArray (0, k) (map f [0..k-1]);
 
-lengthArray :: STArray s a -> ST s Natural;
+lengthArray :: STArray s a -> ST s Integer;
 lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
 
-readArray :: STArray s a -> Natural -> ST s a;
+readArray :: STArray s a -> Integer -> ST s a;
 readArray = Data.Array.ST.readArray;
 
-writeArray :: STArray s a -> Natural -> a -> ST s ();
+writeArray :: STArray s a -> Integer -> a -> ST s ();
 writeArray = Data.Array.ST.writeArray;*}
 
 code_reserved Haskell Heap
@@ -590,16 +587,16 @@
 
 object Array {
   import collection.mutable.ArraySeq
-  def alloc[A](n: Natural)(x: A): ArraySeq[A] =
-    ArraySeq.fill(n.as_Int)(x)
-  def make[A](n: Natural)(f: Natural => A): ArraySeq[A] =
-    ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural(k)))
-  def len[A](a: ArraySeq[A]): Natural =
-    Natural(a.length)
-  def nth[A](a: ArraySeq[A], n: Natural): A =
-    a(n.as_Int)
-  def upd[A](a: ArraySeq[A], n: Natural, x: A): Unit =
-    a.update(n.as_Int, x)
+  def alloc[A](n: BigInt)(x: A): ArraySeq[A] =
+    ArraySeq.fill(n.toInt)(x)
+  def make[A](n: BigInt)(f: BigInt => A): ArraySeq[A] =
+    ArraySeq.tabulate(n.toInt)((k: Int) => f(BigInt(k)))
+  def len[A](a: ArraySeq[A]): BigInt =
+    BigInt(a.length)
+  def nth[A](a: ArraySeq[A], n: BigInt): A =
+    a(n.toInt)
+  def upd[A](a: ArraySeq[A], n: BigInt, x: A): Unit =
+    a.update(n.toInt, x)
   def freeze[A](a: ArraySeq[A]): List[A] =
     a.toList
 }
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -9,7 +9,7 @@
   "~~/src/HOL/Imperative_HOL/Imperative_HOL"
   Subarray
   "~~/src/HOL/Library/Multiset"
-  "~~/src/HOL/Library/Efficient_Nat"
+  "~~/src/HOL/Library/Code_Target_Numeral"
 begin
 
 text {* We prove QuickSort correct in the Relational Calculus. *}
@@ -657,7 +657,12 @@
 
 code_reserved SML upto
 
-ML {* @{code qsort} (Array.fromList [42, 2, 3, 5, 0, 1705, 8, 3, 15]) () *}
+definition "example = do {
+    a \<leftarrow> Array.of_list [42, 2, 3, 5, 0, 1705, 8, 3, 15];
+    qsort a
+  }"
+
+ML {* @{code example} () *}
 
 export_code qsort checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp
 
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* Linked Lists by ML references *}
 
 theory Linked_Lists
-imports "../Imperative_HOL" "~~/src/HOL/Library/Code_Integer"
+imports "../Imperative_HOL" "~~/src/HOL/Library/Code_Target_Int"
 begin
 
 section {* Definition of Linked Lists *}
--- a/src/HOL/Int.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Int.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -853,7 +853,7 @@
 apply (rule nat_mono, simp_all)
 done
 
-lemma nat_numeral [simp, code_abbrev]:
+lemma nat_numeral [simp]:
   "nat (numeral k) = numeral k"
   by (simp add: nat_eq_iff)
 
--- a/src/HOL/Lazy_Sequence.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Lazy_Sequence.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -169,13 +169,14 @@
 where
   "those xq = Option.map lazy_sequence_of_list (List.those (list_of_lazy_sequence xq))"
   
-function iterate_upto :: "(code_numeral \<Rightarrow> 'a) \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> 'a lazy_sequence"
+function iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a lazy_sequence"
 where
   "iterate_upto f n m =
     Lazy_Sequence (\<lambda>_. if n > m then None else Some (f n, iterate_upto f (n + 1) m))"
   by pat_completeness auto
 
-termination by (relation "measure (%(f, n, m). Code_Numeral.nat_of (m + 1 - n))") auto
+termination by (relation "measure (\<lambda>(f, n, m). nat_of_natural (m + 1 - n))")
+  (auto simp add: less_natural_def)
 
 definition not_seq :: "unit lazy_sequence \<Rightarrow> unit lazy_sequence"
 where
@@ -225,7 +226,7 @@
 subsubsection {* Small lazy typeclasses *}
 
 class small_lazy =
-  fixes small_lazy :: "code_numeral \<Rightarrow> 'a lazy_sequence"
+  fixes small_lazy :: "natural \<Rightarrow> 'a lazy_sequence"
 
 instantiation unit :: small_lazy
 begin
@@ -252,7 +253,7 @@
   by (relation "measure (%(d, i). nat (d + 1 - i))") auto
 
 definition
-  "small_lazy d = small_lazy' (Code_Numeral.int_of d) (- (Code_Numeral.int_of d))"
+  "small_lazy d = small_lazy' (int (nat_of_natural d)) (- (int (nat_of_natural d)))"
 
 instance ..
 
@@ -271,7 +272,7 @@
 instantiation list :: (small_lazy) small_lazy
 begin
 
-fun small_lazy_list :: "code_numeral \<Rightarrow> 'a list lazy_sequence"
+fun small_lazy_list :: "natural \<Rightarrow> 'a list lazy_sequence"
 where
   "small_lazy_list d = append (single [])
     (if d > 0 then bind (product (small_lazy (d - 1))
--- a/src/HOL/Library/Cardinality.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Cardinality.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -225,11 +225,24 @@
 instance by intro_classes (simp_all add: card_UNIV_int_def finite_UNIV_int_def infinite_UNIV_int)
 end
 
-instantiation code_numeral :: card_UNIV begin
-definition "finite_UNIV = Phantom(code_numeral) False"
-definition "card_UNIV = Phantom(code_numeral) 0"
-instance
-  by(intro_classes)(auto simp add: card_UNIV_code_numeral_def finite_UNIV_code_numeral_def type_definition.univ[OF type_definition_code_numeral] card_eq_0_iff dest!: finite_imageD intro: inj_onI)
+instantiation natural :: card_UNIV begin
+definition "finite_UNIV = Phantom(natural) False"
+definition "card_UNIV = Phantom(natural) 0"
+instance proof
+qed (auto simp add: finite_UNIV_natural_def card_UNIV_natural_def card_eq_0_iff
+  type_definition.univ [OF type_definition_natural] natural_eq_iff
+  dest!: finite_imageD intro: inj_onI)
+end
+
+declare [[show_consts]]
+
+instantiation integer :: card_UNIV begin
+definition "finite_UNIV = Phantom(integer) False"
+definition "card_UNIV = Phantom(integer) 0"
+instance proof
+qed (auto simp add: finite_UNIV_integer_def card_UNIV_integer_def card_eq_0_iff
+  type_definition.univ [OF type_definition_integer] infinite_UNIV_int
+  dest!: finite_imageD intro: inj_onI)
 end
 
 instantiation list :: (type) card_UNIV begin
--- a/src/HOL/Library/Code_Binary_Nat.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Binary_Nat.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -21,10 +21,6 @@
 
 code_datatype "0::nat" nat_of_num
 
-lemma [code_abbrev]:
-  "nat_of_num = numeral"
-  by (fact nat_of_num_numeral)
-
 lemma [code]:
   "num_of_nat 0 = Num.One"
   "num_of_nat (nat_of_num k) = k"
--- a/src/HOL/Library/Code_Char_chr.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Char_chr.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* Code generation of pretty characters with character codes *}
 
 theory Code_Char_chr
-imports Char_nat Code_Char Code_Integer Main
+imports Char_nat Code_Char Code_Target_Int Main
 begin
 
 definition
--- a/src/HOL/Library/Code_Integer.thy	Fri Feb 15 08:31:30 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,161 +0,0 @@
-(*  Title:      HOL/Library/Code_Integer.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Pretty integer literals for code generation *}
-
-theory Code_Integer
-imports Main Code_Natural
-begin
-
-text {*
-  Representation-ignorant code equations for conversions.
-*}
-
-lemma nat_code [code]:
-  "nat k = (if k \<le> 0 then 0 else
-     let
-       (l, j) = divmod_int k 2;
-       n = nat l;
-       l' = n + n
-     in if j = 0 then l' else Suc l')"
-proof -
-  have "2 = nat 2" by simp
-  show ?thesis
-    apply (subst mult_2 [symmetric])
-    apply (auto simp add: Let_def divmod_int_mod_div not_le
-     nat_div_distrib nat_mult_distrib mult_div_cancel mod_2_not_eq_zero_eq_one_int)
-    apply (unfold `2 = nat 2`)
-    apply (subst nat_mod_distrib [symmetric])
-    apply simp_all
-  done
-qed
-
-lemma (in ring_1) of_int_code:
-  "of_int k = (if k = 0 then 0
-     else if k < 0 then - of_int (- k)
-     else let
-       (l, j) = divmod_int k 2;
-       l' = 2 * of_int l
-     in if j = 0 then l' else l' + 1)"
-proof -
-  from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
-  show ?thesis
-    by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int
-      of_int_add [symmetric]) (simp add: * mult_commute)
-qed
-
-declare of_int_code [code]
-
-text {*
-  HOL numeral expressions are mapped to integer literals
-  in target languages, using predefined target language
-  operations for abstract integer operations.
-*}
-
-code_type int
-  (SML "IntInf.int")
-  (OCaml "Big'_int.big'_int")
-  (Haskell "Integer")
-  (Scala "BigInt")
-  (Eval "int")
-
-code_instance int :: equal
-  (Haskell -)
-
-code_const "0::int"
-  (SML "0")
-  (OCaml "Big'_int.zero'_big'_int")
-  (Haskell "0")
-  (Scala "BigInt(0)")
-
-setup {*
-  fold (Numeral.add_code @{const_name Int.Pos}
-    false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
-*}
-
-setup {*
-  fold (Numeral.add_code @{const_name Int.Neg}
-    true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
-*}
-
-code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
-  (SML "IntInf.+ ((_), (_))")
-  (OCaml "Big'_int.add'_big'_int")
-  (Haskell infixl 6 "+")
-  (Scala infixl 7 "+")
-  (Eval infixl 8 "+")
-
-code_const "uminus \<Colon> int \<Rightarrow> int"
-  (SML "IntInf.~")
-  (OCaml "Big'_int.minus'_big'_int")
-  (Haskell "negate")
-  (Scala "!(- _)")
-  (Eval "~/ _")
-
-code_const "op - \<Colon> int \<Rightarrow> int \<Rightarrow> int"
-  (SML "IntInf.- ((_), (_))")
-  (OCaml "Big'_int.sub'_big'_int")
-  (Haskell infixl 6 "-")
-  (Scala infixl 7 "-")
-  (Eval infixl 8 "-")
-
-code_const Int.dup
-  (SML "IntInf.*/ (2,/ (_))")
-  (OCaml "Big'_int.mult'_big'_int/ 2")
-  (Haskell "!(2 * _)")
-  (Scala "!(2 * _)")
-  (Eval "!(2 * _)")
-
-code_const Int.sub
-  (SML "!(raise/ Fail/ \"sub\")")
-  (OCaml "failwith/ \"sub\"")
-  (Haskell "error/ \"sub\"")
-  (Scala "!sys.error(\"sub\")")
-
-code_const "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"
-  (SML "IntInf.* ((_), (_))")
-  (OCaml "Big'_int.mult'_big'_int")
-  (Haskell infixl 7 "*")
-  (Scala infixl 8 "*")
-  (Eval infixl 9 "*")
-
-code_const pdivmod
-  (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
-  (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
-  (Haskell "divMod/ (abs _)/ (abs _)")
-  (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
-  (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
-
-code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
-  (SML "!((_ : IntInf.int) = _)")
-  (OCaml "Big'_int.eq'_big'_int")
-  (Haskell infix 4 "==")
-  (Scala infixl 5 "==")
-  (Eval infixl 6 "=")
-
-code_const "op \<le> \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
-  (SML "IntInf.<= ((_), (_))")
-  (OCaml "Big'_int.le'_big'_int")
-  (Haskell infix 4 "<=")
-  (Scala infixl 4 "<=")
-  (Eval infixl 6 "<=")
-
-code_const "op < \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
-  (SML "IntInf.< ((_), (_))")
-  (OCaml "Big'_int.lt'_big'_int")
-  (Haskell infix 4 "<")
-  (Scala infixl 4 "<")
-  (Eval infixl 6 "<")
-
-code_const Code_Numeral.int_of
-  (SML "IntInf.fromInt")
-  (OCaml "_")
-  (Haskell "Prelude.toInteger")
-  (Scala "!_.as'_BigInt")
-  (Eval "_")
-
-code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term"
-  (Eval "HOLogic.mk'_number/ HOLogic.intT")
-
-end
--- a/src/HOL/Library/Code_Natural.thy	Fri Feb 15 08:31:30 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,148 +0,0 @@
-(*  Title:      HOL/Library/Code_Natural.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-theory Code_Natural
-imports "../Main"
-begin
-
-section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
-
-code_include Haskell "Natural"
-{*import Data.Array.ST;
-
-newtype Natural = Natural Integer deriving (Eq, Show, Read);
-
-instance Num Natural where {
-  fromInteger k = Natural (if k >= 0 then k else 0);
-  Natural n + Natural m = Natural (n + m);
-  Natural n - Natural m = fromInteger (n - m);
-  Natural n * Natural m = Natural (n * m);
-  abs n = n;
-  signum _ = 1;
-  negate n = error "negate Natural";
-};
-
-instance Ord Natural where {
-  Natural n <= Natural m = n <= m;
-  Natural n < Natural m = n < m;
-};
-
-instance Ix Natural where {
-  range (Natural n, Natural m) = map Natural (range (n, m));
-  index (Natural n, Natural m) (Natural q) = index (n, m) q;
-  inRange (Natural n, Natural m) (Natural q) = inRange (n, m) q;
-  rangeSize (Natural n, Natural m) = rangeSize (n, m);
-};
-
-instance Real Natural where {
-  toRational (Natural n) = toRational n;
-};
-
-instance Enum Natural where {
-  toEnum k = fromInteger (toEnum k);
-  fromEnum (Natural n) = fromEnum n;
-};
-
-instance Integral Natural where {
-  toInteger (Natural n) = n;
-  divMod n m = quotRem n m;
-  quotRem (Natural n) (Natural m)
-    | (m == 0) = (0, Natural n)
-    | otherwise = (Natural k, Natural l) where (k, l) = quotRem n m;
-};*}
-
-
-code_reserved Haskell Natural
-
-code_include Scala "Natural"
-{*object Natural {
-
-  def apply(numeral: BigInt): Natural = new Natural(numeral max 0)
-  def apply(numeral: Int): Natural = Natural(BigInt(numeral))
-  def apply(numeral: String): Natural = Natural(BigInt(numeral))
-
-}
-
-class Natural private(private val value: BigInt) {
-
-  override def hashCode(): Int = this.value.hashCode()
-
-  override def equals(that: Any): Boolean = that match {
-    case that: Natural => this equals that
-    case _ => false
-  }
-
-  override def toString(): String = this.value.toString
-
-  def equals(that: Natural): Boolean = this.value == that.value
-
-  def as_BigInt: BigInt = this.value
-  def as_Int: Int = if (this.value >= scala.Int.MinValue && this.value <= scala.Int.MaxValue)
-      this.value.intValue
-    else error("Int value out of range: " + this.value.toString)
-
-  def +(that: Natural): Natural = new Natural(this.value + that.value)
-  def -(that: Natural): Natural = Natural(this.value - that.value)
-  def *(that: Natural): Natural = new Natural(this.value * that.value)
-
-  def /%(that: Natural): (Natural, Natural) = if (that.value == 0) (new Natural(0), this)
-    else {
-      val (k, l) = this.value /% that.value
-      (new Natural(k), new Natural(l))
-    }
-
-  def <=(that: Natural): Boolean = this.value <= that.value
-
-  def <(that: Natural): Boolean = this.value < that.value
-
-}
-*}
-
-code_reserved Scala Natural
-
-code_type code_numeral
-  (Haskell "Natural.Natural")
-  (Scala "Natural")
-
-setup {*
-  fold (Numeral.add_code @{const_name Code_Numeral.Num}
-    false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
-*}
-
-code_instance code_numeral :: equal
-  (Haskell -)
-
-code_const "0::code_numeral"
-  (Haskell "0")
-  (Scala "Natural(0)")
-
-code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (Haskell infixl 6 "+")
-  (Scala infixl 7 "+")
-
-code_const "minus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (Haskell infixl 6 "-")
-  (Scala infixl 7 "-")
-
-code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
-  (Haskell infixl 7 "*")
-  (Scala infixl 8 "*")
-
-code_const Code_Numeral.div_mod
-  (Haskell "divMod")
-  (Scala infixl 8 "/%")
-
-code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (Haskell infix 4 "==")
-  (Scala infixl 5 "==")
-
-code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (Haskell infix 4 "<=")
-  (Scala infixl 4 "<=")
-
-code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
-  (Haskell infix 4 "<")
-  (Scala infixl 4 "<")
-
-end
--- a/src/HOL/Library/Code_Numeral_Types.thy	Fri Feb 15 08:31:30 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,942 +0,0 @@
-(*  Title:      HOL/Library/Code_Numeral_Types.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Numeric types for code generation onto target language numerals only *}
-
-theory Code_Numeral_Types
-imports Main Nat_Transfer Divides Lifting
-begin
-
-subsection {* Type of target language integers *}
-
-typedef integer = "UNIV \<Colon> int set"
-  morphisms int_of_integer integer_of_int ..
-
-setup_lifting (no_code) type_definition_integer
-
-lemma integer_eq_iff:
-  "k = l \<longleftrightarrow> int_of_integer k = int_of_integer l"
-  by transfer rule
-
-lemma integer_eqI:
-  "int_of_integer k = int_of_integer l \<Longrightarrow> k = l"
-  using integer_eq_iff [of k l] by simp
-
-lemma int_of_integer_integer_of_int [simp]:
-  "int_of_integer (integer_of_int k) = k"
-  by transfer rule
-
-lemma integer_of_int_int_of_integer [simp]:
-  "integer_of_int (int_of_integer k) = k"
-  by transfer rule
-
-instantiation integer :: ring_1
-begin
-
-lift_definition zero_integer :: integer
-  is "0 :: int"
-  .
-
-declare zero_integer.rep_eq [simp]
-
-lift_definition one_integer :: integer
-  is "1 :: int"
-  .
-
-declare one_integer.rep_eq [simp]
-
-lift_definition plus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
-  is "plus :: int \<Rightarrow> int \<Rightarrow> int"
-  .
-
-declare plus_integer.rep_eq [simp]
-
-lift_definition uminus_integer :: "integer \<Rightarrow> integer"
-  is "uminus :: int \<Rightarrow> int"
-  .
-
-declare uminus_integer.rep_eq [simp]
-
-lift_definition minus_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
-  is "minus :: int \<Rightarrow> int \<Rightarrow> int"
-  .
-
-declare minus_integer.rep_eq [simp]
-
-lift_definition times_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
-  is "times :: int \<Rightarrow> int \<Rightarrow> int"
-  .
-
-declare times_integer.rep_eq [simp]
-
-instance proof
-qed (transfer, simp add: algebra_simps)+
-
-end
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer (of_nat :: nat \<Rightarrow> int) (of_nat :: nat \<Rightarrow> integer)"
-  by (unfold of_nat_def [abs_def])  transfer_prover
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer (\<lambda>k :: int. k :: int) (of_int :: int \<Rightarrow> integer)"
-proof -
-  have "fun_rel HOL.eq cr_integer (of_int :: int \<Rightarrow> int) (of_int :: int \<Rightarrow> integer)"
-    by (unfold of_int_of_nat [abs_def]) transfer_prover
-  then show ?thesis by (simp add: id_def)
-qed
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer (numeral :: num \<Rightarrow> int) (numeral :: num \<Rightarrow> integer)"
-proof -
-  have "fun_rel HOL.eq cr_integer (numeral :: num \<Rightarrow> int) (\<lambda>n. of_int (numeral n))"
-    by transfer_prover
-  then show ?thesis by simp
-qed
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer (neg_numeral :: num \<Rightarrow> int) (neg_numeral :: num \<Rightarrow> integer)"
-  by (unfold neg_numeral_def [abs_def]) transfer_prover
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq (fun_rel HOL.eq cr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
-  by (unfold Num.sub_def [abs_def]) transfer_prover
-
-lemma int_of_integer_of_nat [simp]:
-  "int_of_integer (of_nat n) = of_nat n"
-  by transfer rule
-
-lift_definition integer_of_nat :: "nat \<Rightarrow> integer"
-  is "of_nat :: nat \<Rightarrow> int"
-  .
-
-lemma integer_of_nat_eq_of_nat [code]:
-  "integer_of_nat = of_nat"
-  by transfer rule
-
-lemma int_of_integer_integer_of_nat [simp]:
-  "int_of_integer (integer_of_nat n) = of_nat n"
-  by transfer rule
-
-lift_definition nat_of_integer :: "integer \<Rightarrow> nat"
-  is Int.nat
-  .
-
-lemma nat_of_integer_of_nat [simp]:
-  "nat_of_integer (of_nat n) = n"
-  by transfer simp
-
-lemma int_of_integer_of_int [simp]:
-  "int_of_integer (of_int k) = k"
-  by transfer simp
-
-lemma nat_of_integer_integer_of_nat [simp]:
-  "nat_of_integer (integer_of_nat n) = n"
-  by transfer simp
-
-lemma integer_of_int_eq_of_int [simp, code_abbrev]:
-  "integer_of_int = of_int"
-  by transfer (simp add: fun_eq_iff)
-
-lemma of_int_integer_of [simp]:
-  "of_int (int_of_integer k) = (k :: integer)"
-  by transfer rule
-
-lemma int_of_integer_numeral [simp]:
-  "int_of_integer (numeral k) = numeral k"
-  by transfer rule
-
-lemma int_of_integer_neg_numeral [simp]:
-  "int_of_integer (neg_numeral k) = neg_numeral k"
-  by transfer rule
-
-lemma int_of_integer_sub [simp]:
-  "int_of_integer (Num.sub k l) = Num.sub k l"
-  by transfer rule
-
-instantiation integer :: "{ring_div, equal, linordered_idom}"
-begin
-
-lift_definition div_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
-  is "Divides.div :: int \<Rightarrow> int \<Rightarrow> int"
-  .
-
-declare div_integer.rep_eq [simp]
-
-lift_definition mod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
-  is "Divides.mod :: int \<Rightarrow> int \<Rightarrow> int"
-  .
-
-declare mod_integer.rep_eq [simp]
-
-lift_definition abs_integer :: "integer \<Rightarrow> integer"
-  is "abs :: int \<Rightarrow> int"
-  .
-
-declare abs_integer.rep_eq [simp]
-
-lift_definition sgn_integer :: "integer \<Rightarrow> integer"
-  is "sgn :: int \<Rightarrow> int"
-  .
-
-declare sgn_integer.rep_eq [simp]
-
-lift_definition less_eq_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
-  is "less_eq :: int \<Rightarrow> int \<Rightarrow> bool"
-  .
-
-lift_definition less_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
-  is "less :: int \<Rightarrow> int \<Rightarrow> bool"
-  .
-
-lift_definition equal_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
-  is "HOL.equal :: int \<Rightarrow> int \<Rightarrow> bool"
-  .
-
-instance proof
-qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] mult_strict_right_mono linear)+
-
-end
-
-lemma [transfer_rule]:
-  "fun_rel cr_integer (fun_rel cr_integer cr_integer) (min :: _ \<Rightarrow> _ \<Rightarrow> int) (min :: _ \<Rightarrow> _ \<Rightarrow> integer)"
-  by (unfold min_def [abs_def]) transfer_prover
-
-lemma [transfer_rule]:
-  "fun_rel cr_integer (fun_rel cr_integer cr_integer) (max :: _ \<Rightarrow> _ \<Rightarrow> int) (max :: _ \<Rightarrow> _ \<Rightarrow> integer)"
-  by (unfold max_def [abs_def]) transfer_prover
-
-lemma int_of_integer_min [simp]:
-  "int_of_integer (min k l) = min (int_of_integer k) (int_of_integer l)"
-  by transfer rule
-
-lemma int_of_integer_max [simp]:
-  "int_of_integer (max k l) = max (int_of_integer k) (int_of_integer l)"
-  by transfer rule
-
-lemma nat_of_integer_non_positive [simp]:
-  "k \<le> 0 \<Longrightarrow> nat_of_integer k = 0"
-  by transfer simp
-
-lemma of_nat_of_integer [simp]:
-  "of_nat (nat_of_integer k) = max 0 k"
-  by transfer auto
-
-
-subsection {* Code theorems for target language integers *}
-
-text {* Constructors *}
-
-definition Pos :: "num \<Rightarrow> integer"
-where
-  [simp, code_abbrev]: "Pos = numeral"
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer numeral Pos"
-  by simp transfer_prover
-
-definition Neg :: "num \<Rightarrow> integer"
-where
-  [simp, code_abbrev]: "Neg = neg_numeral"
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_integer neg_numeral Neg"
-  by simp transfer_prover
-
-code_datatype "0::integer" Pos Neg
-
-
-text {* Auxiliary operations *}
-
-lift_definition dup :: "integer \<Rightarrow> integer"
-  is "\<lambda>k::int. k + k"
-  .
-
-lemma dup_code [code]:
-  "dup 0 = 0"
-  "dup (Pos n) = Pos (Num.Bit0 n)"
-  "dup (Neg n) = Neg (Num.Bit0 n)"
-  by (transfer, simp only: neg_numeral_def numeral_Bit0 minus_add_distrib)+
-
-lift_definition sub :: "num \<Rightarrow> num \<Rightarrow> integer"
-  is "\<lambda>m n. numeral m - numeral n :: int"
-  .
-
-lemma sub_code [code]:
-  "sub Num.One Num.One = 0"
-  "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)"
-  "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)"
-  "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)"
-  "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)"
-  "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)"
-  "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
-  "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
-  "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
-  by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+
-
-
-text {* Implementations *}
-
-lemma one_integer_code [code, code_unfold]:
-  "1 = Pos Num.One"
-  by simp
-
-lemma plus_integer_code [code]:
-  "k + 0 = (k::integer)"
-  "0 + l = (l::integer)"
-  "Pos m + Pos n = Pos (m + n)"
-  "Pos m + Neg n = sub m n"
-  "Neg m + Pos n = sub n m"
-  "Neg m + Neg n = Neg (m + n)"
-  by (transfer, simp)+
-
-lemma uminus_integer_code [code]:
-  "uminus 0 = (0::integer)"
-  "uminus (Pos m) = Neg m"
-  "uminus (Neg m) = Pos m"
-  by simp_all
-
-lemma minus_integer_code [code]:
-  "k - 0 = (k::integer)"
-  "0 - l = uminus (l::integer)"
-  "Pos m - Pos n = sub m n"
-  "Pos m - Neg n = Pos (m + n)"
-  "Neg m - Pos n = Neg (m + n)"
-  "Neg m - Neg n = sub n m"
-  by (transfer, simp)+
-
-lemma abs_integer_code [code]:
-  "\<bar>k\<bar> = (if (k::integer) < 0 then - k else k)"
-  by simp
-
-lemma sgn_integer_code [code]:
-  "sgn k = (if k = 0 then 0 else if (k::integer) < 0 then - 1 else 1)"
-  by simp
-
-lemma times_integer_code [code]:
-  "k * 0 = (0::integer)"
-  "0 * l = (0::integer)"
-  "Pos m * Pos n = Pos (m * n)"
-  "Pos m * Neg n = Neg (m * n)"
-  "Neg m * Pos n = Neg (m * n)"
-  "Neg m * Neg n = Pos (m * n)"
-  by simp_all
-
-definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
-where
-  "divmod_integer k l = (k div l, k mod l)"
-
-lemma fst_divmod [simp]:
-  "fst (divmod_integer k l) = k div l"
-  by (simp add: divmod_integer_def)
-
-lemma snd_divmod [simp]:
-  "snd (divmod_integer k l) = k mod l"
-  by (simp add: divmod_integer_def)
-
-definition divmod_abs :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
-where
-  "divmod_abs k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)"
-
-lemma fst_divmod_abs [simp]:
-  "fst (divmod_abs k l) = \<bar>k\<bar> div \<bar>l\<bar>"
-  by (simp add: divmod_abs_def)
-
-lemma snd_divmod_abs [simp]:
-  "snd (divmod_abs k l) = \<bar>k\<bar> mod \<bar>l\<bar>"
-  by (simp add: divmod_abs_def)
-
-lemma divmod_abs_terminate_code [code]:
-  "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)"
-  "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)"
-  "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)"
-  "divmod_abs j 0 = (0, \<bar>j\<bar>)"
-  "divmod_abs 0 j = (0, 0)"
-  by (simp_all add: prod_eq_iff)
-
-lemma divmod_abs_rec_code [code]:
-  "divmod_abs (Pos k) (Pos l) =
-    (let j = sub k l in
-       if j < 0 then (0, Pos k)
-       else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))"
-  apply (simp add: prod_eq_iff Let_def prod_case_beta)
-  apply transfer
-  apply (simp add: sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq)
-  done
-
-lemma divmod_integer_code [code]:
-  "divmod_integer k l =
-    (if k = 0 then (0, 0) else if l = 0 then (0, k) else
-    (apsnd \<circ> times \<circ> sgn) l (if sgn k = sgn l
-      then divmod_abs k l
-      else (let (r, s) = divmod_abs k l in
-        if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))"
-proof -
-  have aux1: "\<And>k l::int. sgn k = sgn l \<longleftrightarrow> k = 0 \<and> l = 0 \<or> 0 < l \<and> 0 < k \<or> l < 0 \<and> k < 0"
-    by (auto simp add: sgn_if)
-  have aux2: "\<And>q::int. - int_of_integer k = int_of_integer l * q \<longleftrightarrow> int_of_integer k = int_of_integer l * - q" by auto
-  show ?thesis
-    by (simp add: prod_eq_iff integer_eq_iff prod_case_beta aux1)
-      (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if div_minus_right mod_minus_right aux2)
-qed
-
-lemma div_integer_code [code]:
-  "k div l = fst (divmod_integer k l)"
-  by simp
-
-lemma mod_integer_code [code]:
-  "k mod l = snd (divmod_integer k l)"
-  by simp
-
-lemma equal_integer_code [code]:
-  "HOL.equal 0 (0::integer) \<longleftrightarrow> True"
-  "HOL.equal 0 (Pos l) \<longleftrightarrow> False"
-  "HOL.equal 0 (Neg l) \<longleftrightarrow> False"
-  "HOL.equal (Pos k) 0 \<longleftrightarrow> False"
-  "HOL.equal (Pos k) (Pos l) \<longleftrightarrow> HOL.equal k l"
-  "HOL.equal (Pos k) (Neg l) \<longleftrightarrow> False"
-  "HOL.equal (Neg k) 0 \<longleftrightarrow> False"
-  "HOL.equal (Neg k) (Pos l) \<longleftrightarrow> False"
-  "HOL.equal (Neg k) (Neg l) \<longleftrightarrow> HOL.equal k l"
-  by (simp_all add: equal)
-
-lemma equal_integer_refl [code nbe]:
-  "HOL.equal (k::integer) k \<longleftrightarrow> True"
-  by (fact equal_refl)
-
-lemma less_eq_integer_code [code]:
-  "0 \<le> (0::integer) \<longleftrightarrow> True"
-  "0 \<le> Pos l \<longleftrightarrow> True"
-  "0 \<le> Neg l \<longleftrightarrow> False"
-  "Pos k \<le> 0 \<longleftrightarrow> False"
-  "Pos k \<le> Pos l \<longleftrightarrow> k \<le> l"
-  "Pos k \<le> Neg l \<longleftrightarrow> False"
-  "Neg k \<le> 0 \<longleftrightarrow> True"
-  "Neg k \<le> Pos l \<longleftrightarrow> True"
-  "Neg k \<le> Neg l \<longleftrightarrow> l \<le> k"
-  by simp_all
-
-lemma less_integer_code [code]:
-  "0 < (0::integer) \<longleftrightarrow> False"
-  "0 < Pos l \<longleftrightarrow> True"
-  "0 < Neg l \<longleftrightarrow> False"
-  "Pos k < 0 \<longleftrightarrow> False"
-  "Pos k < Pos l \<longleftrightarrow> k < l"
-  "Pos k < Neg l \<longleftrightarrow> False"
-  "Neg k < 0 \<longleftrightarrow> True"
-  "Neg k < Pos l \<longleftrightarrow> True"
-  "Neg k < Neg l \<longleftrightarrow> l < k"
-  by simp_all
-
-lift_definition integer_of_num :: "num \<Rightarrow> integer"
-  is "numeral :: num \<Rightarrow> int"
-  .
-
-lemma integer_of_num [code]:
-  "integer_of_num num.One = 1"
-  "integer_of_num (num.Bit0 n) = (let k = integer_of_num n in k + k)"
-  "integer_of_num (num.Bit1 n) = (let k = integer_of_num n in k + k + 1)"
-  by (transfer, simp only: numeral.simps Let_def)+
-
-lift_definition num_of_integer :: "integer \<Rightarrow> num"
-  is "num_of_nat \<circ> nat"
-  .
-
-lemma num_of_integer_code [code]:
-  "num_of_integer k = (if k \<le> 1 then Num.One
-     else let
-       (l, j) = divmod_integer k 2;
-       l' = num_of_integer l;
-       l'' = l' + l'
-     in if j = 0 then l'' else l'' + Num.One)"
-proof -
-  {
-    assume "int_of_integer k mod 2 = 1"
-    then have "nat (int_of_integer k mod 2) = nat 1" by simp
-    moreover assume *: "1 < int_of_integer k"
-    ultimately have **: "nat (int_of_integer k) mod 2 = 1" by (simp add: nat_mod_distrib)
-    have "num_of_nat (nat (int_of_integer k)) =
-      num_of_nat (2 * (nat (int_of_integer k) div 2) + nat (int_of_integer k) mod 2)"
-      by simp
-    then have "num_of_nat (nat (int_of_integer k)) =
-      num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + nat (int_of_integer k) mod 2)"
-      by (simp add: mult_2)
-    with ** have "num_of_nat (nat (int_of_integer k)) =
-      num_of_nat (nat (int_of_integer k) div 2 + nat (int_of_integer k) div 2 + 1)"
-      by simp
-  }
-  note aux = this
-  show ?thesis
-    by (auto simp add: num_of_integer_def nat_of_integer_def Let_def prod_case_beta
-      not_le integer_eq_iff less_eq_integer_def
-      nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib
-       mult_2 [where 'a=nat] aux add_One)
-qed
-
-lemma nat_of_integer_code [code]:
-  "nat_of_integer k = (if k \<le> 0 then 0
-     else let
-       (l, j) = divmod_integer k 2;
-       l' = nat_of_integer l;
-       l'' = l' + l'
-     in if j = 0 then l'' else l'' + 1)"
-proof -
-  obtain j where "k = integer_of_int j"
-  proof
-    show "k = integer_of_int (int_of_integer k)" by simp
-  qed
-  moreover have "2 * (j div 2) = j - j mod 2"
-    by (simp add: zmult_div_cancel mult_commute)
-  ultimately show ?thesis
-    by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le
-      nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1)
-qed
-
-lemma int_of_integer_code [code]:
-  "int_of_integer k = (if k < 0 then - (int_of_integer (- k))
-     else if k = 0 then 0
-     else let
-       (l, j) = divmod_integer k 2;
-       l' = 2 * int_of_integer l
-     in if j = 0 then l' else l' + 1)"
-  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
-
-lemma integer_of_int_code [code]:
-  "integer_of_int k = (if k < 0 then - (integer_of_int (- k))
-     else if k = 0 then 0
-     else let
-       (l, j) = divmod_int k 2;
-       l' = 2 * integer_of_int l
-     in if j = 0 then l' else l' + 1)"
-  by (auto simp add: split_def Let_def integer_eq_iff zmult_div_cancel)
-
-hide_const (open) Pos Neg sub dup divmod_abs
-
-
-subsection {* Serializer setup for target language integers *}
-
-code_reserved Eval abs
-
-code_type integer
-  (SML "IntInf.int")
-  (OCaml "Big'_int.big'_int")
-  (Haskell "Integer")
-  (Scala "BigInt")
-  (Eval "int")
-
-code_instance integer :: equal
-  (Haskell -)
-
-code_const "0::integer"
-  (SML "0")
-  (OCaml "Big'_int.zero'_big'_int")
-  (Haskell "0")
-  (Scala "BigInt(0)")
-
-setup {*
-  fold (Numeral.add_code @{const_name Code_Numeral_Types.Pos}
-    false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
-*}
-
-setup {*
-  fold (Numeral.add_code @{const_name Code_Numeral_Types.Neg}
-    true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"]
-*}
-
-code_const "plus :: integer \<Rightarrow> _ \<Rightarrow> _"
-  (SML "IntInf.+ ((_), (_))")
-  (OCaml "Big'_int.add'_big'_int")
-  (Haskell infixl 6 "+")
-  (Scala infixl 7 "+")
-  (Eval infixl 8 "+")
-
-code_const "uminus :: integer \<Rightarrow> _"
-  (SML "IntInf.~")
-  (OCaml "Big'_int.minus'_big'_int")
-  (Haskell "negate")
-  (Scala "!(- _)")
-  (Eval "~/ _")
-
-code_const "minus :: integer \<Rightarrow> _"
-  (SML "IntInf.- ((_), (_))")
-  (OCaml "Big'_int.sub'_big'_int")
-  (Haskell infixl 6 "-")
-  (Scala infixl 7 "-")
-  (Eval infixl 8 "-")
-
-code_const Code_Numeral_Types.dup
-  (SML "IntInf.*/ (2,/ (_))")
-  (OCaml "Big'_int.mult'_big'_int/ (Big'_int.big'_int'_of'_int/ 2)")
-  (Haskell "!(2 * _)")
-  (Scala "!(2 * _)")
-  (Eval "!(2 * _)")
-
-code_const Code_Numeral_Types.sub
-  (SML "!(raise/ Fail/ \"sub\")")
-  (OCaml "failwith/ \"sub\"")
-  (Haskell "error/ \"sub\"")
-  (Scala "!sys.error(\"sub\")")
-
-code_const "times :: integer \<Rightarrow> _ \<Rightarrow> _"
-  (SML "IntInf.* ((_), (_))")
-  (OCaml "Big'_int.mult'_big'_int")
-  (Haskell infixl 7 "*")
-  (Scala infixl 8 "*")
-  (Eval infixl 9 "*")
-
-code_const Code_Numeral_Types.divmod_abs
-  (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)")
-  (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
-  (Haskell "divMod/ (abs _)/ (abs _)")
-  (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
-  (Eval "Integer.div'_mod/ (abs _)/ (abs _)")
-
-code_const "HOL.equal :: integer \<Rightarrow> _ \<Rightarrow> bool"
-  (SML "!((_ : IntInf.int) = _)")
-  (OCaml "Big'_int.eq'_big'_int")
-  (Haskell infix 4 "==")
-  (Scala infixl 5 "==")
-  (Eval infixl 6 "=")
-
-code_const "less_eq :: integer \<Rightarrow> _ \<Rightarrow> bool"
-  (SML "IntInf.<= ((_), (_))")
-  (OCaml "Big'_int.le'_big'_int")
-  (Haskell infix 4 "<=")
-  (Scala infixl 4 "<=")
-  (Eval infixl 6 "<=")
-
-code_const "less :: integer \<Rightarrow> _ \<Rightarrow> bool"
-  (SML "IntInf.< ((_), (_))")
-  (OCaml "Big'_int.lt'_big'_int")
-  (Haskell infix 4 "<")
-  (Scala infixl 4 "<")
-  (Eval infixl 6 "<")
-
-code_modulename SML
-  Code_Numeral_Types Arith
-
-code_modulename OCaml
-  Code_Numeral_Types Arith
-
-code_modulename Haskell
-  Code_Numeral_Types Arith
-
-
-subsection {* Type of target language naturals *}
-
-typedef natural = "UNIV \<Colon> nat set"
-  morphisms nat_of_natural natural_of_nat ..
-
-setup_lifting (no_code) type_definition_natural
-
-lemma natural_eq_iff [termination_simp]:
-  "m = n \<longleftrightarrow> nat_of_natural m = nat_of_natural n"
-  by transfer rule
-
-lemma natural_eqI:
-  "nat_of_natural m = nat_of_natural n \<Longrightarrow> m = n"
-  using natural_eq_iff [of m n] by simp
-
-lemma nat_of_natural_of_nat_inverse [simp]:
-  "nat_of_natural (natural_of_nat n) = n"
-  by transfer rule
-
-lemma natural_of_nat_of_natural_inverse [simp]:
-  "natural_of_nat (nat_of_natural n) = n"
-  by transfer rule
-
-instantiation natural :: "{comm_monoid_diff, semiring_1}"
-begin
-
-lift_definition zero_natural :: natural
-  is "0 :: nat"
-  .
-
-declare zero_natural.rep_eq [simp]
-
-lift_definition one_natural :: natural
-  is "1 :: nat"
-  .
-
-declare one_natural.rep_eq [simp]
-
-lift_definition plus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
-  is "plus :: nat \<Rightarrow> nat \<Rightarrow> nat"
-  .
-
-declare plus_natural.rep_eq [simp]
-
-lift_definition minus_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
-  is "minus :: nat \<Rightarrow> nat \<Rightarrow> nat"
-  .
-
-declare minus_natural.rep_eq [simp]
-
-lift_definition times_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
-  is "times :: nat \<Rightarrow> nat \<Rightarrow> nat"
-  .
-
-declare times_natural.rep_eq [simp]
-
-instance proof
-qed (transfer, simp add: algebra_simps)+
-
-end
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_natural (\<lambda>n::nat. n) (of_nat :: nat \<Rightarrow> natural)"
-proof -
-  have "fun_rel HOL.eq cr_natural (of_nat :: nat \<Rightarrow> nat) (of_nat :: nat \<Rightarrow> natural)"
-    by (unfold of_nat_def [abs_def]) transfer_prover
-  then show ?thesis by (simp add: id_def)
-qed
-
-lemma [transfer_rule]:
-  "fun_rel HOL.eq cr_natural (numeral :: num \<Rightarrow> nat) (numeral :: num \<Rightarrow> natural)"
-proof -
-  have "fun_rel HOL.eq cr_natural (numeral :: num \<Rightarrow> nat) (\<lambda>n. of_nat (numeral n))"
-    by transfer_prover
-  then show ?thesis by simp
-qed
-
-lemma nat_of_natural_of_nat [simp]:
-  "nat_of_natural (of_nat n) = n"
-  by transfer rule
-
-lemma natural_of_nat_of_nat [simp, code_abbrev]:
-  "natural_of_nat = of_nat"
-  by transfer rule
-
-lemma of_nat_of_natural [simp]:
-  "of_nat (nat_of_natural n) = n"
-  by transfer rule
-
-lemma nat_of_natural_numeral [simp]:
-  "nat_of_natural (numeral k) = numeral k"
-  by transfer rule
-
-instantiation natural :: "{semiring_div, equal, linordered_semiring}"
-begin
-
-lift_definition div_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
-  is "Divides.div :: nat \<Rightarrow> nat \<Rightarrow> nat"
-  .
-
-declare div_natural.rep_eq [simp]
-
-lift_definition mod_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
-  is "Divides.mod :: nat \<Rightarrow> nat \<Rightarrow> nat"
-  .
-
-declare mod_natural.rep_eq [simp]
-
-lift_definition less_eq_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
-  is "less_eq :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  .
-
-declare less_eq_natural.rep_eq [termination_simp]
-
-lift_definition less_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
-  is "less :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  .
-
-declare less_natural.rep_eq [termination_simp]
-
-lift_definition equal_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
-  is "HOL.equal :: nat \<Rightarrow> nat \<Rightarrow> bool"
-  .
-
-instance proof
-qed (transfer, simp add: algebra_simps equal less_le_not_le [symmetric] linear)+
-
-end
-
-lemma [transfer_rule]:
-  "fun_rel cr_natural (fun_rel cr_natural cr_natural) (min :: _ \<Rightarrow> _ \<Rightarrow> nat) (min :: _ \<Rightarrow> _ \<Rightarrow> natural)"
-  by (unfold min_def [abs_def]) transfer_prover
-
-lemma [transfer_rule]:
-  "fun_rel cr_natural (fun_rel cr_natural cr_natural) (max :: _ \<Rightarrow> _ \<Rightarrow> nat) (max :: _ \<Rightarrow> _ \<Rightarrow> natural)"
-  by (unfold max_def [abs_def]) transfer_prover
-
-lemma nat_of_natural_min [simp]:
-  "nat_of_natural (min k l) = min (nat_of_natural k) (nat_of_natural l)"
-  by transfer rule
-
-lemma nat_of_natural_max [simp]:
-  "nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"
-  by transfer rule
-
-lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
-  is "nat :: int \<Rightarrow> nat"
-  .
-
-lift_definition integer_of_natural :: "natural \<Rightarrow> integer"
-  is "of_nat :: nat \<Rightarrow> int"
-  .
-
-lemma natural_of_integer_of_natural [simp]:
-  "natural_of_integer (integer_of_natural n) = n"
-  by transfer simp
-
-lemma integer_of_natural_of_integer [simp]:
-  "integer_of_natural (natural_of_integer k) = max 0 k"
-  by transfer auto
-
-lemma int_of_integer_of_natural [simp]:
-  "int_of_integer (integer_of_natural n) = of_nat (nat_of_natural n)"
-  by transfer rule
-
-lemma integer_of_natural_of_nat [simp]:
-  "integer_of_natural (of_nat n) = of_nat n"
-  by transfer rule
-
-lemma [measure_function]:
-  "is_measure nat_of_natural"
-  by (rule is_measure_trivial)
-
-
-subsection {* Inductive represenation of target language naturals *}
-
-lift_definition Suc :: "natural \<Rightarrow> natural"
-  is Nat.Suc
-  .
-
-declare Suc.rep_eq [simp]
-
-rep_datatype "0::natural" Suc
-  by (transfer, fact nat.induct nat.inject nat.distinct)+
-
-lemma natural_case [case_names nat, cases type: natural]:
-  fixes m :: natural
-  assumes "\<And>n. m = of_nat n \<Longrightarrow> P"
-  shows P
-  using assms by transfer blast
-
-lemma [simp, code]:
-  "natural_size = nat_of_natural"
-proof (rule ext)
-  fix n
-  show "natural_size n = nat_of_natural n"
-    by (induct n) simp_all
-qed
-
-lemma [simp, code]:
-  "size = nat_of_natural"
-proof (rule ext)
-  fix n
-  show "size n = nat_of_natural n"
-    by (induct n) simp_all
-qed
-
-lemma natural_decr [termination_simp]:
-  "n \<noteq> 0 \<Longrightarrow> nat_of_natural n - Nat.Suc 0 < nat_of_natural n"
-  by transfer simp
-
-lemma natural_zero_minus_one:
-  "(0::natural) - 1 = 0"
-  by simp
-
-lemma Suc_natural_minus_one:
-  "Suc n - 1 = n"
-  by transfer simp
-
-hide_const (open) Suc
-
-
-subsection {* Code refinement for target language naturals *}
-
-lift_definition Nat :: "integer \<Rightarrow> natural"
-  is nat
-  .
-
-lemma [code_post]:
-  "Nat 0 = 0"
-  "Nat 1 = 1"
-  "Nat (numeral k) = numeral k"
-  by (transfer, simp)+
-
-lemma [code abstype]:
-  "Nat (integer_of_natural n) = n"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural (natural_of_nat n) = of_nat n"
-  by simp
-
-lemma [code abstract]:
-  "integer_of_natural (natural_of_integer k) = max 0 k"
-  by simp
-
-lemma [code_abbrev]:
-  "natural_of_integer (Code_Numeral_Types.Pos k) = numeral k"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural 0 = 0"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural 1 = 1"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural (Code_Numeral_Types.Suc n) = integer_of_natural n + 1"
-  by transfer simp
-
-lemma [code]:
-  "nat_of_natural = nat_of_integer \<circ> integer_of_natural"
-  by transfer (simp add: fun_eq_iff)
-
-lemma [code, code_unfold]:
-  "natural_case f g n = (if n = 0 then f else g (n - 1))"
-  by (cases n rule: natural.exhaust) (simp_all, simp add: Suc_def)
-
-declare natural.recs [code del]
-
-lemma [code abstract]:
-  "integer_of_natural (m + n) = integer_of_natural m + integer_of_natural n"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural (m - n) = max 0 (integer_of_natural m - integer_of_natural n)"
-  by transfer simp
-
-lemma [code abstract]:
-  "integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"
-  by transfer (simp add: of_nat_mult)
-
-lemma [code abstract]:
-  "integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"
-  by transfer (simp add: zdiv_int)
-
-lemma [code abstract]:
-  "integer_of_natural (m mod n) = integer_of_natural m mod integer_of_natural n"
-  by transfer (simp add: zmod_int)
-
-lemma [code]:
-  "HOL.equal m n \<longleftrightarrow> HOL.equal (integer_of_natural m) (integer_of_natural n)"
-  by transfer (simp add: equal)
-
-lemma [code nbe]:
-  "HOL.equal n (n::natural) \<longleftrightarrow> True"
-  by (simp add: equal)
-
-lemma [code]:
-  "m \<le> n \<longleftrightarrow> integer_of_natural m \<le> integer_of_natural n"
-  by transfer simp
-
-lemma [code]:
-  "m < n \<longleftrightarrow> integer_of_natural m < integer_of_natural n"
-  by transfer simp
-
-hide_const (open) Nat
-
-
-code_reflect Code_Numeral_Types
-  datatypes natural = _
-  functions integer_of_natural natural_of_integer
-
-end
-
--- a/src/HOL/Library/Code_Real_Approx_By_Float.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Real_Approx_By_Float.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1,7 +1,7 @@
 (* Authors: Florian Haftmann, Johannes Hölzl, Tobias Nipkow *)
 
 theory Code_Real_Approx_By_Float
-imports Complex_Main "~~/src/HOL/Library/Code_Integer"
+imports Complex_Main "~~/src/HOL/Library/Code_Target_Int"
 begin
 
 text{* \textbf{WARNING} This theory implements mathematical reals by machine
@@ -119,15 +119,19 @@
   (OCaml "Pervasives.asin")
 declare arcsin_def[code del]
 
-definition real_of_int :: "int \<Rightarrow> real" where
-  "real_of_int \<equiv> of_int"
+definition real_of_integer :: "integer \<Rightarrow> real" where
+  "real_of_integer = of_int \<circ> int_of_integer"
 
-code_const real_of_int
+code_const real_of_integer
   (SML "Real.fromInt")
   (OCaml "Pervasives.float (Big'_int.int'_of'_big'_int (_))")
 
-lemma of_int_eq_real_of_int[code_unfold]: "of_int = real_of_int"
-  unfolding real_of_int_def ..
+definition real_of_int :: "int \<Rightarrow> real" where
+  [code_abbrev]: "real_of_int = of_int"
+
+lemma [code]:
+  "real_of_int = real_of_integer \<circ> integer_of_int"
+  by (simp add: fun_eq_iff real_of_integer_def real_of_int_def)
 
 lemma [code_unfold del]:
   "0 \<equiv> (of_rat 0 :: real)"
@@ -155,3 +159,4 @@
 end
 
 end
+
--- a/src/HOL/Library/Code_Target_Int.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Target_Int.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* Implementation of integer numbers by target-language integers *}
 
 theory Code_Target_Int
-imports Main "~~/src/HOL/Library/Code_Numeral_Types"
+imports Main
 begin
 
 code_datatype int_of_integer
@@ -54,7 +54,7 @@
   by transfer simp
 
 lemma [code]:
-  "Int.dup k = int_of_integer (Code_Numeral_Types.dup (of_int k))"
+  "Int.dup k = int_of_integer (Code_Numeral.dup (of_int k))"
   by transfer simp
 
 lemma [code, code del]:
@@ -66,7 +66,7 @@
 
 lemma [code]:
   "pdivmod k l = map_pair int_of_integer int_of_integer
-    (Code_Numeral_Types.divmod_abs (of_int k) (of_int l))"
+    (Code_Numeral.divmod_abs (of_int k) (of_int l))"
   by (simp add: prod_eq_iff pdivmod_def)
 
 lemma [code]:
--- a/src/HOL/Library/Code_Target_Nat.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Target_Nat.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* Implementation of natural numbers by target-language integers *}
 
 theory Code_Target_Nat
-imports Code_Abstract_Nat Code_Numeral_Types
+imports Code_Abstract_Nat
 begin
 
 subsection {* Implementation for @{typ nat} *}
--- a/src/HOL/Library/DAList.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/DAList.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -136,7 +136,7 @@
 instantiation alist :: (exhaustive, exhaustive) exhaustive
 begin
 
-fun exhaustive_alist :: "(('a, 'b) alist => (bool * term list) option) => code_numeral => (bool * term list) option"
+fun exhaustive_alist :: "(('a, 'b) alist => (bool * term list) option) => natural => (bool * term list) option"
 where
   "exhaustive_alist f i = (if i = 0 then None else case f empty of Some ts => Some ts | None =>
      exhaustive_alist (%a. Quickcheck_Exhaustive.exhaustive (%k. Quickcheck_Exhaustive.exhaustive (%v. f (update k v a)) (i - 1)) (i - 1)) (i - 1))"
@@ -148,7 +148,7 @@
 instantiation alist :: (full_exhaustive, full_exhaustive) full_exhaustive
 begin
 
-fun full_exhaustive_alist :: "(('a, 'b) alist * (unit => term) => (bool * term list) option) => code_numeral => (bool * term list) option"
+fun full_exhaustive_alist :: "(('a, 'b) alist * (unit => term) => (bool * term list) option) => natural => (bool * term list) option"
 where
   "full_exhaustive_alist f i = (if i = 0 then None else case f valterm_empty of Some ts => Some ts | None =>
      full_exhaustive_alist (%a. Quickcheck_Exhaustive.full_exhaustive (%k. Quickcheck_Exhaustive.full_exhaustive (%v. f (valterm_update k v a)) (i - 1)) (i - 1)) (i - 1))"
--- a/src/HOL/Library/Efficient_Nat.thy	Fri Feb 15 08:31:30 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,305 +0,0 @@
-(*  Title:      HOL/Library/Efficient_Nat.thy
-    Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
-*)
-
-header {* Implementation of natural numbers by target-language integers *}
-
-theory Efficient_Nat
-imports Code_Binary_Nat Code_Integer Main
-begin
-
-text {*
-  The efficiency of the generated code for natural numbers can be improved
-  drastically by implementing natural numbers by target-language
-  integers.  To do this, just include this theory.
-*}
-
-subsection {* Target language fundamentals *}
-
-text {*
-  For ML, we map @{typ nat} to target language integers, where we
-  ensure that values are always non-negative.
-*}
-
-code_type nat
-  (SML "IntInf.int")
-  (OCaml "Big'_int.big'_int")
-  (Eval "int")
-
-text {*
-  For Haskell and Scala we define our own @{typ nat} type.  The reason
-  is that we have to distinguish type class instances for @{typ nat}
-  and @{typ int}.
-*}
-
-code_include Haskell "Nat"
-{*newtype Nat = Nat Integer deriving (Eq, Show, Read);
-
-instance Num Nat where {
-  fromInteger k = Nat (if k >= 0 then k else 0);
-  Nat n + Nat m = Nat (n + m);
-  Nat n - Nat m = fromInteger (n - m);
-  Nat n * Nat m = Nat (n * m);
-  abs n = n;
-  signum _ = 1;
-  negate n = error "negate Nat";
-};
-
-instance Ord Nat where {
-  Nat n <= Nat m = n <= m;
-  Nat n < Nat m = n < m;
-};
-
-instance Real Nat where {
-  toRational (Nat n) = toRational n;
-};
-
-instance Enum Nat where {
-  toEnum k = fromInteger (toEnum k);
-  fromEnum (Nat n) = fromEnum n;
-};
-
-instance Integral Nat where {
-  toInteger (Nat n) = n;
-  divMod n m = quotRem n m;
-  quotRem (Nat n) (Nat m)
-    | (m == 0) = (0, Nat n)
-    | otherwise = (Nat k, Nat l) where (k, l) = quotRem n m;
-};
-*}
-
-code_reserved Haskell Nat
-
-code_include Scala "Nat"
-{*object Nat {
-
-  def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
-  def apply(numeral: Int): Nat = Nat(BigInt(numeral))
-  def apply(numeral: String): Nat = Nat(BigInt(numeral))
-
-}
-
-class Nat private(private val value: BigInt) {
-
-  override def hashCode(): Int = this.value.hashCode()
-
-  override def equals(that: Any): Boolean = that match {
-    case that: Nat => this equals that
-    case _ => false
-  }
-
-  override def toString(): String = this.value.toString
-
-  def equals(that: Nat): Boolean = this.value == that.value
-
-  def as_BigInt: BigInt = this.value
-  def as_Int: Int = if (this.value >= scala.Int.MinValue && this.value <= scala.Int.MaxValue)
-      this.value.intValue
-    else error("Int value out of range: " + this.value.toString)
-
-  def +(that: Nat): Nat = new Nat(this.value + that.value)
-  def -(that: Nat): Nat = Nat(this.value - that.value)
-  def *(that: Nat): Nat = new Nat(this.value * that.value)
-
-  def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
-    else {
-      val (k, l) = this.value /% that.value
-      (new Nat(k), new Nat(l))
-    }
-
-  def <=(that: Nat): Boolean = this.value <= that.value
-
-  def <(that: Nat): Boolean = this.value < that.value
-
-}
-*}
-
-code_reserved Scala Nat
-
-code_type nat
-  (Haskell "Nat.Nat")
-  (Scala "Nat")
-
-code_instance nat :: equal
-  (Haskell -)
-
-setup {*
-  fold (Numeral.add_code @{const_name nat_of_num}
-    false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell", "Scala"]
-*}
-
-code_const "0::nat"
-  (SML "0")
-  (OCaml "Big'_int.zero'_big'_int")
-  (Haskell "0")
-  (Scala "Nat(0)")
-
-
-subsection {* Conversions *}
-
-text {*
-  Since natural numbers are implemented
-  using integers in ML, the coercion function @{term "int"} of type
-  @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
-  For the @{const nat} function for converting an integer to a natural
-  number, we give a specific implementation using an ML expression that
-  returns its input value, provided that it is non-negative, and otherwise
-  returns @{text "0"}.
-*}
-
-definition int :: "nat \<Rightarrow> int" where
-  [code_abbrev]: "int = of_nat"
-
-code_const int
-  (SML "_")
-  (OCaml "_")
-
-code_const nat
-  (SML "IntInf.max/ (0,/ _)")
-  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
-  (Eval "Integer.max/ 0")
-
-text {* For Haskell and Scala, things are slightly different again. *}
-
-code_const int and nat
-  (Haskell "Prelude.toInteger" and "Prelude.fromInteger")
-  (Scala "!_.as'_BigInt" and "Nat")
-
-text {* Alternativ implementation for @{const of_nat} *}
-
-lemma [code]:
-  "of_nat n = (if n = 0 then 0 else
-     let
-       (q, m) = divmod_nat n 2;
-       q' = 2 * of_nat q
-     in if m = 0 then q' else q' + 1)"
-proof -
-  from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp
-  show ?thesis
-    apply (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat
-      of_nat_mult
-      of_nat_add [symmetric])
-    apply (auto simp add: of_nat_mult)
-    apply (simp add: * of_nat_mult add_commute mult_commute)
-    done
-qed
-
-text {* Conversion from and to code numerals *}
-
-code_const Code_Numeral.of_nat
-  (SML "IntInf.toInt")
-  (OCaml "_")
-  (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
-  (Scala "!Natural(_.as'_BigInt)")
-  (Eval "_")
-
-code_const Code_Numeral.nat_of
-  (SML "IntInf.fromInt")
-  (OCaml "_")
-  (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
-  (Scala "!Nat(_.as'_BigInt)")
-  (Eval "_")
-
-
-subsection {* Target language arithmetic *}
-
-code_const "plus \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
-  (SML "IntInf.+/ ((_),/ (_))")
-  (OCaml "Big'_int.add'_big'_int")
-  (Haskell infixl 6 "+")
-  (Scala infixl 7 "+")
-  (Eval infixl 8 "+")
-
-code_const "minus \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
-  (SML "IntInf.max/ (0, IntInf.-/ ((_),/ (_)))")
-  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)")
-  (Haskell infixl 6 "-")
-  (Scala infixl 7 "-")
-  (Eval "Integer.max/ 0/ (_ -/ _)")
-
-code_const Code_Binary_Nat.dup
-  (SML "IntInf.*/ (2,/ (_))")
-  (OCaml "Big'_int.mult'_big'_int/ 2")
-  (Haskell "!(2 * _)")
-  (Scala "!(2 * _)")
-  (Eval "!(2 * _)")
-
-code_const Code_Binary_Nat.sub
-  (SML "!(raise/ Fail/ \"sub\")")
-  (OCaml "failwith/ \"sub\"")
-  (Haskell "error/ \"sub\"")
-  (Scala "!sys.error(\"sub\")")
-
-code_const "times \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
-  (SML "IntInf.*/ ((_),/ (_))")
-  (OCaml "Big'_int.mult'_big'_int")
-  (Haskell infixl 7 "*")
-  (Scala infixl 8 "*")
-  (Eval infixl 9 "*")
-
-code_const divmod_nat
-  (SML "IntInf.divMod/ ((_),/ (_))")
-  (OCaml "Big'_int.quomod'_big'_int")
-  (Haskell "divMod")
-  (Scala infixl 8 "/%")
-  (Eval "Integer.div'_mod")
-
-code_const "HOL.equal \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
-  (SML "!((_ : IntInf.int) = _)")
-  (OCaml "Big'_int.eq'_big'_int")
-  (Haskell infix 4 "==")
-  (Scala infixl 5 "==")
-  (Eval infixl 6 "=")
-
-code_const "less_eq \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
-  (SML "IntInf.<=/ ((_),/ (_))")
-  (OCaml "Big'_int.le'_big'_int")
-  (Haskell infix 4 "<=")
-  (Scala infixl 4 "<=")
-  (Eval infixl 6 "<=")
-
-code_const "less \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
-  (SML "IntInf.</ ((_),/ (_))")
-  (OCaml "Big'_int.lt'_big'_int")
-  (Haskell infix 4 "<")
-  (Scala infixl 4 "<")
-  (Eval infixl 6 "<")
-
-code_const Num.num_of_nat
-  (SML "!(raise/ Fail/ \"num'_of'_nat\")")
-  (OCaml "failwith/ \"num'_of'_nat\"")
-  (Haskell "error/ \"num'_of'_nat\"")
-  (Scala "!sys.error(\"num'_of'_nat\")")
-
-
-subsection {* Evaluation *}
-
-lemma [code, code del]:
-  "(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..
-
-code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"
-  (SML "HOLogic.mk'_number/ HOLogic.natT")
-
-text {*
-  Evaluation with @{text "Quickcheck_Narrowing"} does not work yet,
-  since a couple of questions how to perform evaluations in Haskell are not that
-  clear yet.  Therefore, we simply deactivate the narrowing-based quickcheck
-  from here on.
-*}
-
-declare [[quickcheck_narrowing_active = false]] 
-
-
-code_modulename SML
-  Efficient_Nat Arith
-
-code_modulename OCaml
-  Efficient_Nat Arith
-
-code_modulename Haskell
-  Efficient_Nat Arith
-
-hide_const (open) int
-
-end
-
--- a/src/HOL/Library/IArray.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/IArray.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -46,34 +46,34 @@
 code_const IArray
   (SML "Vector.fromList")
 
-primrec tabulate :: "code_numeral \<times> (code_numeral \<Rightarrow> 'a) \<Rightarrow> 'a iarray" where
-"tabulate (n, f) = IArray (map (f \<circ> Code_Numeral.of_nat) [0..<Code_Numeral.nat_of n])"
+primrec tabulate :: "integer \<times> (integer \<Rightarrow> 'a) \<Rightarrow> 'a iarray" where
+"tabulate (n, f) = IArray (map (f \<circ> integer_of_nat) [0..<nat_of_integer n])"
 hide_const (open) tabulate
 
 lemma [code]:
-"IArray.of_fun f n = IArray.tabulate (Code_Numeral.of_nat n, f \<circ> Code_Numeral.nat_of)"
+"IArray.of_fun f n = IArray.tabulate (integer_of_nat n, f \<circ> nat_of_integer)"
 by simp 
 
 code_const IArray.tabulate
   (SML "Vector.tabulate")
 
-primrec sub' :: "'a iarray \<times> code_numeral \<Rightarrow> 'a" where
-"sub' (as, n) = IArray.list_of as ! Code_Numeral.nat_of n"
+primrec sub' :: "'a iarray \<times> integer \<Rightarrow> 'a" where
+"sub' (as, n) = IArray.list_of as ! nat_of_integer n"
 hide_const (open) sub'
 
 lemma [code]:
-"as !! n = IArray.sub' (as, Code_Numeral.of_nat n)"
+"as !! n = IArray.sub' (as, integer_of_nat n)"
 by simp
 
 code_const IArray.sub'
   (SML "Vector.sub")
 
-definition length' :: "'a iarray \<Rightarrow> code_numeral" where
-[simp]: "length' as = Code_Numeral.of_nat (List.length (IArray.list_of as))"
+definition length' :: "'a iarray \<Rightarrow> integer" where
+[simp]: "length' as = integer_of_nat (List.length (IArray.list_of as))"
 hide_const (open) length'
 
 lemma [code]:
-"IArray.length as = Code_Numeral.nat_of (IArray.length' as)"
+"IArray.length as = nat_of_integer (IArray.length' as)"
 by simp
 
 code_const IArray.length'
--- a/src/HOL/Library/Multiset.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Multiset.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1326,7 +1326,7 @@
 instantiation multiset :: (exhaustive) exhaustive
 begin
 
-definition exhaustive_multiset :: "('a multiset => (bool * term list) option) => code_numeral => (bool * term list) option"
+definition exhaustive_multiset :: "('a multiset => (bool * term list) option) => natural => (bool * term list) option"
 where
   "exhaustive_multiset f i = Quickcheck_Exhaustive.exhaustive (%xs. f (Bag xs)) i"
 
@@ -1337,7 +1337,7 @@
 instantiation multiset :: (full_exhaustive) full_exhaustive
 begin
 
-definition full_exhaustive_multiset :: "('a multiset * (unit => term) => (bool * term list) option) => code_numeral => (bool * term list) option"
+definition full_exhaustive_multiset :: "('a multiset * (unit => term) => (bool * term list) option) => natural => (bool * term list) option"
 where
   "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (%xs. f (bagify xs)) i"
 
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -99,10 +99,10 @@
     end
   fun enumerate_addups_nat compfuns (_ : typ) =
     absdummy @{typ nat} (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ "nat * nat"}
-    (absdummy @{typ code_numeral} (@{term "Pair :: nat => nat => nat * nat"} $
-      (@{term "Code_Numeral.nat_of"} $ Bound 0) $
-      (@{term "op - :: nat => nat => nat"} $ Bound 1 $ (@{term "Code_Numeral.nat_of"} $ Bound 0))),
-      @{term "0 :: code_numeral"}, @{term "Code_Numeral.of_nat"} $ Bound 0))
+    (absdummy @{typ natural} (@{term "Pair :: nat => nat => nat * nat"} $
+      (@{term "nat_of_natural"} $ Bound 0) $
+      (@{term "op - :: nat => nat => nat"} $ Bound 1 $ (@{term "nat_of_natural"} $ Bound 0))),
+      @{term "0 :: natural"}, @{term "natural_of_nat"} $ Bound 0))
   fun enumerate_nats compfuns  (_ : typ) =
     let
       val (single_const, _) = strip_comb (Predicate_Compile_Aux.mk_single compfuns @{term "0 :: nat"})
@@ -111,8 +111,8 @@
       absdummy @{typ nat} (absdummy @{typ nat}
         (Const (@{const_name If}, @{typ bool} --> T --> T --> T) $
           (@{term "op = :: nat => nat => bool"} $ Bound 0 $ @{term "0::nat"}) $
-          (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ nat} (@{term "Code_Numeral.nat_of"},
-            @{term "0::code_numeral"}, @{term "Code_Numeral.of_nat"} $ Bound 1)) $
+          (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ nat} (@{term "nat_of_natural"},
+            @{term "0::natural"}, @{term "natural_of_nat"} $ Bound 1)) $
             (single_const $ (@{term "op + :: nat => nat => nat"} $ Bound 1 $ Bound 0))))
     end
 in
--- a/src/HOL/Library/Predicate_Compile_Quickcheck.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Predicate_Compile_Quickcheck.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -10,4 +10,4 @@
 
 setup {* Predicate_Compile_Quickcheck.setup *}
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Limited_Sequence.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Limited_Sequence.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -9,7 +9,7 @@
 
 subsection {* Depth-Limited Sequence *}
 
-type_synonym 'a dseq = "code_numeral \<Rightarrow> bool \<Rightarrow> 'a lazy_sequence option"
+type_synonym 'a dseq = "natural \<Rightarrow> bool \<Rightarrow> 'a lazy_sequence option"
 
 definition empty :: "'a dseq"
 where
@@ -19,11 +19,11 @@
 where
   "single x = (\<lambda>_ _. Some (Lazy_Sequence.single x))"
 
-definition eval :: "'a dseq \<Rightarrow> code_numeral \<Rightarrow> bool \<Rightarrow> 'a lazy_sequence option"
+definition eval :: "'a dseq \<Rightarrow> natural \<Rightarrow> bool \<Rightarrow> 'a lazy_sequence option"
 where
   [simp]: "eval f i pol = f i pol"
 
-definition yield :: "'a dseq \<Rightarrow> code_numeral \<Rightarrow> bool \<Rightarrow> ('a \<times> 'a dseq) option" 
+definition yield :: "'a dseq \<Rightarrow> natural \<Rightarrow> bool \<Rightarrow> ('a \<times> 'a dseq) option" 
 where
   "yield f i pol = (case eval f i pol of
     None \<Rightarrow> None
@@ -82,7 +82,7 @@
 
 subsection {* Positive Depth-Limited Sequence *}
 
-type_synonym 'a pos_dseq = "code_numeral \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
+type_synonym 'a pos_dseq = "natural \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
 
 definition pos_empty :: "'a pos_dseq"
 where
@@ -112,7 +112,7 @@
 where
   "pos_if_seq b = (if b then pos_single () else pos_empty)"
 
-definition pos_iterate_upto :: "(code_numeral \<Rightarrow> 'a) \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> 'a pos_dseq"
+definition pos_iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a pos_dseq"
 where
   "pos_iterate_upto f n m = (\<lambda>i. Lazy_Sequence.iterate_upto f n m)"
  
@@ -123,7 +123,7 @@
 
 subsection {* Negative Depth-Limited Sequence *}
 
-type_synonym 'a neg_dseq = "code_numeral \<Rightarrow> 'a Lazy_Sequence.hit_bound_lazy_sequence"
+type_synonym 'a neg_dseq = "natural \<Rightarrow> 'a Lazy_Sequence.hit_bound_lazy_sequence"
 
 definition neg_empty :: "'a neg_dseq"
 where
@@ -178,16 +178,16 @@
 ML {*
 signature LIMITED_SEQUENCE =
 sig
-  type 'a dseq = int -> bool -> 'a Lazy_Sequence.lazy_sequence option
+  type 'a dseq = Code_Numeral.natural -> bool -> 'a Lazy_Sequence.lazy_sequence option
   val map : ('a -> 'b) -> 'a dseq -> 'b dseq
-  val yield : 'a dseq -> int -> bool -> ('a * 'a dseq) option
-  val yieldn : int -> 'a dseq -> int -> bool -> 'a list * 'a dseq
+  val yield : 'a dseq -> Code_Numeral.natural -> bool -> ('a * 'a dseq) option
+  val yieldn : int -> 'a dseq -> Code_Numeral.natural -> bool -> 'a list * 'a dseq
 end;
 
 structure Limited_Sequence : LIMITED_SEQUENCE =
 struct
 
-type 'a dseq = int -> bool -> 'a Lazy_Sequence.lazy_sequence option
+type 'a dseq = Code_Numeral.natural -> bool -> 'a Lazy_Sequence.lazy_sequence option
 
 fun map f = @{code Limited_Sequence.map} f;
 
--- a/src/HOL/Num.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Num.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -545,7 +545,8 @@
 
 end
 
-lemma nat_of_num_numeral: "nat_of_num = numeral"
+lemma nat_of_num_numeral [code_abbrev]:
+  "nat_of_num = numeral"
 proof
   fix n
   have "numeral n = nat_of_num n"
@@ -553,6 +554,12 @@
   then show "nat_of_num n = numeral n" by simp
 qed
 
+lemma nat_of_num_code [code]:
+  "nat_of_num One = 1"
+  "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)"
+  "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))"
+  by (simp_all add: Let_def)
+
 subsubsection {*
   Equality: class @{text semiring_char_0}
 *}
@@ -1097,6 +1104,7 @@
 
 hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec
 
+
 subsection {* code module namespace *}
 
 code_modulename SML
@@ -1110,3 +1118,4 @@
 
 end
 
+
--- a/src/HOL/Predicate.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Predicate.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -679,14 +679,15 @@
 
 text {* Lazy Evaluation of an indexed function *}
 
-function iterate_upto :: "(code_numeral \<Rightarrow> 'a) \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> 'a Predicate.pred"
+function iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a Predicate.pred"
 where
   "iterate_upto f n m =
     Predicate.Seq (%u. if n > m then Predicate.Empty
      else Predicate.Insert (f n) (iterate_upto f (n + 1) m))"
 by pat_completeness auto
 
-termination by (relation "measure (%(f, n, m). Code_Numeral.nat_of (m + 1 - n))") auto
+termination by (relation "measure (%(f, n, m). nat_of_natural (m + 1 - n))")
+  (auto simp add: less_natural_def)
 
 text {* Misc *}
 
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -29,7 +29,7 @@
 thm True'.equation
 thm True'.dseq_equation
 thm True'.random_dseq_equation
-values [expected "{()}" ]"{x. True'}"
+values [expected "{()}"] "{x. True'}"
 values [expected "{}" dseq 0] "{x. True'}"
 values [expected "{()}" dseq 1] "{x. True'}"
 values [expected "{()}" dseq 2] "{x. True'}"
@@ -119,22 +119,22 @@
 
 code_pred nested_tuples .
 
-inductive JamesBond :: "nat => int => code_numeral => bool"
+inductive JamesBond :: "nat => int => natural => bool"
 where
   "JamesBond 0 0 7"
 
 code_pred JamesBond .
 
-values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
-values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
-values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
-values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
-values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
-values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
+values [expected "{(0::nat, 0::int , 7::natural)}"] "{(a, b, c). JamesBond a b c}"
+values [expected "{(0::nat, 7::natural, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
+values [expected "{(0::int, 0::nat, 7::natural)}"] "{(b, a, c). JamesBond a b c}"
+values [expected "{(0::int, 7::natural, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
+values [expected "{(7::natural, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
+values [expected "{(7::natural, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
 
-values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
-values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
-values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
+values [expected "{(7::natural, 0::int)}"] "{(a, b). JamesBond 0 b a}"
+values [expected "{(7::natural, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
+values [expected "{(0::nat, 7::natural)}"] "{(a, c). JamesBond a 0 c}"
 
 
 subsection {* Alternative Rules *}
--- a/src/HOL/Proofs/Extraction/Euclid.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Proofs/Extraction/Euclid.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -10,7 +10,7 @@
 imports
   "~~/src/HOL/Number_Theory/UniqueFactorization"
   Util
-  "~~/src/HOL/Library/Efficient_Nat"
+  "~~/src/HOL/Library/Code_Target_Numeral"
 begin
 
 text {*
--- a/src/HOL/Proofs/Extraction/Pigeonhole.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Proofs/Extraction/Pigeonhole.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* The pigeonhole principle *}
 
 theory Pigeonhole
-imports Util "~~/src/HOL/Library/Efficient_Nat"
+imports Util "~~/src/HOL/Library/Code_Target_Numeral"
 begin
 
 text {*
@@ -237,11 +237,11 @@
 end
 
 definition
-  "test n u = pigeonhole n (\<lambda>m. m - 1)"
+  "test n u = pigeonhole (nat_of_integer n) (\<lambda>m. m - 1)"
 definition
-  "test' n u = pigeonhole_slow n (\<lambda>m. m - 1)"
+  "test' n u = pigeonhole_slow (nat_of_integer n) (\<lambda>m. m - 1)"
 definition
-  "test'' u = pigeonhole 8 (op ! [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])"
+  "test'' u = pigeonhole 8 (List.nth [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])"
 
 ML "timeit (@{code test} 10)" 
 ML "timeit (@{code test'} 10)"
--- a/src/HOL/Proofs/Lambda/WeakNorm.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Proofs/Lambda/WeakNorm.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -6,7 +6,7 @@
 header {* Weak normalization for simply-typed lambda calculus *}
 
 theory WeakNorm
-imports LambdaType NormalForm "~~/src/HOL/Library/Code_Integer"
+imports LambdaType NormalForm "~~/src/HOL/Library/Code_Target_Int"
 begin
 
 text {*
@@ -387,14 +387,16 @@
 *}
 
 ML {*
+val nat_of_integer = @{code nat} o @{code int_of_integer};
+
 fun dBtype_of_typ (Type ("fun", [T, U])) =
       @{code Fun} (dBtype_of_typ T, dBtype_of_typ U)
   | dBtype_of_typ (TFree (s, _)) = (case raw_explode s of
-        ["'", a] => @{code Atom} (@{code nat} (ord a - 97))
+        ["'", a] => @{code Atom} (nat_of_integer (ord a - 97))
       | _ => error "dBtype_of_typ: variable name")
   | dBtype_of_typ _ = error "dBtype_of_typ: bad type";
 
-fun dB_of_term (Bound i) = @{code dB.Var} (@{code nat} i)
+fun dB_of_term (Bound i) = @{code dB.Var} (nat_of_integer i)
   | dB_of_term (t $ u) = @{code dB.App} (dB_of_term t, dB_of_term u)
   | dB_of_term (Abs (_, _, t)) = @{code dB.Abs} (dB_of_term t)
   | dB_of_term _ = error "dB_of_term: bad term";
@@ -402,23 +404,23 @@
 fun term_of_dB Ts (Type ("fun", [T, U])) (@{code dB.Abs} dBt) =
       Abs ("x", T, term_of_dB (T :: Ts) U dBt)
   | term_of_dB Ts _ dBt = term_of_dB' Ts dBt
-and term_of_dB' Ts (@{code dB.Var} n) = Bound (@{code int_of_nat} n)
+and term_of_dB' Ts (@{code dB.Var} n) = Bound (@{code integer_of_nat} n)
   | term_of_dB' Ts (@{code dB.App} (dBt, dBu)) =
       let val t = term_of_dB' Ts dBt
       in case fastype_of1 (Ts, t) of
-          Type ("fun", [T, U]) => t $ term_of_dB Ts T dBu
+          Type ("fun", [T, _]) => t $ term_of_dB Ts T dBu
         | _ => error "term_of_dB: function type expected"
       end
   | term_of_dB' _ _ = error "term_of_dB: term not in normal form";
 
 fun typing_of_term Ts e (Bound i) =
-      @{code Var} (e, @{code nat} i, dBtype_of_typ (nth Ts i))
+      @{code Var} (e, nat_of_integer i, dBtype_of_typ (nth Ts i))
   | typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of
         Type ("fun", [T, U]) => @{code App} (e, dB_of_term t,
           dBtype_of_typ T, dBtype_of_typ U, dB_of_term u,
           typing_of_term Ts e t, typing_of_term Ts e u)
       | _ => error "typing_of_term: function type expected")
-  | typing_of_term Ts e (Abs (s, T, t)) =
+  | typing_of_term Ts e (Abs (_, T, t)) =
       let val dBT = dBtype_of_typ T
       in @{code Abs} (e, dBT, dB_of_term t,
         dBtype_of_typ (fastype_of1 (T :: Ts, t)),
--- a/src/HOL/Quickcheck_Benchmark/Needham_Schroeder_Base.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Benchmark/Needham_Schroeder_Base.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -125,8 +125,8 @@
   [code]: "cps_of_set (set xs) f = find_first f xs"
 sorry
 
-consts pos_cps_of_set :: "'a set => ('a => (bool * term list) option) => code_numeral => (bool * term list) option"
-consts neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => code_numeral => term list Quickcheck_Exhaustive.three_valued"
+consts pos_cps_of_set :: "'a set => ('a => (bool * term list) option) => natural => (bool * term list) option"
+consts neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => natural => term list Quickcheck_Exhaustive.three_valued"
 
 lemma
   [code]: "pos_cps_of_set (set xs) f i = find_first f xs"
--- a/src/HOL/Quickcheck_Examples/Completeness.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Examples/Completeness.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -16,6 +16,10 @@
   "is_some (Some t) = True"
 | "is_some None = False"
 
+lemma is_some_is_not_None:
+  "is_some x \<longleftrightarrow> x \<noteq> None"
+  by (cases x) simp_all
+
 setup {* Exhaustive_Generators.setup_exhaustive_datatype_interpretation *} 
 
 subsection {* Defining the size of values *}
@@ -36,12 +40,12 @@
 
 end
 
-instantiation code_numeral :: size
+instantiation natural :: size
 begin
 
-definition size_code_numeral :: "code_numeral => nat"
+definition size_natural :: "natural => nat"
 where
-  "size_code_numeral = Code_Numeral.nat_of"
+  "size_natural = nat_of_natural"
 
 instance ..
 
@@ -74,96 +78,86 @@
 
 class complete = exhaustive + size +
 assumes
-   complete: "(\<exists> v. size v \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
+   complete: "(\<exists> v. size v \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
 
 lemma complete_imp1:
-  "size (v :: 'a :: complete) \<le> n \<Longrightarrow> is_some (f v) \<Longrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
+  "size (v :: 'a :: complete) \<le> n \<Longrightarrow> is_some (f v) \<Longrightarrow> is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
 by (metis complete)
 
 lemma complete_imp2:
-  assumes "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
+  assumes "is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
   obtains v where "size (v :: 'a :: complete) \<le> n" "is_some (f v)"
 using assms by (metis complete)
 
 subsubsection {* Instance Proofs *}
 
-declare exhaustive'.simps [simp del]
+declare exhaustive_int'.simps [simp del]
 
 lemma complete_exhaustive':
-  "(\<exists> i. j <= i & i <= k & is_some (f i)) \<longleftrightarrow> is_some (Quickcheck_Exhaustive.exhaustive' f k j)"
-proof (induct rule: Quickcheck_Exhaustive.exhaustive'.induct[of _ f k j])
+  "(\<exists> i. j <= i & i <= k & is_some (f i)) \<longleftrightarrow> is_some (Quickcheck_Exhaustive.exhaustive_int' f k j)"
+proof (induct rule: Quickcheck_Exhaustive.exhaustive_int'.induct[of _ f k j])
   case (1 f d i)
   show ?case
   proof (cases "f i")
     case None
     from this 1 show ?thesis
-    unfolding exhaustive'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def
+    unfolding exhaustive_int'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def
+    apply (auto simp add: add1_zle_eq dest: less_imp_le)
     apply auto
-    apply (metis is_some.simps(2) order_le_neq_trans zless_imp_add1_zle)
-    apply (metis add1_zle_eq less_int_def)
     done
   next
     case Some
     from this show ?thesis
-    unfolding exhaustive'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def by auto
+    unfolding exhaustive_int'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def by auto
   qed
 qed
 
-lemma int_of_nat:
-  "Code_Numeral.int_of (Code_Numeral.of_nat n) = int n"
-unfolding int_of_def by simp
-
 instance int :: complete
 proof
   fix n f
-  show "(\<exists>v. size (v :: int) \<le> n \<and> is_some (f v)) = is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
+  show "(\<exists>v. size (v :: int) \<le> n \<and> is_some (f v)) = is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
     unfolding exhaustive_int_def complete_exhaustive'[symmetric]
     apply auto apply (rule_tac x="v" in exI)
-    unfolding size_int_def by (metis int_of_nat abs_le_iff minus_le_iff nat_le_iff)+
+    unfolding size_int_def by (metis abs_le_iff minus_le_iff nat_le_iff)+
 qed
 
-declare exhaustive_code_numeral'.simps[simp del]
+declare exhaustive_natural'.simps[simp del]
 
-lemma complete_code_numeral':
+lemma complete_natural':
   "(\<exists>n. j \<le> n \<and> n \<le> k \<and> is_some (f n)) =
-    is_some (Quickcheck_Exhaustive.exhaustive_code_numeral' f k j)"
-proof (induct rule: exhaustive_code_numeral'.induct[of _ f k j])
+    is_some (Quickcheck_Exhaustive.exhaustive_natural' f k j)"
+proof (induct rule: exhaustive_natural'.induct[of _ f k j])
   case (1 f k j)
-  show "(\<exists>n\<ge>j. n \<le> k \<and> is_some (f n)) = is_some (Quickcheck_Exhaustive.exhaustive_code_numeral' f k j)"
-  unfolding exhaustive_code_numeral'.simps[of f k j] Quickcheck_Exhaustive.orelse_def
-  apply auto
+  show "(\<exists>n\<ge>j. n \<le> k \<and> is_some (f n)) = is_some (Quickcheck_Exhaustive.exhaustive_natural' f k j)"
+  unfolding exhaustive_natural'.simps [of f k j] Quickcheck_Exhaustive.orelse_def
   apply (auto split: option.split)
-  apply (insert 1[symmetric])
-  apply simp
-  apply (metis is_some.simps(2) less_eq_code_numeral_def not_less_eq_eq order_antisym)
-  apply simp
-  apply (split option.split_asm)
-  defer apply fastforce
-  apply simp by (metis Suc_leD)
+  apply (auto simp add: less_eq_natural_def less_natural_def 1 [symmetric] dest: Suc_leD)
+  apply (metis is_some.simps(2) natural_eqI not_less_eq_eq order_antisym)
+  done
 qed
 
-instance code_numeral :: complete
+instance natural :: complete
 proof
   fix n f
-  show "(\<exists>v. size (v :: code_numeral) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
-  unfolding exhaustive_code_numeral_def complete_code_numeral'[symmetric]
-  by (auto simp add: size_code_numeral_def)
+  show "(\<exists>v. size (v :: natural) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
+  unfolding exhaustive_natural_def complete_natural' [symmetric]
+    by (auto simp add: size_natural_def less_eq_natural_def)
 qed  
 
 instance nat :: complete
 proof
   fix n f
-  show "(\<exists>v. size (v :: nat) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
-    unfolding exhaustive_nat_def complete[of n "%x. f (Code_Numeral.nat_of x)", symmetric]
+  show "(\<exists>v. size (v :: nat) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
+    unfolding exhaustive_nat_def complete[of n "%x. f (nat_of_natural x)", symmetric]
     apply auto
-    apply (rule_tac x="Code_Numeral.of_nat v" in exI)
-    apply (auto simp add: size_code_numeral_def size_nat_def) done
+    apply (rule_tac x="natural_of_nat v" in exI)
+    apply (auto simp add: size_natural_def size_nat_def) done
 qed
 
 instance list :: (complete) complete
 proof
   fix n f
-  show "(\<exists> v. size (v :: 'a list) \<le> n \<and> is_some (f (v :: 'a list))) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))"
+  show "(\<exists> v. size (v :: 'a list) \<le> n \<and> is_some (f (v :: 'a list))) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (natural_of_nat n))"
   proof (induct n arbitrary: f)
     case 0
     { fix v have "size (v :: 'a list) > 0" by (induct v) auto }
@@ -174,25 +168,25 @@
     proof
       assume "\<exists>v. Completeness.size_class.size v \<le> Suc n \<and> is_some (f v)"
       then obtain v where v: "size v \<le> Suc n" "is_some (f v)" by blast
-      show "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat (Suc n)))"
+      show "is_some (exhaustive_class.exhaustive f (natural_of_nat (Suc n)))"
       proof (cases "v")
       case Nil
         from this v show ?thesis
-          unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
-          by (auto split: option.split)
+          unfolding list.exhaustive_list.simps[of _ "natural_of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
+          by (auto split: option.split simp add: less_natural_def)
       next 
       case (Cons v' vs')
         from Cons v have size_v': "Completeness.size_class.size v' \<le> n"
           and "Completeness.size_class.size vs' \<le> n" by auto
-        from Suc v Cons this have "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (Code_Numeral.of_nat n))"
+        from Suc v Cons this have "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (natural_of_nat n))"
           by metis
-        from complete_imp1[OF size_v', of "%x. (exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (Code_Numeral.of_nat n))", OF this]
+        from complete_imp1[OF size_v', of "%x. (exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (natural_of_nat n))", OF this]
         show ?thesis
-          unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
-          by (auto split: option.split)
+          unfolding list.exhaustive_list.simps[of _ "natural_of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
+          by (auto split: option.split simp add: less_natural_def)
       qed
     next
-      assume is_some: "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat (Suc n)))"
+      assume is_some: "is_some (exhaustive_class.exhaustive f (natural_of_nat (Suc n)))"
       show "\<exists>v. size v \<le> Suc n \<and> is_some (f v)"
       proof (cases "f []")
         case Some
@@ -201,12 +195,12 @@
       next
         case None
         with is_some have
-          "is_some (exhaustive_class.exhaustive (\<lambda>x. exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (Code_Numeral.of_nat n)) (Code_Numeral.of_nat n))"
-          unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
-          by simp
+          "is_some (exhaustive_class.exhaustive (\<lambda>x. exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (natural_of_nat n)) (natural_of_nat n))"
+          unfolding list.exhaustive_list.simps[of _ "natural_of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def
+          by (simp add: less_natural_def)
         then obtain v' where
             v: "size v' \<le> n"
-              "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (Code_Numeral.of_nat n))"
+              "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (natural_of_nat n))"
           by (rule complete_imp2)
         with Suc[of "%xs. f (v' # xs)"]
         obtain vs' where vs': "size vs' \<le> n" "is_some (f (v' # vs'))"
@@ -219,3 +213,4 @@
 qed
 
 end
+
--- a/src/HOL/Quickcheck_Examples/Hotel_Example.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Examples/Hotel_Example.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -105,8 +105,8 @@
   [code]: "cps_of_set (set xs) f = find_first f xs"
 sorry
 
-consts pos_cps_of_set :: "'a set => ('a => (bool * term list) option) => code_numeral => (bool * term list) option"
-consts neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => code_numeral => term list Quickcheck_Exhaustive.three_valued"
+consts pos_cps_of_set :: "'a set => ('a => (bool * term list) option) => natural => (bool * term list) option"
+consts neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => natural => term list Quickcheck_Exhaustive.three_valued"
 
 lemma
   [code]: "pos_cps_of_set (set xs) f i = find_first f xs"
--- a/src/HOL/Quickcheck_Exhaustive.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Exhaustive.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -16,42 +16,78 @@
 subsection {* exhaustive generator type classes *}
 
 class exhaustive = term_of +
-  fixes exhaustive :: "('a \<Rightarrow> (bool * term list) option) \<Rightarrow> code_numeral \<Rightarrow> (bool * term list) option"
+  fixes exhaustive :: "('a \<Rightarrow> (bool * term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option"
   
 class full_exhaustive = term_of +
-  fixes full_exhaustive :: "('a * (unit => term) \<Rightarrow> (bool * term list) option) \<Rightarrow> code_numeral \<Rightarrow> (bool * term list) option"
+  fixes full_exhaustive :: "('a * (unit => term) \<Rightarrow> (bool * term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option"
 
-instantiation code_numeral :: full_exhaustive
+instantiation natural :: full_exhaustive
 begin
 
-function full_exhaustive_code_numeral' :: "(code_numeral * (unit => term) => (bool * term list) option) => code_numeral => code_numeral => (bool * term list) option"
-  where "full_exhaustive_code_numeral' f d i =
+function full_exhaustive_natural' :: "(natural * (unit => term) => (bool * term list) option) => natural => natural => (bool * term list) option"
+  where "full_exhaustive_natural' f d i =
     (if d < i then None
-    else (f (i, %_. Code_Evaluation.term_of i)) orelse (full_exhaustive_code_numeral' f d (i + 1)))"
+    else (f (i, %_. Code_Evaluation.term_of i)) orelse (full_exhaustive_natural' f d (i + 1)))"
 by pat_completeness auto
 
 termination
-  by (relation "measure (%(_, d, i). Code_Numeral.nat_of (d + 1 - i))") auto
+  by (relation "measure (%(_, d, i). nat_of_natural (d + 1 - i))")
+    (auto simp add: less_natural_def)
 
-definition "full_exhaustive f d = full_exhaustive_code_numeral' f d 0"
+definition "full_exhaustive f d = full_exhaustive_natural' f d 0"
 
 instance ..
 
 end
 
-instantiation code_numeral :: exhaustive
+instantiation natural :: exhaustive
 begin
 
-function exhaustive_code_numeral' :: "(code_numeral => (bool * term list) option) => code_numeral => code_numeral => (bool * term list) option"
-  where "exhaustive_code_numeral' f d i =
+function exhaustive_natural' :: "(natural => (bool * term list) option) => natural => natural => (bool * term list) option"
+  where "exhaustive_natural' f d i =
     (if d < i then None
-    else (f i orelse exhaustive_code_numeral' f d (i + 1)))"
+    else (f i orelse exhaustive_natural' f d (i + 1)))"
 by pat_completeness auto
 
 termination
-  by (relation "measure (%(_, d, i). Code_Numeral.nat_of (d + 1 - i))") auto
+  by (relation "measure (%(_, d, i). nat_of_natural (d + 1 - i))")
+    (auto simp add: less_natural_def)
+
+definition "exhaustive f d = exhaustive_natural' f d 0"
+
+instance ..
+
+end
+
+instantiation integer :: exhaustive
+begin
+
+function exhaustive_integer' :: "(integer => (bool * term list) option) => integer => integer => (bool * term list) option"
+  where "exhaustive_integer' f d i = (if d < i then None else (f i orelse exhaustive_integer' f d (i + 1)))"
+by pat_completeness auto
 
-definition "exhaustive f d = exhaustive_code_numeral' f d 0"
+termination 
+  by (relation "measure (%(_, d, i). nat_of_integer (d + 1 - i))")
+    (auto simp add: less_integer_def nat_of_integer_def)
+
+definition "exhaustive f d = exhaustive_integer' f (integer_of_natural d) (- (integer_of_natural d))"
+
+instance ..
+
+end
+
+instantiation integer :: full_exhaustive
+begin
+
+function full_exhaustive_integer' :: "(integer * (unit => term) => (bool * term list) option) => integer => integer => (bool * term list) option"
+  where "full_exhaustive_integer' f d i = (if d < i then None else (case f (i, %_. Code_Evaluation.term_of i) of Some t => Some t | None => full_exhaustive_integer' f d (i + 1)))"
+by pat_completeness auto
+
+termination 
+  by (relation "measure (%(_, d, i). nat_of_integer (d + 1 - i))")
+    (auto simp add: less_integer_def nat_of_integer_def)
+
+definition "full_exhaustive f d = full_exhaustive_integer' f (integer_of_natural d) (- (integer_of_natural d))"
 
 instance ..
 
@@ -60,7 +96,7 @@
 instantiation nat :: exhaustive
 begin
 
-definition "exhaustive f d = exhaustive (%x. f (Code_Numeral.nat_of x)) d"
+definition "exhaustive f d = exhaustive (%x. f (nat_of_natural x)) d"
 
 instance ..
 
@@ -69,7 +105,7 @@
 instantiation nat :: full_exhaustive
 begin
 
-definition "full_exhaustive f d = full_exhaustive (%(x, xt). f (Code_Numeral.nat_of x, %_. Code_Evaluation.term_of (Code_Numeral.nat_of x))) d"
+definition "full_exhaustive f d = full_exhaustive (%(x, xt). f (nat_of_natural x, %_. Code_Evaluation.term_of (nat_of_natural x))) d"
 
 instance ..
 
@@ -78,14 +114,15 @@
 instantiation int :: exhaustive
 begin
 
-function exhaustive' :: "(int => (bool * term list) option) => int => int => (bool * term list) option"
-  where "exhaustive' f d i = (if d < i then None else (f i orelse exhaustive' f d (i + 1)))"
+function exhaustive_int' :: "(int => (bool * term list) option) => int => int => (bool * term list) option"
+  where "exhaustive_int' f d i = (if d < i then None else (f i orelse exhaustive_int' f d (i + 1)))"
 by pat_completeness auto
 
 termination 
   by (relation "measure (%(_, d, i). nat (d + 1 - i))") auto
 
-definition "exhaustive f d = exhaustive' f (Code_Numeral.int_of d) (- (Code_Numeral.int_of d))"
+definition "exhaustive f d = exhaustive_int' f (int_of_integer (integer_of_natural d))
+  (- (int_of_integer (integer_of_natural d)))"
 
 instance ..
 
@@ -94,14 +131,15 @@
 instantiation int :: full_exhaustive
 begin
 
-function full_exhaustive' :: "(int * (unit => term) => (bool * term list) option) => int => int => (bool * term list) option"
-  where "full_exhaustive' f d i = (if d < i then None else (case f (i, %_. Code_Evaluation.term_of i) of Some t => Some t | None => full_exhaustive' f d (i + 1)))"
+function full_exhaustive_int' :: "(int * (unit => term) => (bool * term list) option) => int => int => (bool * term list) option"
+  where "full_exhaustive_int' f d i = (if d < i then None else (case f (i, %_. Code_Evaluation.term_of i) of Some t => Some t | None => full_exhaustive_int' f d (i + 1)))"
 by pat_completeness auto
 
 termination 
   by (relation "measure (%(_, d, i). nat (d + 1 - i))") auto
 
-definition "full_exhaustive f d = full_exhaustive' f (Code_Numeral.int_of d) (- (Code_Numeral.int_of d))"
+definition "full_exhaustive f d = full_exhaustive_int' f (int_of_integer (integer_of_natural d))
+  (- (int_of_integer (integer_of_natural d)))"
 
 instance ..
 
@@ -154,14 +192,14 @@
 instantiation "fun" :: ("{equal, exhaustive}", exhaustive) exhaustive
 begin
 
-fun exhaustive_fun' :: "(('a => 'b) => (bool * term list) option) => code_numeral => code_numeral => (bool * term list) option"
+fun exhaustive_fun' :: "(('a => 'b) => (bool * term list) option) => natural => natural => (bool * term list) option"
 where
   "exhaustive_fun' f i d = (exhaustive (%b. f (%_. b)) d)
    orelse (if i > 1 then
      exhaustive_fun' (%g. exhaustive (%a. exhaustive (%b.
        f (g(a := b))) d) d) (i - 1) d else None)"
 
-definition exhaustive_fun :: "(('a => 'b) => (bool * term list) option) => code_numeral => (bool * term list) option"
+definition exhaustive_fun :: "(('a => 'b) => (bool * term list) option) => natural => (bool * term list) option"
 where
   "exhaustive_fun f d = exhaustive_fun' f d d" 
 
@@ -176,14 +214,14 @@
 instantiation "fun" :: ("{equal, full_exhaustive}", full_exhaustive) full_exhaustive
 begin
 
-fun full_exhaustive_fun' :: "(('a => 'b) * (unit => term) => (bool * term list) option) => code_numeral => code_numeral => (bool * term list) option"
+fun full_exhaustive_fun' :: "(('a => 'b) * (unit => term) => (bool * term list) option) => natural => natural => (bool * term list) option"
 where
   "full_exhaustive_fun' f i d = (full_exhaustive (%v. f (valtermify_absdummy v)) d)
    orelse (if i > 1 then
      full_exhaustive_fun' (%g. full_exhaustive (%a. full_exhaustive (%b.
        f (valtermify_fun_upd g a b)) d) d) (i - 1) d else None)"
 
-definition full_exhaustive_fun :: "(('a => 'b) * (unit => term) => (bool * term list) option) => code_numeral => (bool * term list) option"
+definition full_exhaustive_fun :: "(('a => 'b) * (unit => term) => (bool * term list) option) => natural => (bool * term list) option"
 where
   "full_exhaustive_fun f d = full_exhaustive_fun' f d d" 
 
@@ -197,7 +235,7 @@
   fixes check_all :: "('a * (unit \<Rightarrow> term) \<Rightarrow> (bool * term list) option) \<Rightarrow> (bool * term list) option"
   fixes enum_term_of :: "'a itself \<Rightarrow> unit \<Rightarrow> term list"
   
-fun check_all_n_lists :: "(('a :: check_all) list * (unit \<Rightarrow> term list) \<Rightarrow> (bool * term list) option) \<Rightarrow> code_numeral \<Rightarrow> (bool * term list) option"
+fun check_all_n_lists :: "(('a :: check_all) list * (unit \<Rightarrow> term list) \<Rightarrow> (bool * term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option"
 where
   "check_all_n_lists f n =
      (if n = 0 then f ([], (%_. [])) else check_all (%(x, xt). check_all_n_lists (%(xs, xst). f ((x # xs), (%_. (xt () # xst ())))) (n - 1)))"
@@ -227,7 +265,7 @@
     (let
       mk_term = mk_map_term (%_. Typerep.typerep (TYPE('a))) (%_. Typerep.typerep (TYPE('b))) (enum_term_of (TYPE('a)));
       enum = (Enum.enum :: 'a list)
-    in check_all_n_lists (\<lambda>(ys, yst). f (the o map_of (zip enum ys), mk_term yst)) (Code_Numeral.of_nat (length enum)))"
+    in check_all_n_lists (\<lambda>(ys, yst). f (the o map_of (zip enum ys), mk_term yst)) (natural_of_nat (length enum)))"
 
 definition enum_term_of_fun :: "('a => 'b) itself => unit => term list"
 where
@@ -470,12 +508,12 @@
 subsection {* Bounded universal quantifiers *}
 
 class bounded_forall =
-  fixes bounded_forall :: "('a \<Rightarrow> bool) \<Rightarrow> code_numeral \<Rightarrow> bool"
+  fixes bounded_forall :: "('a \<Rightarrow> bool) \<Rightarrow> natural \<Rightarrow> bool"
 
 subsection {* Fast exhaustive combinators *}
 
 class fast_exhaustive = term_of +
-  fixes fast_exhaustive :: "('a \<Rightarrow> unit) \<Rightarrow> code_numeral \<Rightarrow> unit"
+  fixes fast_exhaustive :: "('a \<Rightarrow> unit) \<Rightarrow> natural \<Rightarrow> unit"
 
 axiomatization throw_Counterexample :: "term list => unit"
 axiomatization catch_Counterexample :: "unit => term list option"
@@ -513,7 +551,7 @@
 where
   "cps_not n = (%c. case n (%u. Some []) of None => c () | Some _ => None)"
 
-type_synonym 'a pos_bound_cps = "('a => (bool * term list) option) => code_numeral => (bool * term list) option"
+type_synonym 'a pos_bound_cps = "('a => (bool * term list) option) => natural => (bool * term list) option"
 
 definition pos_bound_cps_empty :: "'a pos_bound_cps"
 where
@@ -538,7 +576,7 @@
 datatype 'a unknown = Unknown | Known 'a
 datatype 'a three_valued = Unknown_value | Value 'a | No_value
 
-type_synonym 'a neg_bound_cps = "('a unknown => term list three_valued) => code_numeral => term list three_valued"
+type_synonym 'a neg_bound_cps = "('a unknown => term list three_valued) => natural => term list three_valued"
 
 definition neg_bound_cps_empty :: "'a neg_bound_cps"
 where
@@ -573,7 +611,7 @@
 axiomatization unknown :: 'a
 
 notation (output) unknown  ("?")
- 
+
 ML_file "Tools/Quickcheck/exhaustive_generators.ML"
 
 setup {* Exhaustive_Generators.setup *}
@@ -588,15 +626,19 @@
 no_notation orelse (infixr "orelse" 55)
 
 hide_fact
-  exhaustive'_def
-  exhaustive_code_numeral'_def
+  exhaustive_int'_def
+  exhaustive_integer'_def
+  exhaustive_natural'_def
 
 hide_const valtermify_absdummy valtermify_fun_upd valterm_emptyset valtermify_insert valtermify_pair
   valtermify_Inl valtermify_Inr
   termify_fun_upd term_emptyset termify_insert termify_pair setify
 
 hide_const (open)
-  exhaustive full_exhaustive exhaustive' exhaustive_code_numeral' full_exhaustive_code_numeral'
+  exhaustive full_exhaustive
+  exhaustive_int' full_exhaustive_int'
+  exhaustive_integer' full_exhaustive_integer'
+  exhaustive_natural' full_exhaustive_natural'
   throw_Counterexample catch_Counterexample
   check_all enum_term_of
   orelse unknown mk_map_term check_all_n_lists check_all_subsets
--- a/src/HOL/Quickcheck_Narrowing.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Narrowing.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -9,188 +9,26 @@
 
 subsection {* Counterexample generator *}
 
-text {* We create a new target for the necessary code generation setup. *}
+subsubsection {* Code generation setup *}
 
 setup {* Code_Target.extend_target ("Haskell_Quickcheck", (Code_Haskell.target, K I)) *}
 
-subsubsection {* Code generation setup *}
-
 code_type typerep
   (Haskell_Quickcheck "Typerep")
 
 code_const Typerep.Typerep
   (Haskell_Quickcheck "Typerep")
 
+code_type integer
+  (Haskell_Quickcheck "Prelude.Int")
+
 code_reserved Haskell_Quickcheck Typerep
 
-subsubsection {* Type @{text "code_int"} for Haskell Quickcheck's Int type *}
-
-typedef code_int = "UNIV \<Colon> int set"
-  morphisms int_of of_int by rule
-
-lemma of_int_int_of [simp]:
-  "of_int (int_of k) = k"
-  by (rule int_of_inverse)
-
-lemma int_of_of_int [simp]:
-  "int_of (of_int n) = n"
-  by (rule of_int_inverse) (rule UNIV_I)
-
-lemma code_int:
-  "(\<And>n\<Colon>code_int. PROP P n) \<equiv> (\<And>n\<Colon>int. PROP P (of_int n))"
-proof
-  fix n :: int
-  assume "\<And>n\<Colon>code_int. PROP P n"
-  then show "PROP P (of_int n)" .
-next
-  fix n :: code_int
-  assume "\<And>n\<Colon>int. PROP P (of_int n)"
-  then have "PROP P (of_int (int_of n))" .
-  then show "PROP P n" by simp
-qed
-
-
-lemma int_of_inject [simp]:
-  "int_of k = int_of l \<longleftrightarrow> k = l"
-  by (rule int_of_inject)
-
-lemma of_int_inject [simp]:
-  "of_int n = of_int m \<longleftrightarrow> n = m"
-  by (rule of_int_inject) (rule UNIV_I)+
-
-instantiation code_int :: equal
-begin
-
-definition
-  "HOL.equal k l \<longleftrightarrow> HOL.equal (int_of k) (int_of l)"
-
-instance proof
-qed (auto simp add: equal_code_int_def equal_int_def equal_int_refl)
-
-end
-
-definition nat_of :: "code_int => nat"
-where
-  "nat_of i = nat (int_of i)"
-  
-instantiation code_int :: "{minus, linordered_semidom, semiring_div, neg_numeral, linorder}"
-begin
-
-definition [simp, code del]:
-  "0 = of_int 0"
-
-definition [simp, code del]:
-  "1 = of_int 1"
-
-definition [simp, code del]:
-  "n + m = of_int (int_of n + int_of m)"
-
-definition [simp, code del]:
-  "- n = of_int (- int_of n)"
-
-definition [simp, code del]:
-  "n - m = of_int (int_of n - int_of m)"
-
-definition [simp, code del]:
-  "n * m = of_int (int_of n * int_of m)"
-
-definition [simp, code del]:
-  "n div m = of_int (int_of n div int_of m)"
-
-definition [simp, code del]:
-  "n mod m = of_int (int_of n mod int_of m)"
-
-definition [simp, code del]:
-  "n \<le> m \<longleftrightarrow> int_of n \<le> int_of m"
-
-definition [simp, code del]:
-  "n < m \<longleftrightarrow> int_of n < int_of m"
-
-instance proof
-qed (auto simp add: code_int distrib_right zmult_zless_mono2)
-
-end
-
-lemma int_of_numeral [simp]:
-  "int_of (numeral k) = numeral k"
-  by (induct k) (simp_all only: numeral.simps plus_code_int_def
-    one_code_int_def of_int_inverse UNIV_I)
-
-definition Num :: "num \<Rightarrow> code_int"
-  where [code_abbrev]: "Num = numeral"
-
-lemma [code_abbrev]:
-  "- numeral k = (neg_numeral k :: code_int)"
-  by (unfold neg_numeral_def) simp
-
-code_datatype "0::code_int" Num
-
-lemma one_code_int_code [code, code_unfold]:
-  "(1\<Colon>code_int) = Numeral1"
-  by (simp only: numeral.simps)
-
-definition div_mod :: "code_int \<Rightarrow> code_int \<Rightarrow> code_int \<times> code_int" where
-  [code del]: "div_mod n m = (n div m, n mod m)"
-
-lemma [code]:
-  "n div m = fst (div_mod n m)"
-  unfolding div_mod_def by simp
-
-lemma [code]:
-  "n mod m = snd (div_mod n m)"
-  unfolding div_mod_def by simp
-
-lemma int_of_code [code]:
-  "int_of k = (if k = 0 then 0
-    else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
-proof -
-  have 1: "(int_of k div 2) * 2 + int_of k mod 2 = int_of k" 
-    by (rule mod_div_equality)
-  have "int_of k mod 2 = 0 \<or> int_of k mod 2 = 1" by auto
-  from this show ?thesis
-    apply auto
-    apply (insert 1) by (auto simp add: mult_ac)
-qed
-
-
-code_instance code_numeral :: equal
-  (Haskell_Quickcheck -)
-
-setup {* fold (Numeral.add_code @{const_name Num}
-  false Code_Printer.literal_numeral) ["Haskell_Quickcheck"]  *}
-
-code_type code_int
-  (Haskell_Quickcheck "Prelude.Int")
-
-code_const "0 \<Colon> code_int"
-  (Haskell_Quickcheck "0")
-
-code_const "1 \<Colon> code_int"
-  (Haskell_Quickcheck "1")
-
-code_const "minus \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> code_int"
-  (Haskell_Quickcheck infixl 6 "-")
-
-code_const div_mod
-  (Haskell_Quickcheck "divMod")
-
-code_const "HOL.equal \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
-  (Haskell_Quickcheck infix 4 "==")
-
-code_const "less_eq \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
-  (Haskell_Quickcheck infix 4 "<=")
-
-code_const "less \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
-  (Haskell_Quickcheck infix 4 "<")
-
-code_abort of_int
-
-hide_const (open) Num div_mod
 
 subsubsection {* Narrowing's deep representation of types and terms *}
 
 datatype narrowing_type = Narrowing_sum_of_products "narrowing_type list list"
-datatype narrowing_term = Narrowing_variable "code_int list" narrowing_type | Narrowing_constructor code_int "narrowing_term list"
+datatype narrowing_term = Narrowing_variable "integer list" narrowing_type | Narrowing_constructor integer "narrowing_term list"
 datatype 'a narrowing_cons = Narrowing_cons narrowing_type "(narrowing_term list => 'a) list"
 
 primrec map_cons :: "('a => 'b) => 'a narrowing_cons => 'b narrowing_cons"
@@ -207,7 +45,7 @@
  
 subsubsection {* Auxilary functions for Narrowing *}
 
-consts nth :: "'a list => code_int => 'a"
+consts nth :: "'a list => integer => 'a"
 
 code_const nth (Haskell_Quickcheck infixl 9  "!!")
 
@@ -215,7 +53,7 @@
 
 code_const error (Haskell_Quickcheck "error")
 
-consts toEnum :: "code_int => char"
+consts toEnum :: "integer => char"
 
 code_const toEnum (Haskell_Quickcheck "Prelude.toEnum")
 
@@ -225,7 +63,7 @@
 
 subsubsection {* Narrowing's basic operations *}
 
-type_synonym 'a narrowing = "code_int => 'a narrowing_cons"
+type_synonym 'a narrowing = "integer => 'a narrowing_cons"
 
 definition empty :: "'a narrowing"
 where
@@ -267,35 +105,33 @@
 using assms unfolding sum_def by (auto split: narrowing_cons.split narrowing_type.split)
 
 lemma [fundef_cong]:
-  assumes "f d = f' d" "(\<And>d'. 0 <= d' & d' < d ==> a d' = a' d')"
+  assumes "f d = f' d" "(\<And>d'. 0 \<le> d' \<and> d' < d \<Longrightarrow> a d' = a' d')"
   assumes "d = d'"
   shows "apply f a d = apply f' a' d'"
 proof -
-  note assms moreover
-  have "int_of (of_int 0) < int_of d' ==> int_of (of_int 0) <= int_of (of_int (int_of d' - int_of (of_int 1)))"
-    by (simp add: of_int_inverse)
-  moreover
-  have "int_of (of_int (int_of d' - int_of (of_int 1))) < int_of d'"
-    by (simp add: of_int_inverse)
+  note assms
+  moreover have "0 < d' \<Longrightarrow> 0 \<le> d' - 1"
+    by (simp add: less_integer_def less_eq_integer_def)
   ultimately show ?thesis
-    unfolding apply_def by (auto split: narrowing_cons.split narrowing_type.split simp add: Let_def)
+    by (auto simp add: apply_def Let_def
+      split: narrowing_cons.split narrowing_type.split)
 qed
 
 subsubsection {* Narrowing generator type class *}
 
 class narrowing =
-  fixes narrowing :: "code_int => 'a narrowing_cons"
+  fixes narrowing :: "integer => 'a narrowing_cons"
 
 datatype property = Universal narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Existential narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Property bool
 
 (* FIXME: hard-wired maximal depth of 100 here *)
 definition exists :: "('a :: {narrowing, partial_term_of} => property) => property"
 where
-  "exists f = (case narrowing (100 :: code_int) of Narrowing_cons ty cs => Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
+  "exists f = (case narrowing (100 :: integer) of Narrowing_cons ty cs => Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 
 definition "all" :: "('a :: {narrowing, partial_term_of} => property) => property"
 where
-  "all f = (case narrowing (100 :: code_int) of Narrowing_cons ty cs => Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
+  "all f = (case narrowing (100 :: integer) of Narrowing_cons ty cs => Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 
 subsubsection {* class @{text is_testable} *}
 
@@ -343,14 +179,14 @@
 where
   "narrowing_dummy_partial_term_of = partial_term_of"
 
-definition narrowing_dummy_narrowing :: "code_int => ('a :: narrowing) narrowing_cons"
+definition narrowing_dummy_narrowing :: "integer => ('a :: narrowing) narrowing_cons"
 where
   "narrowing_dummy_narrowing = narrowing"
 
 lemma [code]:
   "ensure_testable f =
     (let
-      x = narrowing_dummy_narrowing :: code_int => bool narrowing_cons;
+      x = narrowing_dummy_narrowing :: integer => bool narrowing_cons;
       y = narrowing_dummy_partial_term_of :: bool itself => narrowing_term => term;
       z = (conv :: _ => _ => unit)  in f)"
 unfolding Let_def ensure_testable_def ..
@@ -369,47 +205,76 @@
 subsection {* Narrowing for integers *}
 
 
-definition drawn_from :: "'a list => 'a narrowing_cons"
-where "drawn_from xs = Narrowing_cons (Narrowing_sum_of_products (map (%_. []) xs)) (map (%x y. x) xs)"
+definition drawn_from :: "'a list \<Rightarrow> 'a narrowing_cons"
+where
+  "drawn_from xs =
+    Narrowing_cons (Narrowing_sum_of_products (map (\<lambda>_. []) xs)) (map (\<lambda>x _. x) xs)"
 
-function around_zero :: "int => int list"
+function around_zero :: "int \<Rightarrow> int list"
 where
   "around_zero i = (if i < 0 then [] else (if i = 0 then [0] else around_zero (i - 1) @ [i, -i]))"
-by pat_completeness auto
+  by pat_completeness auto
 termination by (relation "measure nat") auto
 
-declare around_zero.simps[simp del]
+declare around_zero.simps [simp del]
 
 lemma length_around_zero:
   assumes "i >= 0" 
   shows "length (around_zero i) = 2 * nat i + 1"
-proof (induct rule: int_ge_induct[OF assms])
+proof (induct rule: int_ge_induct [OF assms])
   case 1
   from 1 show ?case by (simp add: around_zero.simps)
 next
   case (2 i)
   from 2 show ?case
-    by (simp add: around_zero.simps[of "i + 1"])
+    by (simp add: around_zero.simps [of "i + 1"])
 qed
 
 instantiation int :: narrowing
 begin
 
 definition
-  "narrowing_int d = (let (u :: _ => _ => unit) = conv; i = Quickcheck_Narrowing.int_of d in drawn_from (around_zero i))"
+  "narrowing_int d = (let (u :: _ \<Rightarrow> _ \<Rightarrow> unit) = conv; i = int_of_integer d
+    in drawn_from (around_zero i))"
 
 instance ..
 
 end
 
-lemma [code, code del]: "partial_term_of (ty :: int itself) t == undefined"
-by (rule partial_term_of_anything)+
+lemma [code, code del]: "partial_term_of (ty :: int itself) t \<equiv> undefined"
+  by (rule partial_term_of_anything)+
 
 lemma [code]:
-  "partial_term_of (ty :: int itself) (Narrowing_variable p t) == Code_Evaluation.Free (STR ''_'') (Typerep.Typerep (STR ''Int.int'') [])"
-  "partial_term_of (ty :: int itself) (Narrowing_constructor i []) == (if i mod 2 = 0 then
-     Code_Evaluation.term_of (- (int_of i) div 2) else Code_Evaluation.term_of ((int_of i + 1) div 2))"
-by (rule partial_term_of_anything)+
+  "partial_term_of (ty :: int itself) (Narrowing_variable p t) \<equiv>
+    Code_Evaluation.Free (STR ''_'') (Typerep.Typerep (STR ''Int.int'') [])"
+  "partial_term_of (ty :: int itself) (Narrowing_constructor i []) \<equiv>
+    (if i mod 2 = 0
+     then Code_Evaluation.term_of (- (int_of_integer i) div 2)
+     else Code_Evaluation.term_of ((int_of_integer i + 1) div 2))"
+  by (rule partial_term_of_anything)+
+
+instantiation integer :: narrowing
+begin
+
+definition
+  "narrowing_integer d = (let (u :: _ \<Rightarrow> _ \<Rightarrow> unit) = conv; i = int_of_integer d
+    in drawn_from (map integer_of_int (around_zero i)))"
+
+instance ..
+
+end
+
+lemma [code, code del]: "partial_term_of (ty :: integer itself) t \<equiv> undefined"
+  by (rule partial_term_of_anything)+
+
+lemma [code]:
+  "partial_term_of (ty :: integer itself) (Narrowing_variable p t) \<equiv>
+    Code_Evaluation.Free (STR ''_'') (Typerep.Typerep (STR ''Code_Numeral.integer'') [])"
+  "partial_term_of (ty :: integer itself) (Narrowing_constructor i []) \<equiv>
+    (if i mod 2 = 0
+     then Code_Evaluation.term_of (- i div 2)
+     else Code_Evaluation.term_of ((i + 1) div 2))"
+  by (rule partial_term_of_anything)+
 
 
 subsection {* The @{text find_unused_assms} command *}
@@ -418,9 +283,10 @@
 
 subsection {* Closing up *}
 
-hide_type code_int narrowing_type narrowing_term narrowing_cons property
-hide_const int_of of_int nat_of map_cons nth error toEnum marker empty Narrowing_cons conv non_empty ensure_testable all exists drawn_from around_zero
+hide_type narrowing_type narrowing_term narrowing_cons property
+hide_const map_cons nth error toEnum marker empty Narrowing_cons conv non_empty ensure_testable all exists drawn_from around_zero
 hide_const (open) Narrowing_variable Narrowing_constructor "apply" sum cons
 hide_fact empty_def cons_def conv.simps non_empty.simps apply_def sum_def ensure_testable_def all_def exists_def
 
 end
+
--- a/src/HOL/Quickcheck_Random.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Quickcheck_Random.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -21,7 +21,7 @@
 subsection {* The @{text random} class *}
 
 class random = typerep +
-  fixes random :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+  fixes random :: "natural \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
 
 
 subsection {* Fundamental and numeric types*}
@@ -41,7 +41,7 @@
 begin
 
 definition
-  random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+  random_itself :: "natural \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
 where "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))"
 
 instance ..
@@ -71,11 +71,11 @@
 instantiation nat :: random
 begin
 
-definition random_nat :: "code_numeral \<Rightarrow> Random.seed
+definition random_nat :: "natural \<Rightarrow> Random.seed
   \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed"
 where
   "random_nat i = Random.range (i + 1) \<circ>\<rightarrow> (\<lambda>k. Pair (
-     let n = Code_Numeral.nat_of k
+     let n = nat_of_natural k
      in (n, \<lambda>_. Code_Evaluation.term_of n)))"
 
 instance ..
@@ -87,13 +87,39 @@
 
 definition
   "random i = Random.range (2 * i + 1) \<circ>\<rightarrow> (\<lambda>k. Pair (
-     let j = (if k \<ge> i then Code_Numeral.int_of (k - i) else - Code_Numeral.int_of (i - k))
+     let j = (if k \<ge> i then int (nat_of_natural (k - i)) else - (int (nat_of_natural (i - k))))
      in (j, \<lambda>_. Code_Evaluation.term_of j)))"
 
 instance ..
 
 end
 
+instantiation natural :: random
+begin
+
+definition random_natural :: "natural \<Rightarrow> Random.seed
+  \<Rightarrow> (natural \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed"
+where
+  "random_natural i = Random.range (i + 1) \<circ>\<rightarrow> (\<lambda>n. Pair (n, \<lambda>_. Code_Evaluation.term_of n))"
+
+instance ..
+
+end
+
+instantiation integer :: random
+begin
+
+definition random_integer :: "natural \<Rightarrow> Random.seed
+  \<Rightarrow> (integer \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed"
+where
+  "random_integer i = Random.range (2 * i + 1) \<circ>\<rightarrow> (\<lambda>k. Pair (
+     let j = (if k \<ge> i then integer_of_natural (k - i) else - (integer_of_natural (i - k)))
+      in (j, \<lambda>_. Code_Evaluation.term_of j)))"
+
+instance ..
+
+end
+
 
 subsection {* Complex generators *}
 
@@ -114,7 +140,7 @@
 begin
 
 definition
-  random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+  random_fun :: "natural \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
   where "random i = random_fun_lift (random i)"
 
 instance ..
@@ -126,7 +152,7 @@
 definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a"
   where "collapse f = (f \<circ>\<rightarrow> id)"
 
-definition beyond :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
+definition beyond :: "natural \<Rightarrow> natural \<Rightarrow> natural"
   where "beyond k l = (if l > k then l else 0)"
 
 lemma beyond_zero: "beyond k 0 = 0"
@@ -155,13 +181,13 @@
   "random_aux_set i j =
     collapse (Random.select_weight [(1, Pair valterm_emptyset),
       (i, random j \<circ>\<rightarrow> (%x. random_aux_set (i - 1) j \<circ>\<rightarrow> (%s. Pair (valtermify_insert x s))))])"
-proof (induct i rule: code_numeral.induct)
+proof (induct i rule: natural.induct)
   case zero
   show ?case by (subst select_weight_drop_zero [symmetric])
-    (simp add: random_aux_set.simps [simplified])
+    (simp add: random_aux_set.simps [simplified] less_natural_def)
 next
   case (Suc i)
-  show ?case by (simp only: random_aux_set.simps(2) [of "i"] Suc_code_numeral_minus_one)
+  show ?case by (simp only: random_aux_set.simps(2) [of "i"] Suc_natural_minus_one)
 qed
 
 definition "random_set i = random_aux_set i i"
@@ -171,11 +197,11 @@
 end
 
 lemma random_aux_rec:
-  fixes random_aux :: "code_numeral \<Rightarrow> 'a"
+  fixes random_aux :: "natural \<Rightarrow> 'a"
   assumes "random_aux 0 = rhs 0"
     and "\<And>k. random_aux (Code_Numeral.Suc k) = rhs (Code_Numeral.Suc k)"
   shows "random_aux k = rhs k"
-  using assms by (rule code_numeral.induct)
+  using assms by (rule natural.induct)
 
 subsection {* Deriving random generators for datatypes *}
 
--- a/src/HOL/ROOT	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/ROOT	Fri Feb 15 08:31:31 2013 +0100
@@ -28,8 +28,6 @@
     Code_Char_ord
     Product_Lexorder
     Product_Order
-    Code_Integer
-    Efficient_Nat
     (* Code_Prolog  FIXME cf. 76965c356d2a *)
     Code_Real_Approx_By_Float
     Code_Target_Numeral
@@ -282,7 +280,6 @@
   theories [document = false]
     "~~/src/HOL/Library/Countable"
     "~~/src/HOL/Library/Monad_Syntax"
-    "~~/src/HOL/Library/Code_Natural"
     "~~/src/HOL/Library/LaTeXsugar"
   theories Imperative_HOL_ex
   files "document/root.bib" "document/root.tex"
@@ -299,7 +296,7 @@
   description {* Examples for program extraction in Higher-Order Logic *}
   options [condition = ISABELLE_POLYML, proofs = 2, parallel_proofs = 0]
   theories [document = false]
-    "~~/src/HOL/Library/Efficient_Nat"
+    "~~/src/HOL/Library/Code_Target_Numeral"
     "~~/src/HOL/Library/Monad_Syntax"
     "~~/src/HOL/Number_Theory/Primes"
     "~~/src/HOL/Number_Theory/UniqueFactorization"
@@ -315,7 +312,7 @@
 session "HOL-Proofs-Lambda" in "Proofs/Lambda" = "HOL-Proofs" +
   options [document_graph, print_mode = "no_brackets", proofs = 2, parallel_proofs = 0]
   theories [document = false]
-    "~~/src/HOL/Library/Code_Integer"
+    "~~/src/HOL/Library/Code_Target_Int"
   theories
     Eta
     StrongNorm
--- a/src/HOL/Random.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Random.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -13,21 +13,21 @@
 
 subsection {* Auxiliary functions *}
 
-fun log :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
+fun log :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
   "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
 
-definition inc_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
+definition inc_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
   "inc_shift v k = (if v = k then 1 else k + 1)"
 
-definition minus_shift :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
+definition minus_shift :: "natural \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> natural" where
   "minus_shift r k l = (if k < l then r + k - l else k - l)"
 
 
 subsection {* Random seeds *}
 
-type_synonym seed = "code_numeral \<times> code_numeral"
+type_synonym seed = "natural \<times> natural"
 
-primrec "next" :: "seed \<Rightarrow> code_numeral \<times> seed" where
+primrec "next" :: "seed \<Rightarrow> natural \<times> seed" where
   "next (v, w) = (let
      k =  v div 53668;
      v' = minus_shift 2147483563 ((v mod 53668) * 40014) (k * 12211);
@@ -47,55 +47,55 @@
 
 subsection {* Base selectors *}
 
-fun iterate :: "code_numeral \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
+fun iterate :: "natural \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
   "iterate k f x = (if k = 0 then Pair x else f x \<circ>\<rightarrow> iterate (k - 1) f)"
 
-definition range :: "code_numeral \<Rightarrow> seed \<Rightarrow> code_numeral \<times> seed" where
+definition range :: "natural \<Rightarrow> seed \<Rightarrow> natural \<times> seed" where
   "range k = iterate (log 2147483561 k)
       (\<lambda>l. next \<circ>\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
     \<circ>\<rightarrow> (\<lambda>v. Pair (v mod k))"
 
 lemma range:
   "k > 0 \<Longrightarrow> fst (range k s) < k"
-  by (simp add: range_def split_def del: log.simps iterate.simps)
+  by (simp add: range_def split_def less_natural_def del: log.simps iterate.simps)
 
 definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
-  "select xs = range (Code_Numeral.of_nat (length xs))
-    \<circ>\<rightarrow> (\<lambda>k. Pair (nth xs (Code_Numeral.nat_of k)))"
+  "select xs = range (natural_of_nat (length xs))
+    \<circ>\<rightarrow> (\<lambda>k. Pair (nth xs (nat_of_natural k)))"
      
 lemma select:
   assumes "xs \<noteq> []"
   shows "fst (select xs s) \<in> set xs"
 proof -
-  from assms have "Code_Numeral.of_nat (length xs) > 0" by simp
+  from assms have "natural_of_nat (length xs) > 0" by (simp add: less_natural_def)
   with range have
-    "fst (range (Code_Numeral.of_nat (length xs)) s) < Code_Numeral.of_nat (length xs)" by best
+    "fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)" by best
   then have
-    "Code_Numeral.nat_of (fst (range (Code_Numeral.of_nat (length xs)) s)) < length xs" by simp
+    "nat_of_natural (fst (range (natural_of_nat (length xs)) s)) < length xs" by (simp add: less_natural_def)
   then show ?thesis
     by (simp add: split_beta select_def)
 qed
 
-primrec pick :: "(code_numeral \<times> 'a) list \<Rightarrow> code_numeral \<Rightarrow> 'a" where
+primrec pick :: "(natural \<times> 'a) list \<Rightarrow> natural \<Rightarrow> 'a" where
   "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"
 
 lemma pick_member:
   "i < listsum (map fst xs) \<Longrightarrow> pick xs i \<in> set (map snd xs)"
-  by (induct xs arbitrary: i) simp_all
+  by (induct xs arbitrary: i) (simp_all add: less_natural_def)
 
 lemma pick_drop_zero:
   "pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"
-  by (induct xs) (auto simp add: fun_eq_iff)
+  by (induct xs) (auto simp add: fun_eq_iff less_natural_def minus_natural_def)
 
 lemma pick_same:
-  "l < length xs \<Longrightarrow> Random.pick (map (Pair 1) xs) (Code_Numeral.of_nat l) = nth xs l"
+  "l < length xs \<Longrightarrow> Random.pick (map (Pair 1) xs) (natural_of_nat l) = nth xs l"
 proof (induct xs arbitrary: l)
   case Nil then show ?case by simp
 next
-  case (Cons x xs) then show ?case by (cases l) simp_all
+  case (Cons x xs) then show ?case by (cases l) (simp_all add: less_natural_def)
 qed
 
-definition select_weight :: "(code_numeral \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
+definition select_weight :: "(natural \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
   "select_weight xs = range (listsum (map fst xs))
    \<circ>\<rightarrow> (\<lambda>k. Pair (pick xs k))"
 
@@ -112,13 +112,13 @@
 
 lemma select_weight_cons_zero:
   "select_weight ((0, x) # xs) = select_weight xs"
-  by (simp add: select_weight_def)
+  by (simp add: select_weight_def less_natural_def)
 
 lemma select_weight_drop_zero:
   "select_weight (filter (\<lambda>(k, _). k > 0) xs) = select_weight xs"
 proof -
   have "listsum (map fst [(k, _)\<leftarrow>xs . 0 < k]) = listsum (map fst xs)"
-    by (induct xs) auto
+    by (induct xs) (auto simp add: less_natural_def, simp add: plus_natural_def)
   then show ?thesis by (simp only: select_weight_def pick_drop_zero)
 qed
 
@@ -126,13 +126,13 @@
   assumes "xs \<noteq> []"
   shows "select_weight (map (Pair 1) xs) = select xs"
 proof -
-  have less: "\<And>s. fst (range (Code_Numeral.of_nat (length xs)) s) < Code_Numeral.of_nat (length xs)"
-    using assms by (intro range) simp
-  moreover have "listsum (map fst (map (Pair 1) xs)) = Code_Numeral.of_nat (length xs)"
+  have less: "\<And>s. fst (range (natural_of_nat (length xs)) s) < natural_of_nat (length xs)"
+    using assms by (intro range) (simp add: less_natural_def)
+  moreover have "listsum (map fst (map (Pair 1) xs)) = natural_of_nat (length xs)"
     by (induct xs) simp_all
   ultimately show ?thesis
     by (auto simp add: select_weight_def select_def scomp_def split_def
-      fun_eq_iff pick_same [symmetric])
+      fun_eq_iff pick_same [symmetric] less_natural_def)
 qed
 
 
@@ -147,7 +147,7 @@
 
 open Random_Engine;
 
-type seed = int * int;
+type seed = Code_Numeral.natural * Code_Numeral.natural;
 
 local
 
@@ -156,7 +156,7 @@
     val now = Time.toMilliseconds (Time.now ());
     val (q, s1) = IntInf.divMod (now, 2147483562);
     val s2 = q mod 2147483398;
-  in (s1 + 1, s2 + 1) end);
+  in pairself Code_Numeral.natural_of_integer (s1 + 1, s2 + 1) end);
 
 in
 
@@ -188,3 +188,4 @@
 no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
 
 end
+
--- a/src/HOL/Random_Pred.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Random_Pred.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -7,13 +7,13 @@
 imports Quickcheck_Random
 begin
 
-fun iter' :: "'a itself \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a::random) Predicate.pred"
+fun iter' :: "'a itself \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> ('a::random) Predicate.pred"
 where
   "iter' T nrandom sz seed = (if nrandom = 0 then bot_class.bot else
      let ((x, _), seed') = Quickcheck_Random.random sz seed
    in Predicate.Seq (%u. Predicate.Insert x (iter' T (nrandom - 1) sz seed')))"
 
-definition iter :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a::random) Predicate.pred"
+definition iter :: "natural \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> ('a::random) Predicate.pred"
 where
   "iter nrandom sz seed = iter' (TYPE('a)) nrandom sz seed"
 
@@ -48,7 +48,7 @@
 where
   "if_randompred b = (if b then single () else empty)"
 
-definition iterate_upto :: "(code_numeral \<Rightarrow> 'a) => code_numeral \<Rightarrow> code_numeral \<Rightarrow> 'a random_pred"
+definition iterate_upto :: "(natural \<Rightarrow> 'a) => natural \<Rightarrow> natural \<Rightarrow> 'a random_pred"
 where
   "iterate_upto f n m = Pair (Predicate.iterate_upto f n m)"
 
--- a/src/HOL/Random_Sequence.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Random_Sequence.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -7,7 +7,7 @@
 imports Random_Pred
 begin
 
-type_synonym 'a random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a Limited_Sequence.dseq \<times> Random.seed)"
+type_synonym 'a random_dseq = "natural \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> ('a Limited_Sequence.dseq \<times> Random.seed)"
 
 definition empty :: "'a random_dseq"
 where
@@ -44,13 +44,13 @@
 where
   "map f P = bind P (single o f)"
 
-fun Random :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a random_dseq"
+fun Random :: "(natural \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a random_dseq"
 where
   "Random g nrandom = (%size. if nrandom <= 0 then (Pair Limited_Sequence.empty) else
      (scomp (g size) (%r. scomp (Random g (nrandom - 1) size) (%rs. Pair (Limited_Sequence.union (Limited_Sequence.single (fst r)) rs)))))"
 
 
-type_synonym 'a pos_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a Limited_Sequence.pos_dseq"
+type_synonym 'a pos_random_dseq = "natural \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> 'a Limited_Sequence.pos_dseq"
 
 definition pos_empty :: "'a pos_random_dseq"
 where
@@ -76,7 +76,7 @@
 where
   "pos_if_random_dseq b = (if b then pos_single () else pos_empty)"
 
-definition pos_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a pos_random_dseq"
+definition pos_iterate_upto :: "(natural => 'a) => natural => natural => 'a pos_random_dseq"
 where
   "pos_iterate_upto f n m = (\<lambda>nrandom size seed. Limited_Sequence.pos_iterate_upto f n m)"
 
@@ -85,18 +85,18 @@
   "pos_map f P = pos_bind P (pos_single o f)"
 
 fun iter :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
-  \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
+  \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
 where
   "iter random nrandom seed =
     (if nrandom = 0 then Lazy_Sequence.empty else Lazy_Sequence.Lazy_Sequence (%u. let ((x, _), seed') = random seed in Some (x, iter random (nrandom - 1) seed')))"
 
-definition pos_Random :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
+definition pos_Random :: "(natural \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
   \<Rightarrow> 'a pos_random_dseq"
 where
   "pos_Random g = (%nrandom size seed depth. iter (g size) nrandom seed)"
 
 
-type_synonym 'a neg_random_dseq = "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> 'a Limited_Sequence.neg_dseq"
+type_synonym 'a neg_random_dseq = "natural \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> 'a Limited_Sequence.neg_dseq"
 
 definition neg_empty :: "'a neg_random_dseq"
 where
@@ -122,7 +122,7 @@
 where
   "neg_if_random_dseq b = (if b then neg_single () else neg_empty)"
 
-definition neg_iterate_upto :: "(code_numeral => 'a) => code_numeral => code_numeral => 'a neg_random_dseq"
+definition neg_iterate_upto :: "(natural => 'a) => natural => natural => 'a neg_random_dseq"
 where
   "neg_iterate_upto f n m = (\<lambda>nrandom size seed. Limited_Sequence.neg_iterate_upto f n m)"
 
--- a/src/HOL/Rat.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Rat.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1031,7 +1031,7 @@
 
 definition
   "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>num. Random.range i \<circ>\<rightarrow> (\<lambda>denom. Pair (
-     let j = Code_Numeral.int_of (denom + 1)
+     let j = int_of_integer (integer_of_natural (denom + 1))
      in valterm_fract num (j, \<lambda>u. Code_Evaluation.term_of j))))"
 
 instance ..
@@ -1045,7 +1045,8 @@
 begin
 
 definition
-  "exhaustive_rat f d = Quickcheck_Exhaustive.exhaustive (%l. Quickcheck_Exhaustive.exhaustive (%k. f (Fract k (Code_Numeral.int_of l + 1))) d) d"
+  "exhaustive_rat f d = Quickcheck_Exhaustive.exhaustive
+    (\<lambda>l. Quickcheck_Exhaustive.exhaustive (\<lambda>k. f (Fract k (int_of_integer (integer_of_natural l) + 1))) d) d"
 
 instance ..
 
@@ -1056,7 +1057,7 @@
 
 definition
   "full_exhaustive_rat f d = Quickcheck_Exhaustive.full_exhaustive (%(l, _). Quickcheck_Exhaustive.full_exhaustive (%k.
-     f (let j = Code_Numeral.int_of l + 1
+     f (let j = int_of_integer (integer_of_natural l) + 1
         in valterm_fract k (j, %_. Code_Evaluation.term_of j))) d) d"
 
 instance ..
@@ -1135,3 +1136,4 @@
 declare Quotient_rat[quot_del]
 
 end
+
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -589,11 +589,11 @@
 val is_real_datatype = is_some oo Datatype.get_info
 fun is_standard_datatype thy = the oo triple_lookup (type_match thy)
 
-(* FIXME: Use antiquotation for "code_numeral" below or detect "rep_datatype",
+(* FIXME: Use antiquotation for "natural" below or detect "rep_datatype",
    e.g., by adding a field to "Datatype_Aux.info". *)
 fun is_basic_datatype thy stds s =
   member (op =) [@{type_name prod}, @{type_name set}, @{type_name bool},
-                 @{type_name int}, "Code_Numeral.code_numeral"] s orelse
+                 @{type_name int}, @{type_name natural}, @{type_name integer}] s orelse
   (s = @{type_name nat} andalso is_standard_datatype thy stds nat_T)
 
 fun repair_constr_type ctxt body_T' T =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -31,7 +31,7 @@
 
 fun mk_iterate_upto T (f, from, to) =
   list_comb (Const (@{const_name Predicate.iterate_upto},
-      [@{typ code_numeral} --> T, @{typ code_numeral}, @{typ code_numeral}] ---> mk_monadT T),
+      [@{typ natural} --> T, @{typ natural}, @{typ natural}] ---> mk_monadT T),
     [f, from, to])
 
 fun mk_not t =
@@ -115,10 +115,10 @@
 struct
 
 val resultT = @{typ "(bool * Code_Evaluation.term list) option"}
-fun mk_monadT T = (T --> resultT) --> @{typ "code_numeral"} --> resultT
+fun mk_monadT T = (T --> resultT) --> @{typ "natural"} --> resultT
 
 fun dest_monadT (Type ("fun", [Type ("fun", [T, @{typ "(bool * term list) option"}]),
-  @{typ "code_numeral => (bool * term list) option"}])) = T
+  @{typ "natural => (bool * term list) option"}])) = T
   | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 
 fun mk_empty T = Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_empty}, mk_monadT T);
@@ -143,7 +143,7 @@
 fun mk_not t =
   let
     val nT = @{typ "(unit Quickcheck_Exhaustive.unknown =>
-      Code_Evaluation.term list Quickcheck_Exhaustive.three_valued) => code_numeral =>
+      Code_Evaluation.term list Quickcheck_Exhaustive.three_valued) => natural =>
       Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"}
     val T = mk_monadT HOLogic.unitT
   in Const (@{const_name Quickcheck_Exhaustive.pos_bound_cps_not}, nT --> T) $ t end
@@ -169,11 +169,11 @@
 fun mk_monadT T =
   (Type (@{type_name "Quickcheck_Exhaustive.unknown"}, [T])
     --> @{typ "Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"})
-    --> @{typ "code_numeral => Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"}
+    --> @{typ "natural => Code_Evaluation.term list Quickcheck_Exhaustive.three_valued"}
 
 fun dest_monadT (Type ("fun", [Type ("fun", [Type (@{type_name "Quickcheck_Exhaustive.unknown"}, [T]),
     @{typ "term list Quickcheck_Exhaustive.three_valued"}]),
-    @{typ "code_numeral => term list Quickcheck_Exhaustive.three_valued"}])) = T
+    @{typ "natural => term list Quickcheck_Exhaustive.three_valued"}])) = T
   | dest_monadT T = raise TYPE ("dest_monadT", [T], []);
 
 fun mk_empty T = Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_empty}, mk_monadT T);
@@ -199,7 +199,7 @@
   let
     val T = mk_monadT HOLogic.unitT
     val pT = @{typ "(unit => (bool * Code_Evaluation.term list) option)"}
-      --> @{typ "code_numeral => (bool * Code_Evaluation.term list) option"}
+      --> @{typ "natural => (bool * Code_Evaluation.term list) option"}
   in Const (@{const_name Quickcheck_Exhaustive.neg_bound_cps_not}, pT --> T) $ t end
 
 fun mk_Enum _ = error "not implemented"
@@ -251,7 +251,7 @@
 
 fun mk_iterate_upto T (f, from, to) =
   list_comb (Const (@{const_name Random_Pred.iterate_upto},
-      [@{typ code_numeral} --> T, @{typ code_numeral}, @{typ code_numeral}] ---> mk_randompredT T),
+      [@{typ natural} --> T, @{typ natural}, @{typ natural}] ---> mk_randompredT T),
     [f, from, to])
 
 fun mk_not t =
@@ -272,10 +272,10 @@
 structure DSequence_CompFuns =
 struct
 
-fun mk_dseqT T = Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ bool},
+fun mk_dseqT T = Type ("fun", [@{typ natural}, Type ("fun", [@{typ bool},
   Type (@{type_name Option.option}, [Type  (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])
 
-fun dest_dseqT (Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ bool},
+fun dest_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ bool},
   Type (@{type_name Option.option}, [Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])) = T
   | dest_dseqT T = raise TYPE ("dest_dseqT", [T], []);
 
@@ -315,9 +315,9 @@
 struct
 
 fun mk_pos_dseqT T =
-    @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
+    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
 
-fun dest_pos_dseqT (Type ("fun", [@{typ code_numeral},
+fun dest_pos_dseqT (Type ("fun", [@{typ natural},
     Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])) = T
   | dest_pos_dseqT T = raise TYPE ("dest_pos_dseqT", [T], []);
 
@@ -353,7 +353,7 @@
   let
     val pT = mk_pos_dseqT HOLogic.unitT
     val nT =
-      @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
+      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
         [Type (@{type_name Option.option}, [@{typ unit}])])
   in Const (@{const_name Limited_Sequence.pos_not_seq}, nT --> pT) $ t end
 
@@ -375,10 +375,10 @@
 structure New_Neg_DSequence_CompFuns =
 struct
 
-fun mk_neg_dseqT T = @{typ code_numeral} -->
+fun mk_neg_dseqT T = @{typ natural} -->
   Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])
 
-fun dest_neg_dseqT (Type ("fun", [@{typ code_numeral},
+fun dest_neg_dseqT (Type ("fun", [@{typ natural},
     Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])])) = T
   | dest_neg_dseqT T = raise TYPE ("dest_neg_dseqT", [T], []);
 
@@ -414,7 +414,7 @@
   let
     val nT = mk_neg_dseqT HOLogic.unitT
     val pT =
-      @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
+      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
         [@{typ unit}])
   in Const (@{const_name Limited_Sequence.neg_not_seq}, pT --> nT) $ t end
 
@@ -437,11 +437,11 @@
 struct
 
 fun mk_pos_random_dseqT T =
-  @{typ code_numeral} --> @{typ code_numeral} --> @{typ Random.seed} -->
-    @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
+  @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
+    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
 
-fun dest_pos_random_dseqT (Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ code_numeral},
-    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ code_numeral},
+fun dest_pos_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
+    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ natural},
     Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])])) = T
   | dest_pos_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
 
@@ -473,15 +473,15 @@
 
 fun mk_iterate_upto T (f, from, to) =
   list_comb (Const (@{const_name Random_Sequence.pos_iterate_upto},
-      [@{typ code_numeral} --> T, @{typ code_numeral}, @{typ code_numeral}]
+      [@{typ natural} --> T, @{typ natural}, @{typ natural}]
         ---> mk_pos_random_dseqT T),
     [f, from, to])
 
 fun mk_not t =
   let
     val pT = mk_pos_random_dseqT HOLogic.unitT
-    val nT = @{typ code_numeral} --> @{typ code_numeral} --> @{typ Random.seed} -->
-      @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
+    val nT = @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
+      @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
         [Type (@{type_name Option.option}, [@{typ unit}])])
 
   in Const (@{const_name Random_Sequence.pos_not_random_dseq}, nT --> pT) $ t end
@@ -504,12 +504,12 @@
 struct
 
 fun mk_neg_random_dseqT T =
-   @{typ code_numeral} --> @{typ code_numeral} --> @{typ Random.seed} -->
-    @{typ code_numeral} --> 
+   @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
+    @{typ natural} --> 
     Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])
 
-fun dest_neg_random_dseqT (Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ code_numeral},
-    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ code_numeral},
+fun dest_neg_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
+    Type ("fun", [@{typ Random.seed}, Type ("fun", [@{typ natural},
       Type (@{type_name Lazy_Sequence.lazy_sequence},
         [Type (@{type_name Option.option}, [T])])])])])])) = T
   | dest_neg_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
@@ -542,15 +542,15 @@
 
 fun mk_iterate_upto T (f, from, to) =
   list_comb (Const (@{const_name Random_Sequence.neg_iterate_upto},
-      [@{typ code_numeral} --> T, @{typ code_numeral}, @{typ code_numeral}]
+      [@{typ natural} --> T, @{typ natural}, @{typ natural}]
         ---> mk_neg_random_dseqT T),
     [f, from, to])
 
 fun mk_not t =
   let
     val nT = mk_neg_random_dseqT HOLogic.unitT
-    val pT = @{typ code_numeral} --> @{typ code_numeral} --> @{typ Random.seed} -->
-    @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [@{typ unit}])
+    val pT = @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
+    @{typ natural} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [@{typ unit}])
   in Const (@{const_name Random_Sequence.neg_not_random_dseq}, pT --> nT) $ t end
 
 fun mk_map T1 T2 tf tp = Const (@{const_name Random_Sequence.neg_map},
@@ -572,10 +572,10 @@
 struct
 
 fun mk_random_dseqT T =
-  @{typ code_numeral} --> @{typ code_numeral} --> @{typ Random.seed} -->
+  @{typ natural} --> @{typ natural} --> @{typ Random.seed} -->
     HOLogic.mk_prodT (DSequence_CompFuns.mk_dseqT T, @{typ Random.seed})
 
-fun dest_random_dseqT (Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ code_numeral},
+fun dest_random_dseqT (Type ("fun", [@{typ natural}, Type ("fun", [@{typ natural},
   Type ("fun", [@{typ Random.seed},
   Type (@{type_name Product_Type.prod}, [T, @{typ Random.seed}])])])])) = DSequence_CompFuns.dest_dseqT T
   | dest_random_dseqT T = raise TYPE ("dest_random_dseqT", [T], []);
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -27,15 +27,15 @@
   val put_pred_random_result : (unit -> seed -> term Predicate.pred * seed) ->
     Proof.context -> Proof.context
   val put_dseq_result : (unit -> term Limited_Sequence.dseq) -> Proof.context -> Proof.context
-  val put_dseq_random_result : (unit -> int -> int -> seed -> term Limited_Sequence.dseq * seed) ->
+  val put_dseq_random_result : (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term Limited_Sequence.dseq * seed) ->
     Proof.context -> Proof.context
-  val put_new_dseq_result : (unit -> int -> term Lazy_Sequence.lazy_sequence) ->
+  val put_new_dseq_result : (unit -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence) ->
     Proof.context -> Proof.context
   val put_lseq_random_result :
-    (unit -> int -> int -> seed -> int -> term Lazy_Sequence.lazy_sequence) ->
+    (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence) ->
     Proof.context -> Proof.context
   val put_lseq_random_stats_result :
-    (unit -> int -> int -> seed -> int -> (term * int) Lazy_Sequence.lazy_sequence) ->
+    (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> (term * Code_Numeral.natural) Lazy_Sequence.lazy_sequence) ->
     Proof.context -> Proof.context
 
   val code_pred_intro_attrib : attribute
@@ -66,6 +66,8 @@
 structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
 struct
 
+type random_seed = Random_Engine.seed
+
 open Predicate_Compile_Aux;
 open Core_Data;
 open Mode_Inference;
@@ -294,9 +296,9 @@
   additional_arguments = fn names =>
     let
       val depth_name = singleton (Name.variant_list names) "depth"
-    in [Free (depth_name, @{typ code_numeral})] end,
+    in [Free (depth_name, @{typ natural})] end,
   modify_funT = (fn T => let val (Ts, U) = strip_type T
-  val Ts' = [@{typ code_numeral}] in (Ts @ Ts') ---> U end),
+  val Ts' = [@{typ natural}] in (Ts @ Ts') ---> U end),
   wrap_compilation =
     fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
     let
@@ -305,7 +307,7 @@
       val T' = mk_monadT compfuns (HOLogic.mk_tupleT Ts)
       val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
     in
-      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
+      if_const $ HOLogic.mk_eq (depth, @{term "0 :: natural"})
         $ mk_empty compfuns (dest_monadT compfuns T')
         $ compilation
     end,
@@ -314,8 +316,8 @@
     let
       val [depth] = additional_arguments
       val depth' =
-        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
-          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
+        Const (@{const_name Groups.minus}, @{typ "natural => natural => natural"})
+          $ depth $ Const (@{const_name Groups.one}, @{typ "natural"})
     in [depth'] end
   }
 
@@ -326,18 +328,18 @@
   compfuns = Predicate_Comp_Funs.compfuns,
   mk_random = (fn T => fn additional_arguments =>
   list_comb (Const(@{const_name Random_Pred.iter},
-  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
+  [@{typ natural}, @{typ natural}, @{typ Random.seed}] ---> 
     Predicate_Comp_Funs.mk_monadT T), additional_arguments)),
   modify_funT = (fn T =>
     let
       val (Ts, U) = strip_type T
-      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}]
+      val Ts' = [@{typ natural}, @{typ natural}, @{typ Random.seed}]
     in (Ts @ Ts') ---> U end),
   additional_arguments = (fn names =>
     let
       val [nrandom, size, seed] = Name.variant_list names ["nrandom", "size", "seed"]
     in
-      [Free (nrandom, @{typ code_numeral}), Free (size, @{typ code_numeral}),
+      [Free (nrandom, @{typ natural}), Free (size, @{typ natural}),
         Free (seed, @{typ Random.seed})]
     end),
   wrap_compilation = K (K (K (K (K I))))
@@ -352,20 +354,20 @@
   compfuns = Predicate_Comp_Funs.compfuns,
   mk_random = (fn T => fn additional_arguments =>
   list_comb (Const(@{const_name Random_Pred.iter},
-  [@{typ code_numeral}, @{typ code_numeral}, @{typ Random.seed}] ---> 
+  [@{typ natural}, @{typ natural}, @{typ Random.seed}] ---> 
     Predicate_Comp_Funs.mk_monadT T), tl additional_arguments)),
   modify_funT = (fn T =>
     let
       val (Ts, U) = strip_type T
-      val Ts' = [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
+      val Ts' = [@{typ natural}, @{typ natural}, @{typ natural},
         @{typ Random.seed}]
     in (Ts @ Ts') ---> U end),
   additional_arguments = (fn names =>
     let
       val [depth, nrandom, size, seed] = Name.variant_list names ["depth", "nrandom", "size", "seed"]
     in
-      [Free (depth, @{typ code_numeral}), Free (nrandom, @{typ code_numeral}),
-        Free (size, @{typ code_numeral}), Free (seed, @{typ Random.seed})]
+      [Free (depth, @{typ natural}), Free (nrandom, @{typ natural}),
+        Free (size, @{typ natural}), Free (seed, @{typ Random.seed})]
     end),
   wrap_compilation =
   fn compfuns => fn _ => fn T => fn mode => fn additional_arguments => fn compilation =>
@@ -376,7 +378,7 @@
       val T' = mk_monadT compfuns (HOLogic.mk_tupleT Ts)
       val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
     in
-      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
+      if_const $ HOLogic.mk_eq (depth, @{term "0 :: natural"})
         $ mk_empty compfuns (dest_monadT compfuns T')
         $ compilation
     end,
@@ -385,8 +387,8 @@
     let
       val [depth, nrandom, size, seed] = additional_arguments
       val depth' =
-        Const (@{const_name Groups.minus}, @{typ "code_numeral => code_numeral => code_numeral"})
-          $ depth $ Const (@{const_name Groups.one}, @{typ "Code_Numeral.code_numeral"})
+        Const (@{const_name Groups.minus}, @{typ "natural => natural => natural"})
+          $ depth $ Const (@{const_name Groups.one}, @{typ "natural"})
     in [depth', nrandom, size, seed] end
 }
 
@@ -424,10 +426,10 @@
   mk_random = (fn T => fn _ =>
   let
     val random = Const (@{const_name Quickcheck_Random.random},
-      @{typ code_numeral} --> @{typ Random.seed} -->
+      @{typ natural} --> @{typ Random.seed} -->
         HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
   in
-    Const (@{const_name Random_Sequence.Random}, (@{typ code_numeral} --> @{typ Random.seed} -->
+    Const (@{const_name Random_Sequence.Random}, (@{typ natural} --> @{typ Random.seed} -->
       HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
       Random_Sequence_CompFuns.mk_random_dseqT T) $ random
   end),
@@ -461,10 +463,10 @@
   mk_random = (fn T => fn _ =>
   let
     val random = Const (@{const_name Quickcheck_Random.random},
-      @{typ code_numeral} --> @{typ Random.seed} -->
+      @{typ natural} --> @{typ Random.seed} -->
         HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed}))
   in
-    Const (@{const_name Random_Sequence.pos_Random}, (@{typ code_numeral} --> @{typ Random.seed} -->
+    Const (@{const_name Random_Sequence.pos_Random}, (@{typ natural} --> @{typ Random.seed} -->
       HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
       New_Pos_Random_Sequence_CompFuns.mk_pos_random_dseqT T) $ random
   end),
@@ -496,7 +498,7 @@
   mk_random =
     (fn T => fn _ =>
       Const (@{const_name "Lazy_Sequence.small_lazy_class.small_lazy"},
-        @{typ "Code_Numeral.code_numeral"} --> Type (@{type_name "Lazy_Sequence.lazy_sequence"}, [T]))),
+        @{typ "natural"} --> Type (@{type_name "Lazy_Sequence.lazy_sequence"}, [T]))),
   modify_funT = I,
   additional_arguments = K [],
   wrap_compilation = K (K (K (K (K I))))
@@ -526,7 +528,7 @@
     (fn T => fn _ =>
        Const (@{const_name "Quickcheck_Exhaustive.exhaustive"},
        (T --> @{typ "(bool * term list) option"}) -->
-         @{typ "code_numeral => (bool * term list) option"})),
+         @{typ "natural => (bool * term list) option"})),
   modify_funT = I,
   additional_arguments = K [],
   wrap_compilation = K (K (K (K (K I))))
@@ -1655,7 +1657,7 @@
 
 structure Dseq_Random_Result = Proof_Data
 (
-  type T = unit -> int -> int -> seed -> term Limited_Sequence.dseq * seed
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term Limited_Sequence.dseq * seed
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Dseq_Random_Result"
 );
@@ -1663,7 +1665,7 @@
 
 structure New_Dseq_Result = Proof_Data
 (
-  type T = unit -> int -> term Lazy_Sequence.lazy_sequence
+  type T = unit -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "New_Dseq_Random_Result"
 );
@@ -1671,7 +1673,7 @@
 
 structure Lseq_Random_Result = Proof_Data
 (
-  type T = unit -> int -> int -> seed -> int -> term Lazy_Sequence.lazy_sequence
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> term Lazy_Sequence.lazy_sequence
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Lseq_Random_Result"
 );
@@ -1679,7 +1681,7 @@
 
 structure Lseq_Random_Stats_Result = Proof_Data
 (
-  type T = unit -> int -> int -> seed -> int -> (term * int) Lazy_Sequence.lazy_sequence
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> (term * Code_Numeral.natural) Lazy_Sequence.lazy_sequence
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Lseq_Random_Stats_Result"
 );
@@ -1795,7 +1797,7 @@
         fun count' i [] = i
           | count' i (x' :: xs) = if x = x' then count' (i + 1) xs else count' i xs
       in count' 0 xs end
-    fun accumulate xs = map (fn x => (x, count xs x)) (sort int_ord (distinct (op =) xs))
+    fun accumulate xs = (map (fn x => (x, count xs x)) o sort int_ord o distinct (op =)) xs;
     val comp_modifiers =
       case compilation of
           Pred => predicate_comp_modifiers
@@ -1811,12 +1813,12 @@
     val additional_arguments =
       case compilation of
         Pred => []
-      | Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
-        [@{term "(1, 1) :: code_numeral * code_numeral"}]
+      | Random => map (HOLogic.mk_number @{typ "natural"}) arguments @
+        [@{term "(1, 1) :: natural * natural"}]
       | Annotated => []
-      | Depth_Limited => [HOLogic.mk_number @{typ "code_numeral"} (hd arguments)]
-      | Depth_Limited_Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
-        [@{term "(1, 1) :: code_numeral * code_numeral"}]
+      | Depth_Limited => [HOLogic.mk_number @{typ "natural"} (hd arguments)]
+      | Depth_Limited_Random => map (HOLogic.mk_number @{typ "natural"}) arguments @
+        [@{term "(1, 1) :: natural * natural"}]
       | DSeq => []
       | Pos_Random_DSeq => []
       | New_Pos_Random_DSeq => []
@@ -1825,9 +1827,9 @@
     val T = dest_monadT compfuns (fastype_of t);
     val t' =
       if stats andalso compilation = New_Pos_Random_DSeq then
-        mk_map compfuns T (HOLogic.mk_prodT (HOLogic.termT, @{typ code_numeral}))
+        mk_map compfuns T (HOLogic.mk_prodT (HOLogic.termT, @{typ natural}))
           (absdummy T (HOLogic.mk_prod (HOLogic.term_of_const T $ Bound 0,
-            @{term Code_Numeral.of_nat} $ (HOLogic.size_const T $ Bound 0)))) t
+            @{term natural_of_nat} $ (HOLogic.size_const T $ Bound 0)))) t
       else
         mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t
     val thy = Proof_Context.theory_of ctxt
@@ -1841,7 +1843,7 @@
             |> Random_Engine.run))*)
         Pos_Random_DSeq =>
           let
-            val [nrandom, size, depth] = arguments
+            val [nrandom, size, depth] = map Code_Numeral.natural_of_integer arguments
           in
             rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Limited_Sequence.yieldn k
               (Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
@@ -1853,10 +1855,10 @@
       | DSeq =>
           rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Limited_Sequence.yieldn k
             (Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
-              thy NONE Limited_Sequence.map t' []) (the_single arguments) true)) ())
+              thy NONE Limited_Sequence.map t' []) (Code_Numeral.natural_of_integer (the_single arguments)) true)) ())
       | Pos_Generator_DSeq =>
           let
-            val depth = (the_single arguments)
+            val depth = Code_Numeral.natural_of_integer (the_single arguments)
           in
             rpair NONE (TimeLimit.timeLimit time_limit (fn () => fst (Lazy_Sequence.yieldn k
               (Code_Runtime.dynamic_value_strict (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Core.put_new_dseq_result")
@@ -1865,11 +1867,12 @@
           end
       | New_Pos_Random_DSeq =>
           let
-            val [nrandom, size, depth] = arguments
+            val [nrandom, size, depth] = map Code_Numeral.natural_of_integer arguments
             val seed = Random_Engine.next_seed ()
           in
             if stats then
-              apsnd (SOME o accumulate) (split_list (TimeLimit.timeLimit time_limit
+              apsnd (SOME o accumulate o map Code_Numeral.integer_of_natural)
+              (split_list (TimeLimit.timeLimit time_limit
               (fn () => fst (Lazy_Sequence.yieldn k
                 (Code_Runtime.dynamic_value_strict
                   (Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result")
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -9,17 +9,17 @@
   type seed = Random_Engine.seed
   (*val quickcheck : Proof.context -> term -> int -> term list option*)
   val put_pred_result :
-    (unit -> int -> int -> int -> seed -> term list Predicate.pred) ->
+    (unit -> Code_Numeral.natural -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term list Predicate.pred) ->
       Proof.context -> Proof.context;
   val put_dseq_result :
-    (unit -> int -> int -> seed -> term list Limited_Sequence.dseq * seed) ->
+    (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term list Limited_Sequence.dseq * seed) ->
       Proof.context -> Proof.context;
   val put_lseq_result :
-    (unit -> int -> int -> seed -> int -> term list Lazy_Sequence.lazy_sequence) ->
+    (unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence) ->
       Proof.context -> Proof.context;
-  val put_new_dseq_result : (unit -> int -> term list Lazy_Sequence.lazy_sequence) ->
+  val put_new_dseq_result : (unit -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence) ->
     Proof.context -> Proof.context
-  val put_cps_result : (unit -> int -> (bool * term list) option) ->
+  val put_cps_result : (unit -> Code_Numeral.natural -> (bool * term list) option) ->
     Proof.context -> Proof.context
   val test_goals : (Predicate_Compile_Aux.compilation * bool) ->
     Proof.context -> bool * bool -> (string * typ) list -> (term * term list) list
@@ -41,7 +41,7 @@
 
 structure Pred_Result = Proof_Data
 (
-  type T = unit -> int -> int -> int -> seed -> term list Predicate.pred
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term list Predicate.pred
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Pred_Result"
 );
@@ -49,7 +49,7 @@
 
 structure Dseq_Result = Proof_Data
 (
-  type T = unit -> int -> int -> seed -> term list Limited_Sequence.dseq * seed
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> term list Limited_Sequence.dseq * seed
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Dseq_Result"
 );
@@ -57,7 +57,7 @@
 
 structure Lseq_Result = Proof_Data
 (
-  type T = unit -> int -> int -> seed -> int -> term list Lazy_Sequence.lazy_sequence
+  type T = unit -> Code_Numeral.natural -> Code_Numeral.natural -> seed -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Lseq_Result"
 );
@@ -65,7 +65,7 @@
 
 structure New_Dseq_Result = Proof_Data
 (
-  type T = unit -> int -> term list Lazy_Sequence.lazy_sequence
+  type T = unit -> Code_Numeral.natural -> term list Lazy_Sequence.lazy_sequence
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "New_Dseq_Random_Result"
 );
@@ -73,7 +73,7 @@
 
 structure CPS_Result = Proof_Data
 (
-  type T = unit -> int -> (bool * term list) option
+  type T = unit -> Code_Numeral.natural -> (bool * term list) option
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "CPS_Result"
 );
@@ -141,7 +141,7 @@
     show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
     show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
     show_invalid_clauses = s_ic, skip_proof = s_p,
-    compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, 
+    compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = _, 
     fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
     smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
@@ -158,7 +158,7 @@
     show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
     show_invalid_clauses = s_ic, skip_proof = s_p,
     compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, 
-    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    fail_safe_function_flattening = _, no_higher_order_predicate = no_ho,
     smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
     show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
@@ -174,7 +174,7 @@
     show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
     show_invalid_clauses = s_ic, skip_proof = s_p,
     compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f, 
-    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = _,
     smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
   (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
     show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
@@ -214,10 +214,6 @@
 
 val mk_cpsT =
   Predicate_Compile_Aux.mk_monadT Pos_Bounded_CPS_Comp_Funs.compfuns
-val mk_cps_return =
-  Predicate_Compile_Aux.mk_single Pos_Bounded_CPS_Comp_Funs.compfuns
-val mk_cps_bind =
-  Predicate_Compile_Aux.mk_bind Pos_Bounded_CPS_Comp_Funs.compfuns
 
 val mk_split_lambda = HOLogic.tupled_lambda o HOLogic.mk_tuple
 
@@ -232,9 +228,9 @@
     val ((((full_constname, constT), vs'), intro), thy1) =
       Predicate_Compile_Aux.define_quickcheck_predicate t' thy
     val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
-    val (thy3, preproc_time) = cpu_time "predicate preprocessing"
+    val (thy3, _) = cpu_time "predicate preprocessing"
         (fn () => Predicate_Compile.preprocess options (Const (full_constname, constT)) thy2)
-    val (thy4, core_comp_time) = cpu_time "random_dseq core compilation"
+    val (thy4, _) = cpu_time "random_dseq core compilation"
         (fn () =>
           case compilation of
             Pos_Random_DSeq =>
@@ -261,7 +257,7 @@
             | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs'))
             | Pos_Generator_DSeq => mk_new_dseqT (HOLogic.mk_tupleT (map snd vs'))
             | Depth_Limited_Random =>
-              [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
+              [@{typ natural}, @{typ natural}, @{typ natural},
               @{typ Random.seed}] ---> mk_predT (HOLogic.mk_tupleT (map snd vs'))
             | Pos_Generator_CPS => mk_cpsT (HOLogic.mk_tupleT (map snd vs'))
         in
@@ -285,7 +281,7 @@
             HOLogic.mk_prod (@{term "True"}, HOLogic.mk_list @{typ term}
                 (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))     
         | Depth_Limited_Random => fold_rev absdummy
-            [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
+            [@{typ natural}, @{typ natural}, @{typ natural},
              @{typ Random.seed}]
             (mk_bind' (list_comb (prog, map Bound (3 downto 0)),
             mk_split_lambda (map Free vs') (mk_return' (HOLogic.mk_list @{typ term}
@@ -340,7 +336,7 @@
                 (fn proc => fn g => fn depth => g depth |> Option.map (apsnd (map proc)))
                 qc_term []
           in
-            fn size => fn nrandom => Option.map snd o compiled_term
+            fn _ => fn _ => Option.map snd o compiled_term
           end
        | Depth_Limited_Random =>
           let
@@ -382,11 +378,12 @@
 
 (* quickcheck interface functions *)
 
-fun compile_term' compilation options ctxt (t, eval_terms) =
+fun compile_term' compilation options ctxt (t, _) =
   let
     val size = Config.get ctxt Quickcheck.size
     val c = compile_term compilation options ctxt t
-    val counterexample = try_upto_depth ctxt (c size (!nrandom))
+    val counterexample = try_upto_depth ctxt (c (Code_Numeral.natural_of_integer size)
+      (Code_Numeral.natural_of_integer (!nrandom)) o Code_Numeral.natural_of_integer)
   in
     Quickcheck.Result
       {counterexample = Option.map (pair true o (curry (op ~~)) (Term.add_free_names t [])) counterexample,
@@ -403,7 +400,7 @@
   end
 
 
-fun test_goals options ctxt catch_code_errors insts goals =
+fun test_goals options ctxt _ insts goals =
   let
     val (compilation, fail_safe_function_flattening) = options
     val function_flattening = Config.get ctxt (Quickcheck.allow_function_inversion)
--- a/src/HOL/Tools/Qelim/cooper.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -267,23 +267,23 @@
   end
  | _ => addC $ (mulC $ one $ tm) $ zero;
 
-fun lin (vs as x::_) (Const (@{const_name Not}, _) $ (Const (@{const_name Orderings.less}, T) $ s $ t)) =
+fun lin (vs as _::_) (Const (@{const_name Not}, _) $ (Const (@{const_name Orderings.less}, T) $ s $ t)) =
     lin vs (Const (@{const_name Orderings.less_eq}, T) $ t $ s)
-  | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name Orderings.less_eq}, T) $ s $ t)) =
+  | lin (vs as _::_) (Const (@{const_name Not},_) $ (Const(@{const_name Orderings.less_eq}, T) $ s $ t)) =
     lin vs (Const (@{const_name Orderings.less}, T) $ t $ s)
   | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
-  | lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) =
+  | lin (vs as _::_) (Const(@{const_name Rings.dvd},_)$d$t) =
     HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t)
   | lin (vs as x::_) ((b as Const(@{const_name HOL.eq},_))$s$t) =
      (case lint vs (subC$t$s) of
-      (t as a$(m$c$y)$r) =>
+      (t as _$(m$c$y)$r) =>
         if x <> y then b$zero$t
         else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r
         else b$(m$c$y)$(linear_neg r)
       | t => b$zero$t)
   | lin (vs as x::_) (b$s$t) =
      (case lint vs (subC$t$s) of
-      (t as a$(m$c$y)$r) =>
+      (t as _$(m$c$y)$r) =>
         if x <> y then b$zero$t
         else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r
         else b$(linear_neg r)$(m$c$y)
@@ -303,7 +303,7 @@
   | is_intrel _ = false;
 
 fun linearize_conv ctxt vs ct = case term_of ct of
-  Const(@{const_name Rings.dvd},_)$d$t =>
+  Const(@{const_name Rings.dvd},_)$_$_ =>
   let
     val th = Conv.binop_conv (lint_conv ctxt vs) ct
     val (d',t') = Thm.dest_binop (Thm.rhs_of th)
@@ -565,9 +565,9 @@
   Qelim.gen_qelim_conv (Simplifier.rewrite conv_ss) (Simplifier.rewrite presburger_ss) (Simplifier.rewrite conv_ss)
     (cons o term_of) (Misc_Legacy.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt)
     (cooperex_conv ctxt) p
-  handle CTERM s => raise COOPER "bad cterm"
-       | THM s => raise COOPER "bad thm"
-       | TYPE s => raise COOPER "bad type"
+  handle CTERM _ => raise COOPER "bad cterm"
+       | THM _ => raise COOPER "bad thm"
+       | TYPE _ => raise COOPER "bad type"
 
 fun add_bools t =
   let
@@ -593,14 +593,14 @@
 
 local structure Proc = Cooper_Procedure in
 
-fun num_of_term vs (Free vT) = Proc.Bound (find_index (fn vT' => vT' = vT) vs)
-  | num_of_term vs (Term.Bound i) = Proc.Bound i
-  | num_of_term vs @{term "0::int"} = Proc.C 0
-  | num_of_term vs @{term "1::int"} = Proc.C 1
+fun num_of_term vs (Free vT) = Proc.Bound (Proc.nat_of_integer (find_index (fn vT' => vT' = vT) vs))
+  | num_of_term vs (Term.Bound i) = Proc.Bound (Proc.nat_of_integer i)
+  | num_of_term vs @{term "0::int"} = Proc.C (Proc.Int_of_integer 0)
+  | num_of_term vs @{term "1::int"} = Proc.C (Proc.Int_of_integer 1)
   | num_of_term vs (t as Const (@{const_name numeral}, _) $ _) =
-      Proc.C (dest_number t)
+      Proc.C (Proc.Int_of_integer (dest_number t))
   | num_of_term vs (t as Const (@{const_name neg_numeral}, _) $ _) =
-      Proc.Neg (Proc.C (dest_number t))
+      Proc.Neg (Proc.C (Proc.Int_of_integer (dest_number t)))
   | num_of_term vs (Const (@{const_name Groups.uminus}, _) $ t') =
       Proc.Neg (num_of_term vs t')
   | num_of_term vs (Const (@{const_name Groups.plus}, _) $ t1 $ t2) =
@@ -609,9 +609,9 @@
       Proc.Sub (num_of_term vs t1, num_of_term vs t2)
   | num_of_term vs (Const (@{const_name Groups.times}, _) $ t1 $ t2) =
      (case perhaps_number t1
-       of SOME n => Proc.Mul (n, num_of_term vs t2)
+       of SOME n => Proc.Mul (Proc.Int_of_integer n, num_of_term vs t2)
         | NONE => (case perhaps_number t2
-           of SOME n => Proc.Mul (n, num_of_term vs t1)
+           of SOME n => Proc.Mul (Proc.Int_of_integer n, num_of_term vs t1)
             | NONE => raise COOPER "reification: unsupported kind of multiplication"))
   | num_of_term _ _ = raise COOPER "reification: bad term";
 
@@ -639,13 +639,13 @@
       Proc.Lt (Proc.Sub (num_of_term vs t1, num_of_term vs t2))
   | fm_of_term ps vs (Const (@{const_name Rings.dvd}, _) $ t1 $ t2) =
      (case perhaps_number t1
-       of SOME n => Proc.Dvd (n, num_of_term vs t2)
+       of SOME n => Proc.Dvd (Proc.Int_of_integer n, num_of_term vs t2)
         | NONE => raise COOPER "reification: unsupported dvd")
   | fm_of_term ps vs t = let val n = find_index (fn t' => t aconv t') ps
-      in if n > 0 then Proc.Closed n else raise COOPER "reification: unknown term" end;
+      in if n > 0 then Proc.Closed (Proc.nat_of_integer n) else raise COOPER "reification: unknown term" end;
 
-fun term_of_num vs (Proc.C i) = HOLogic.mk_number HOLogic.intT i
-  | term_of_num vs (Proc.Bound n) = Free (nth vs n)
+fun term_of_num vs (Proc.C i) = HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i)
+  | term_of_num vs (Proc.Bound n) = Free (nth vs (Proc.integer_of_nat n))
   | term_of_num vs (Proc.Neg t') =
       @{term "uminus :: int => _"} $ term_of_num vs t'
   | term_of_num vs (Proc.Add (t1, t2)) =
@@ -653,7 +653,7 @@
   | term_of_num vs (Proc.Sub (t1, t2)) =
       @{term "op - :: int => _"} $ term_of_num vs t1 $ term_of_num vs t2
   | term_of_num vs (Proc.Mul (i, t2)) =
-      @{term "op * :: int => _"} $ HOLogic.mk_number HOLogic.intT i $ term_of_num vs t2
+      @{term "op * :: int => _"} $ HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i) $ term_of_num vs t2
   | term_of_num vs (Proc.Cn (n, i, t')) =
       term_of_num vs (Proc.Add (Proc.Mul (i, Proc.Bound n), t'));
 
@@ -671,9 +671,9 @@
   | term_of_fm ps vs (Proc.Gt t') = @{term "op < :: int => _ "} $ @{term "0::int"} $ term_of_num vs t'
   | term_of_fm ps vs (Proc.Ge t') = @{term "op <= :: int => _ "} $ @{term "0::int"} $ term_of_num vs t'
   | term_of_fm ps vs (Proc.Dvd (i, t')) = @{term "op dvd :: int => _ "} $
-      HOLogic.mk_number HOLogic.intT i $ term_of_num vs t'
+      HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i) $ term_of_num vs t'
   | term_of_fm ps vs (Proc.NDvd (i, t')) = term_of_fm ps vs (Proc.Not (Proc.Dvd (i, t')))
-  | term_of_fm ps vs (Proc.Closed n) = nth ps n
+  | term_of_fm ps vs (Proc.Closed n) = nth ps (Proc.integer_of_nat n)
   | term_of_fm ps vs (Proc.NClosed n) = term_of_fm ps vs (Proc.Not (Proc.Closed n));
 
 fun procedure t =
@@ -701,7 +701,7 @@
 
 fun strip_objall ct = 
  case term_of ct of 
-  Const (@{const_name All}, _) $ Abs (xn,xT,p) => 
+  Const (@{const_name All}, _) $ Abs (xn,_,_) => 
    let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
    in apfst (cons (a,v)) (strip_objall t')
    end
@@ -782,7 +782,7 @@
  in h [] ct end
 end;
 
-fun generalize_tac f = CSUBGOAL (fn (p, i) => PRIMITIVE (fn st =>
+fun generalize_tac f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st =>
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.apply (all (ctyp_of_term x)) (Thm.lambda x t)
--- a/src/HOL/Tools/Qelim/cooper_procedure.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Qelim/cooper_procedure.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -1,57 +1,60 @@
 (* Generated from Cooper.thy; DO NOT EDIT! *)
 
 structure Cooper_Procedure : sig
+  val id : 'a -> 'a
   type 'a equal
   val equal : 'a equal -> 'a -> 'a -> bool
   val eq : 'a equal -> 'a -> 'a -> bool
-  val suc : int -> int
-  datatype num = C of int | Bound of int | Cn of int * int * num | Neg of num |
-    Add of num * num | Sub of num * num | Mul of int * num
-  datatype fm = T | F | Lt of num | Le of num | Gt of num | Ge of num |
-    Eq of num | NEq of num | Dvd of int * num | NDvd of int * num | Not of fm |
-    And of fm * fm | Or of fm * fm | Imp of fm * fm | Iff of fm * fm | E of fm |
-    A of fm | Closed of int | NClosed of int
+  datatype inta = Int_of_integer of int
+  datatype nat = Nat of int
+  datatype num = One | Bit0 of num | Bit1 of num
+  type 'a ord
+  val less_eq : 'a ord -> 'a -> 'a -> bool
+  val less : 'a ord -> 'a -> 'a -> bool
+  val ord_integer : int ord
+  val max : 'a ord -> 'a -> 'a -> 'a
+  val nat_of_integer : int -> nat
+  val integer_of_nat : nat -> int
+  val plus_nat : nat -> nat -> nat
+  val suc : nat -> nat
+  datatype numa = C of inta | Bound of nat | Cn of nat * inta * numa |
+    Neg of numa | Add of numa * numa | Sub of numa * numa | Mul of inta * numa
+  datatype fm = T | F | Lt of numa | Le of numa | Gt of numa | Ge of numa |
+    Eq of numa | NEq of numa | Dvd of inta * numa | NDvd of inta * numa |
+    Not of fm | And of fm * fm | Or of fm * fm | Imp of fm * fm | Iff of fm * fm
+    | E of fm | A of fm | Closed of nat | NClosed of nat
   val map : ('a -> 'b) -> 'a list -> 'b list
-  val equal_numa : num -> num -> bool
+  val disjuncts : fm -> fm list
+  val foldr : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
+  val equal_nat : nat -> nat -> bool
+  val integer_of_int : inta -> int
+  val equal_inta : inta -> inta -> bool
+  val equal_numa : numa -> numa -> bool
   val equal_fm : fm -> fm -> bool
   val djf : ('a -> fm) -> 'a -> fm -> fm
-  val foldr : ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b
   val evaldjf : ('a -> fm) -> 'a list -> fm
-  val disjuncts : fm -> fm list
   val dj : (fm -> fm) -> fm -> fm
-  val prep : fm -> fm
-  val conj : fm -> fm -> fm
-  val disj : fm -> fm -> fm
-  val nota : fm -> fm
-  val iffa : fm -> fm -> fm
-  val impa : fm -> fm -> fm
-  type 'a times
-  val times : 'a times -> 'a -> 'a -> 'a
-  type 'a dvd
-  val times_dvd : 'a dvd -> 'a times
-  type 'a diva
-  val dvd_div : 'a diva -> 'a dvd
-  val diva : 'a diva -> 'a -> 'a -> 'a
-  val moda : 'a diva -> 'a -> 'a -> 'a
-  type 'a zero
-  val zero : 'a zero -> 'a
-  type 'a no_zero_divisors
-  val times_no_zero_divisors : 'a no_zero_divisors -> 'a times
-  val zero_no_zero_divisors : 'a no_zero_divisors -> 'a zero
-  type 'a semigroup_mult
-  val times_semigroup_mult : 'a semigroup_mult -> 'a times
+  val minus_nat : nat -> nat -> nat
+  val zero_nat : nat
+  val minusinf : fm -> fm
+  val numsubst0 : numa -> numa -> numa
+  val subst0 : numa -> fm -> fm
   type 'a plus
   val plus : 'a plus -> 'a -> 'a -> 'a
   type 'a semigroup_add
   val plus_semigroup_add : 'a semigroup_add -> 'a plus
+  type 'a cancel_semigroup_add
+  val semigroup_add_cancel_semigroup_add :
+    'a cancel_semigroup_add -> 'a semigroup_add
   type 'a ab_semigroup_add
   val semigroup_add_ab_semigroup_add : 'a ab_semigroup_add -> 'a semigroup_add
-  type 'a semiring
-  val ab_semigroup_add_semiring : 'a semiring -> 'a ab_semigroup_add
-  val semigroup_mult_semiring : 'a semiring -> 'a semigroup_mult
-  type 'a mult_zero
-  val times_mult_zero : 'a mult_zero -> 'a times
-  val zero_mult_zero : 'a mult_zero -> 'a zero
+  type 'a cancel_ab_semigroup_add
+  val ab_semigroup_add_cancel_ab_semigroup_add :
+    'a cancel_ab_semigroup_add -> 'a ab_semigroup_add
+  val cancel_semigroup_add_cancel_ab_semigroup_add :
+    'a cancel_ab_semigroup_add -> 'a cancel_semigroup_add
+  type 'a zero
+  val zero : 'a zero -> 'a
   type 'a monoid_add
   val semigroup_add_monoid_add : 'a monoid_add -> 'a semigroup_add
   val zero_monoid_add : 'a monoid_add -> 'a zero
@@ -59,25 +62,29 @@
   val ab_semigroup_add_comm_monoid_add :
     'a comm_monoid_add -> 'a ab_semigroup_add
   val monoid_add_comm_monoid_add : 'a comm_monoid_add -> 'a monoid_add
+  type 'a cancel_comm_monoid_add
+  val cancel_ab_semigroup_add_cancel_comm_monoid_add :
+    'a cancel_comm_monoid_add -> 'a cancel_ab_semigroup_add
+  val comm_monoid_add_cancel_comm_monoid_add :
+    'a cancel_comm_monoid_add -> 'a comm_monoid_add
+  type 'a times
+  val times : 'a times -> 'a -> 'a -> 'a
+  type 'a mult_zero
+  val times_mult_zero : 'a mult_zero -> 'a times
+  val zero_mult_zero : 'a mult_zero -> 'a zero
+  type 'a semigroup_mult
+  val times_semigroup_mult : 'a semigroup_mult -> 'a times
+  type 'a semiring
+  val ab_semigroup_add_semiring : 'a semiring -> 'a ab_semigroup_add
+  val semigroup_mult_semiring : 'a semiring -> 'a semigroup_mult
   type 'a semiring_0
   val comm_monoid_add_semiring_0 : 'a semiring_0 -> 'a comm_monoid_add
   val mult_zero_semiring_0 : 'a semiring_0 -> 'a mult_zero
   val semiring_semiring_0 : 'a semiring_0 -> 'a semiring
-  type 'a one
-  val one : 'a one -> 'a
-  type 'a power
-  val one_power : 'a power -> 'a one
-  val times_power : 'a power -> 'a times
-  type 'a monoid_mult
-  val semigroup_mult_monoid_mult : 'a monoid_mult -> 'a semigroup_mult
-  val power_monoid_mult : 'a monoid_mult -> 'a power
-  type 'a zero_neq_one
-  val one_zero_neq_one : 'a zero_neq_one -> 'a one
-  val zero_zero_neq_one : 'a zero_neq_one -> 'a zero
-  type 'a semiring_1
-  val monoid_mult_semiring_1 : 'a semiring_1 -> 'a monoid_mult
-  val semiring_0_semiring_1 : 'a semiring_1 -> 'a semiring_0
-  val zero_neq_one_semiring_1 : 'a semiring_1 -> 'a zero_neq_one
+  type 'a semiring_0_cancel
+  val cancel_comm_monoid_add_semiring_0_cancel :
+    'a semiring_0_cancel -> 'a cancel_comm_monoid_add
+  val semiring_0_semiring_0_cancel : 'a semiring_0_cancel -> 'a semiring_0
   type 'a ab_semigroup_mult
   val semigroup_mult_ab_semigroup_mult :
     'a ab_semigroup_mult -> 'a semigroup_mult
@@ -87,42 +94,49 @@
   type 'a comm_semiring_0
   val comm_semiring_comm_semiring_0 : 'a comm_semiring_0 -> 'a comm_semiring
   val semiring_0_comm_semiring_0 : 'a comm_semiring_0 -> 'a semiring_0
+  type 'a comm_semiring_0_cancel
+  val comm_semiring_0_comm_semiring_0_cancel :
+    'a comm_semiring_0_cancel -> 'a comm_semiring_0
+  val semiring_0_cancel_comm_semiring_0_cancel :
+    'a comm_semiring_0_cancel -> 'a semiring_0_cancel
+  type 'a one
+  val one : 'a one -> 'a
+  type 'a power
+  val one_power : 'a power -> 'a one
+  val times_power : 'a power -> 'a times
+  type 'a monoid_mult
+  val semigroup_mult_monoid_mult : 'a monoid_mult -> 'a semigroup_mult
+  val power_monoid_mult : 'a monoid_mult -> 'a power
+  type 'a numeral
+  val one_numeral : 'a numeral -> 'a one
+  val semigroup_add_numeral : 'a numeral -> 'a semigroup_add
+  type 'a semiring_numeral
+  val monoid_mult_semiring_numeral : 'a semiring_numeral -> 'a monoid_mult
+  val numeral_semiring_numeral : 'a semiring_numeral -> 'a numeral
+  val semiring_semiring_numeral : 'a semiring_numeral -> 'a semiring
+  type 'a zero_neq_one
+  val one_zero_neq_one : 'a zero_neq_one -> 'a one
+  val zero_zero_neq_one : 'a zero_neq_one -> 'a zero
+  type 'a semiring_1
+  val semiring_numeral_semiring_1 : 'a semiring_1 -> 'a semiring_numeral
+  val semiring_0_semiring_1 : 'a semiring_1 -> 'a semiring_0
+  val zero_neq_one_semiring_1 : 'a semiring_1 -> 'a zero_neq_one
+  type 'a semiring_1_cancel
+  val semiring_0_cancel_semiring_1_cancel :
+    'a semiring_1_cancel -> 'a semiring_0_cancel
+  val semiring_1_semiring_1_cancel : 'a semiring_1_cancel -> 'a semiring_1
   type 'a comm_monoid_mult
   val ab_semigroup_mult_comm_monoid_mult :
     'a comm_monoid_mult -> 'a ab_semigroup_mult
   val monoid_mult_comm_monoid_mult : 'a comm_monoid_mult -> 'a monoid_mult
+  type 'a dvd
+  val times_dvd : 'a dvd -> 'a times
   type 'a comm_semiring_1
   val comm_monoid_mult_comm_semiring_1 :
     'a comm_semiring_1 -> 'a comm_monoid_mult
   val comm_semiring_0_comm_semiring_1 : 'a comm_semiring_1 -> 'a comm_semiring_0
   val dvd_comm_semiring_1 : 'a comm_semiring_1 -> 'a dvd
   val semiring_1_comm_semiring_1 : 'a comm_semiring_1 -> 'a semiring_1
-  type 'a cancel_semigroup_add
-  val semigroup_add_cancel_semigroup_add :
-    'a cancel_semigroup_add -> 'a semigroup_add
-  type 'a cancel_ab_semigroup_add
-  val ab_semigroup_add_cancel_ab_semigroup_add :
-    'a cancel_ab_semigroup_add -> 'a ab_semigroup_add
-  val cancel_semigroup_add_cancel_ab_semigroup_add :
-    'a cancel_ab_semigroup_add -> 'a cancel_semigroup_add
-  type 'a cancel_comm_monoid_add
-  val cancel_ab_semigroup_add_cancel_comm_monoid_add :
-    'a cancel_comm_monoid_add -> 'a cancel_ab_semigroup_add
-  val comm_monoid_add_cancel_comm_monoid_add :
-    'a cancel_comm_monoid_add -> 'a comm_monoid_add
-  type 'a semiring_0_cancel
-  val cancel_comm_monoid_add_semiring_0_cancel :
-    'a semiring_0_cancel -> 'a cancel_comm_monoid_add
-  val semiring_0_semiring_0_cancel : 'a semiring_0_cancel -> 'a semiring_0
-  type 'a semiring_1_cancel
-  val semiring_0_cancel_semiring_1_cancel :
-    'a semiring_1_cancel -> 'a semiring_0_cancel
-  val semiring_1_semiring_1_cancel : 'a semiring_1_cancel -> 'a semiring_1
-  type 'a comm_semiring_0_cancel
-  val comm_semiring_0_comm_semiring_0_cancel :
-    'a comm_semiring_0_cancel -> 'a comm_semiring_0
-  val semiring_0_cancel_comm_semiring_0_cancel :
-    'a comm_semiring_0_cancel -> 'a semiring_0_cancel
   type 'a comm_semiring_1_cancel
   val comm_semiring_0_cancel_comm_semiring_1_cancel :
     'a comm_semiring_1_cancel -> 'a comm_semiring_0_cancel
@@ -130,107 +144,182 @@
     'a comm_semiring_1_cancel -> 'a comm_semiring_1
   val semiring_1_cancel_comm_semiring_1_cancel :
     'a comm_semiring_1_cancel -> 'a semiring_1_cancel
+  type 'a no_zero_divisors
+  val times_no_zero_divisors : 'a no_zero_divisors -> 'a times
+  val zero_no_zero_divisors : 'a no_zero_divisors -> 'a zero
+  type 'a diva
+  val dvd_div : 'a diva -> 'a dvd
+  val diva : 'a diva -> 'a -> 'a -> 'a
+  val moda : 'a diva -> 'a -> 'a -> 'a
   type 'a semiring_div
   val div_semiring_div : 'a semiring_div -> 'a diva
   val comm_semiring_1_cancel_semiring_div :
     'a semiring_div -> 'a comm_semiring_1_cancel
   val no_zero_divisors_semiring_div : 'a semiring_div -> 'a no_zero_divisors
-  val dvd : 'a semiring_div * 'a equal -> 'a -> 'a -> bool
-  val abs_int : int -> int
-  val equal_int : int equal
-  val numadd : num * num -> num
-  val nummul : int -> num -> num
-  val numneg : num -> num
-  val numsub : num -> num -> num
-  val simpnum : num -> num
-  val one_inta : int
-  val zero_inta : int
-  val times_int : int times
-  val dvd_int : int dvd
-  val fst : 'a * 'b -> 'a
-  val sgn_int : int -> int
+  val plus_inta : inta -> inta -> inta
+  val plus_int : inta plus
+  val semigroup_add_int : inta semigroup_add
+  val cancel_semigroup_add_int : inta cancel_semigroup_add
+  val ab_semigroup_add_int : inta ab_semigroup_add
+  val cancel_ab_semigroup_add_int : inta cancel_ab_semigroup_add
+  val zero_inta : inta
+  val zero_int : inta zero
+  val monoid_add_int : inta monoid_add
+  val comm_monoid_add_int : inta comm_monoid_add
+  val cancel_comm_monoid_add_int : inta cancel_comm_monoid_add
+  val times_inta : inta -> inta -> inta
+  val times_int : inta times
+  val mult_zero_int : inta mult_zero
+  val semigroup_mult_int : inta semigroup_mult
+  val semiring_int : inta semiring
+  val semiring_0_int : inta semiring_0
+  val semiring_0_cancel_int : inta semiring_0_cancel
+  val ab_semigroup_mult_int : inta ab_semigroup_mult
+  val comm_semiring_int : inta comm_semiring
+  val comm_semiring_0_int : inta comm_semiring_0
+  val comm_semiring_0_cancel_int : inta comm_semiring_0_cancel
+  val one_inta : inta
+  val one_int : inta one
+  val power_int : inta power
+  val monoid_mult_int : inta monoid_mult
+  val numeral_int : inta numeral
+  val semiring_numeral_int : inta semiring_numeral
+  val zero_neq_one_int : inta zero_neq_one
+  val semiring_1_int : inta semiring_1
+  val semiring_1_cancel_int : inta semiring_1_cancel
+  val comm_monoid_mult_int : inta comm_monoid_mult
+  val dvd_int : inta dvd
+  val comm_semiring_1_int : inta comm_semiring_1
+  val comm_semiring_1_cancel_int : inta comm_semiring_1_cancel
+  val no_zero_divisors_int : inta no_zero_divisors
+  val sgn_integer : int -> int
+  val abs_integer : int -> int
   val apsnd : ('a -> 'b) -> 'c * 'a -> 'c * 'b
-  val divmod_int : int -> int -> int * int
-  val div_inta : int -> int -> int
+  val divmod_integer : int -> int -> int * int
   val snd : 'a * 'b -> 'b
-  val mod_int : int -> int -> int
-  val div_int : int diva
-  val zero_int : int zero
-  val no_zero_divisors_int : int no_zero_divisors
-  val semigroup_mult_int : int semigroup_mult
-  val plus_int : int plus
-  val semigroup_add_int : int semigroup_add
-  val ab_semigroup_add_int : int ab_semigroup_add
-  val semiring_int : int semiring
-  val mult_zero_int : int mult_zero
-  val monoid_add_int : int monoid_add
-  val comm_monoid_add_int : int comm_monoid_add
-  val semiring_0_int : int semiring_0
-  val one_int : int one
-  val power_int : int power
-  val monoid_mult_int : int monoid_mult
-  val zero_neq_one_int : int zero_neq_one
-  val semiring_1_int : int semiring_1
-  val ab_semigroup_mult_int : int ab_semigroup_mult
-  val comm_semiring_int : int comm_semiring
-  val comm_semiring_0_int : int comm_semiring_0
-  val comm_monoid_mult_int : int comm_monoid_mult
-  val comm_semiring_1_int : int comm_semiring_1
-  val cancel_semigroup_add_int : int cancel_semigroup_add
-  val cancel_ab_semigroup_add_int : int cancel_ab_semigroup_add
-  val cancel_comm_monoid_add_int : int cancel_comm_monoid_add
-  val semiring_0_cancel_int : int semiring_0_cancel
-  val semiring_1_cancel_int : int semiring_1_cancel
-  val comm_semiring_0_cancel_int : int comm_semiring_0_cancel
-  val comm_semiring_1_cancel_int : int comm_semiring_1_cancel
-  val semiring_div_int : int semiring_div
+  val mod_integer : int -> int -> int
+  val mod_int : inta -> inta -> inta
+  val fst : 'a * 'b -> 'a
+  val div_integer : int -> int -> int
+  val div_inta : inta -> inta -> inta
+  val div_int : inta diva
+  val semiring_div_int : inta semiring_div
+  val less_eq_int : inta -> inta -> bool
+  val uminus_int : inta -> inta
+  val nummul : inta -> numa -> numa
+  val numneg : numa -> numa
+  val less_eq_nat : nat -> nat -> bool
+  val numadd : numa * numa -> numa
+  val numsub : numa -> numa -> numa
+  val simpnum : numa -> numa
+  val less_int : inta -> inta -> bool
+  val equal_int : inta equal
+  val abs_int : inta -> inta
+  val nota : fm -> fm
+  val impa : fm -> fm -> fm
+  val iffa : fm -> fm -> fm
+  val disj : fm -> fm -> fm
+  val conj : fm -> fm -> fm
+  val dvd : 'a semiring_div * 'a equal -> 'a -> 'a -> bool
   val simpfm : fm -> fm
-  val qelim : fm -> (fm -> fm) -> fm
-  val maps : ('a -> 'b list) -> 'a list -> 'b list
-  val uptoa : int -> int -> int list
-  val minus_nat : int -> int -> int
-  val decrnum : num -> num
-  val decr : fm -> fm
-  val beta : fm -> num list
-  val gcd_int : int -> int -> int
-  val lcm_int : int -> int -> int
-  val zeta : fm -> int
-  val zsplit0 : num -> int * num
-  val zlfm : fm -> fm
-  val alpha : fm -> num list
-  val delta : fm -> int
+  val equal_num : numa equal
+  val gen_length : nat -> 'a list -> nat
+  val size_list : 'a list -> nat
+  val mirror : fm -> fm
+  val a_beta : fm -> inta -> fm
   val member : 'a equal -> 'a list -> 'a -> bool
   val remdups : 'a equal -> 'a list -> 'a list
-  val a_beta : fm -> int -> fm
-  val mirror : fm -> fm
-  val size_list : 'a list -> int
-  val equal_num : num equal
-  val unita : fm -> fm * (num list * int)
-  val numsubst0 : num -> num -> num
-  val subst0 : num -> fm -> fm
-  val minusinf : fm -> fm
+  val gcd_int : inta -> inta -> inta
+  val lcm_int : inta -> inta -> inta
+  val delta : fm -> inta
+  val alpha : fm -> numa list
+  val minus_int : inta -> inta -> inta
+  val zsplit0 : numa -> inta * numa
+  val zlfm : fm -> fm
+  val zeta : fm -> inta
+  val beta : fm -> numa list
+  val unita : fm -> fm * (numa list * inta)
+  val decrnum : numa -> numa
+  val decr : fm -> fm
+  val uptoa : inta -> inta -> inta list
+  val maps : ('a -> 'b list) -> 'a list -> 'b list
   val cooper : fm -> fm
+  val qelim : fm -> (fm -> fm) -> fm
+  val prep : fm -> fm
   val pa : fm -> fm
 end = struct
 
+fun id x = (fn xa => xa) x;
+
 type 'a equal = {equal : 'a -> 'a -> bool};
 val equal = #equal : 'a equal -> 'a -> 'a -> bool;
 
 fun eq A_ a b = equal A_ a b;
 
-fun suc n = n + (1 : IntInf.int);
+datatype inta = Int_of_integer of int;
+
+datatype nat = Nat of int;
+
+datatype num = One | Bit0 of num | Bit1 of num;
 
-datatype num = C of int | Bound of int | Cn of int * int * num | Neg of num |
-  Add of num * num | Sub of num * num | Mul of int * num;
+type 'a ord = {less_eq : 'a -> 'a -> bool, less : 'a -> 'a -> bool};
+val less_eq = #less_eq : 'a ord -> 'a -> 'a -> bool;
+val less = #less : 'a ord -> 'a -> 'a -> bool;
+
+val ord_integer =
+  {less_eq = (fn a => fn b => a <= b), less = (fn a => fn b => a < b)} :
+  int ord;
+
+fun max A_ a b = (if less_eq A_ a b then b else a);
 
-datatype fm = T | F | Lt of num | Le of num | Gt of num | Ge of num | Eq of num
-  | NEq of num | Dvd of int * num | NDvd of int * num | Not of fm |
-  And of fm * fm | Or of fm * fm | Imp of fm * fm | Iff of fm * fm | E of fm |
-  A of fm | Closed of int | NClosed of int;
+fun nat_of_integer k = Nat (max ord_integer 0 k);
+
+fun integer_of_nat (Nat x) = x;
+
+fun plus_nat m n = Nat (integer_of_nat m + integer_of_nat n);
+
+fun suc n = plus_nat n (nat_of_integer (1 : IntInf.int));
+
+datatype numa = C of inta | Bound of nat | Cn of nat * inta * numa | Neg of numa
+  | Add of numa * numa | Sub of numa * numa | Mul of inta * numa;
+
+datatype fm = T | F | Lt of numa | Le of numa | Gt of numa | Ge of numa |
+  Eq of numa | NEq of numa | Dvd of inta * numa | NDvd of inta * numa |
+  Not of fm | And of fm * fm | Or of fm * fm | Imp of fm * fm | Iff of fm * fm |
+  E of fm | A of fm | Closed of nat | NClosed of nat;
 
 fun map f [] = []
   | map f (x :: xs) = f x :: map f xs;
 
+fun disjuncts (Or (p, q)) = disjuncts p @ disjuncts q
+  | disjuncts F = []
+  | disjuncts T = [T]
+  | disjuncts (Lt v) = [Lt v]
+  | disjuncts (Le v) = [Le v]
+  | disjuncts (Gt v) = [Gt v]
+  | disjuncts (Ge v) = [Ge v]
+  | disjuncts (Eq v) = [Eq v]
+  | disjuncts (NEq v) = [NEq v]
+  | disjuncts (Dvd (v, va)) = [Dvd (v, va)]
+  | disjuncts (NDvd (v, va)) = [NDvd (v, va)]
+  | disjuncts (Not v) = [Not v]
+  | disjuncts (And (v, va)) = [And (v, va)]
+  | disjuncts (Imp (v, va)) = [Imp (v, va)]
+  | disjuncts (Iff (v, va)) = [Iff (v, va)]
+  | disjuncts (E v) = [E v]
+  | disjuncts (A v) = [A v]
+  | disjuncts (Closed v) = [Closed v]
+  | disjuncts (NClosed v) = [NClosed v];
+
+fun foldr f [] = id
+  | foldr f (x :: xs) = f x o foldr f xs;
+
+fun equal_nat m n = integer_of_nat m = integer_of_nat n;
+
+fun integer_of_int (Int_of_integer k) = k;
+
+fun equal_inta k l = integer_of_int k = integer_of_int l;
+
 fun equal_numa (Mul (inta, num)) (Sub (num1, num2)) = false
   | equal_numa (Sub (num1, num2)) (Mul (inta, num)) = false
   | equal_numa (Mul (inta, num)) (Add (num1, num2)) = false
@@ -274,16 +363,17 @@
   | equal_numa (Bound nat) (C inta) = false
   | equal_numa (C inta) (Bound nat) = false
   | equal_numa (Mul (intaa, numa)) (Mul (inta, num)) =
-    intaa = inta andalso equal_numa numa num
+    equal_inta intaa inta andalso equal_numa numa num
   | equal_numa (Sub (num1a, num2a)) (Sub (num1, num2)) =
     equal_numa num1a num1 andalso equal_numa num2a num2
   | equal_numa (Add (num1a, num2a)) (Add (num1, num2)) =
     equal_numa num1a num1 andalso equal_numa num2a num2
   | equal_numa (Neg numa) (Neg num) = equal_numa numa num
   | equal_numa (Cn (nata, intaa, numa)) (Cn (nat, inta, num)) =
-    nata = nat andalso (intaa = inta andalso equal_numa numa num)
-  | equal_numa (Bound nata) (Bound nat) = nata = nat
-  | equal_numa (C intaa) (C inta) = intaa = inta;
+    equal_nat nata nat andalso
+      (equal_inta intaa inta andalso equal_numa numa num)
+  | equal_numa (Bound nata) (Bound nat) = equal_nat nata nat
+  | equal_numa (C intaa) (C inta) = equal_inta intaa inta;
 
 fun equal_fm (NClosed nata) (Closed nat) = false
   | equal_fm (Closed nata) (NClosed nat) = false
@@ -627,8 +717,8 @@
   | equal_fm T (Lt num) = false
   | equal_fm F T = false
   | equal_fm T F = false
-  | equal_fm (NClosed nata) (NClosed nat) = nata = nat
-  | equal_fm (Closed nata) (Closed nat) = nata = nat
+  | equal_fm (NClosed nata) (NClosed nat) = equal_nat nata nat
+  | equal_fm (Closed nata) (Closed nat) = equal_nat nata nat
   | equal_fm (A fma) (A fm) = equal_fm fma fm
   | equal_fm (E fma) (E fm) = equal_fm fma fm
   | equal_fm (Iff (fm1a, fm2a)) (Iff (fm1, fm2)) =
@@ -641,9 +731,9 @@
     equal_fm fm1a fm1 andalso equal_fm fm2a fm2
   | equal_fm (Not fma) (Not fm) = equal_fm fma fm
   | equal_fm (NDvd (intaa, numa)) (NDvd (inta, num)) =
-    intaa = inta andalso equal_numa numa num
+    equal_inta intaa inta andalso equal_numa numa num
   | equal_fm (Dvd (intaa, numa)) (Dvd (inta, num)) =
-    intaa = inta andalso equal_numa numa num
+    equal_inta intaa inta andalso equal_numa numa num
   | equal_fm (NEq numa) (NEq num) = equal_numa numa num
   | equal_fm (Eq numa) (Eq num) = equal_numa numa num
   | equal_fm (Ge numa) (Ge num) = equal_numa numa num
@@ -666,32 +756,1457 @@
                   | E _ => Or (f p, q) | A _ => Or (f p, q)
                   | Closed _ => Or (f p, q) | NClosed _ => Or (f p, q))));
 
-fun foldr f [] a = a
-  | foldr f (x :: xs) a = f x (foldr f xs a);
-
 fun evaldjf f ps = foldr (djf f) ps F;
 
-fun disjuncts (Or (p, q)) = disjuncts p @ disjuncts q
-  | disjuncts F = []
-  | disjuncts T = [T]
-  | disjuncts (Lt v) = [Lt v]
-  | disjuncts (Le v) = [Le v]
-  | disjuncts (Gt v) = [Gt v]
-  | disjuncts (Ge v) = [Ge v]
-  | disjuncts (Eq v) = [Eq v]
-  | disjuncts (NEq v) = [NEq v]
-  | disjuncts (Dvd (v, va)) = [Dvd (v, va)]
-  | disjuncts (NDvd (v, va)) = [NDvd (v, va)]
-  | disjuncts (Not v) = [Not v]
-  | disjuncts (And (v, va)) = [And (v, va)]
-  | disjuncts (Imp (v, va)) = [Imp (v, va)]
-  | disjuncts (Iff (v, va)) = [Iff (v, va)]
-  | disjuncts (E v) = [E v]
-  | disjuncts (A v) = [A v]
-  | disjuncts (Closed v) = [Closed v]
-  | disjuncts (NClosed v) = [NClosed v];
+fun dj f p = evaldjf f (disjuncts p);
+
+fun minus_nat m n =
+  Nat (max ord_integer 0 (integer_of_nat m - integer_of_nat n));
+
+val zero_nat : nat = Nat 0;
+
+fun minusinf (And (p, q)) = And (minusinf p, minusinf q)
+  | minusinf (Or (p, q)) = Or (minusinf p, minusinf q)
+  | minusinf T = T
+  | minusinf F = F
+  | minusinf (Lt (C bo)) = Lt (C bo)
+  | minusinf (Lt (Bound bp)) = Lt (Bound bp)
+  | minusinf (Lt (Neg bt)) = Lt (Neg bt)
+  | minusinf (Lt (Add (bu, bv))) = Lt (Add (bu, bv))
+  | minusinf (Lt (Sub (bw, bx))) = Lt (Sub (bw, bx))
+  | minusinf (Lt (Mul (by, bz))) = Lt (Mul (by, bz))
+  | minusinf (Le (C co)) = Le (C co)
+  | minusinf (Le (Bound cp)) = Le (Bound cp)
+  | minusinf (Le (Neg ct)) = Le (Neg ct)
+  | minusinf (Le (Add (cu, cv))) = Le (Add (cu, cv))
+  | minusinf (Le (Sub (cw, cx))) = Le (Sub (cw, cx))
+  | minusinf (Le (Mul (cy, cz))) = Le (Mul (cy, cz))
+  | minusinf (Gt (C doa)) = Gt (C doa)
+  | minusinf (Gt (Bound dp)) = Gt (Bound dp)
+  | minusinf (Gt (Neg dt)) = Gt (Neg dt)
+  | minusinf (Gt (Add (du, dv))) = Gt (Add (du, dv))
+  | minusinf (Gt (Sub (dw, dx))) = Gt (Sub (dw, dx))
+  | minusinf (Gt (Mul (dy, dz))) = Gt (Mul (dy, dz))
+  | minusinf (Ge (C eo)) = Ge (C eo)
+  | minusinf (Ge (Bound ep)) = Ge (Bound ep)
+  | minusinf (Ge (Neg et)) = Ge (Neg et)
+  | minusinf (Ge (Add (eu, ev))) = Ge (Add (eu, ev))
+  | minusinf (Ge (Sub (ew, ex))) = Ge (Sub (ew, ex))
+  | minusinf (Ge (Mul (ey, ez))) = Ge (Mul (ey, ez))
+  | minusinf (Eq (C fo)) = Eq (C fo)
+  | minusinf (Eq (Bound fp)) = Eq (Bound fp)
+  | minusinf (Eq (Neg ft)) = Eq (Neg ft)
+  | minusinf (Eq (Add (fu, fv))) = Eq (Add (fu, fv))
+  | minusinf (Eq (Sub (fw, fx))) = Eq (Sub (fw, fx))
+  | minusinf (Eq (Mul (fy, fz))) = Eq (Mul (fy, fz))
+  | minusinf (NEq (C go)) = NEq (C go)
+  | minusinf (NEq (Bound gp)) = NEq (Bound gp)
+  | minusinf (NEq (Neg gt)) = NEq (Neg gt)
+  | minusinf (NEq (Add (gu, gv))) = NEq (Add (gu, gv))
+  | minusinf (NEq (Sub (gw, gx))) = NEq (Sub (gw, gx))
+  | minusinf (NEq (Mul (gy, gz))) = NEq (Mul (gy, gz))
+  | minusinf (Dvd (aa, ab)) = Dvd (aa, ab)
+  | minusinf (NDvd (ac, ad)) = NDvd (ac, ad)
+  | minusinf (Not ae) = Not ae
+  | minusinf (Imp (aj, ak)) = Imp (aj, ak)
+  | minusinf (Iff (al, am)) = Iff (al, am)
+  | minusinf (E an) = E an
+  | minusinf (A ao) = A ao
+  | minusinf (Closed ap) = Closed ap
+  | minusinf (NClosed aq) = NClosed aq
+  | minusinf (Lt (Cn (cm, c, e))) =
+    (if equal_nat cm zero_nat then T
+      else Lt (Cn (suc (minus_nat cm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | minusinf (Le (Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat then T
+      else Le (Cn (suc (minus_nat dm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | minusinf (Gt (Cn (em, c, e))) =
+    (if equal_nat em zero_nat then F
+      else Gt (Cn (suc (minus_nat em (nat_of_integer (1 : IntInf.int))), c, e)))
+  | minusinf (Ge (Cn (fm, c, e))) =
+    (if equal_nat fm zero_nat then F
+      else Ge (Cn (suc (minus_nat fm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | minusinf (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat then F
+      else Eq (Cn (suc (minus_nat gm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | minusinf (NEq (Cn (hm, c, e))) =
+    (if equal_nat hm zero_nat then T
+      else NEq (Cn (suc (minus_nat hm (nat_of_integer (1 : IntInf.int))), c,
+                     e)));
+
+fun numsubst0 t (C c) = C c
+  | numsubst0 t (Bound n) = (if equal_nat n zero_nat then t else Bound n)
+  | numsubst0 t (Neg a) = Neg (numsubst0 t a)
+  | numsubst0 t (Add (a, b)) = Add (numsubst0 t a, numsubst0 t b)
+  | numsubst0 t (Sub (a, b)) = Sub (numsubst0 t a, numsubst0 t b)
+  | numsubst0 t (Mul (i, a)) = Mul (i, numsubst0 t a)
+  | numsubst0 t (Cn (v, i, a)) =
+    (if equal_nat v zero_nat then Add (Mul (i, t), numsubst0 t a)
+      else Cn (suc (minus_nat v (nat_of_integer (1 : IntInf.int))), i,
+                numsubst0 t a));
+
+fun subst0 t T = T
+  | subst0 t F = F
+  | subst0 t (Lt a) = Lt (numsubst0 t a)
+  | subst0 t (Le a) = Le (numsubst0 t a)
+  | subst0 t (Gt a) = Gt (numsubst0 t a)
+  | subst0 t (Ge a) = Ge (numsubst0 t a)
+  | subst0 t (Eq a) = Eq (numsubst0 t a)
+  | subst0 t (NEq a) = NEq (numsubst0 t a)
+  | subst0 t (Dvd (i, a)) = Dvd (i, numsubst0 t a)
+  | subst0 t (NDvd (i, a)) = NDvd (i, numsubst0 t a)
+  | subst0 t (Not p) = Not (subst0 t p)
+  | subst0 t (And (p, q)) = And (subst0 t p, subst0 t q)
+  | subst0 t (Or (p, q)) = Or (subst0 t p, subst0 t q)
+  | subst0 t (Imp (p, q)) = Imp (subst0 t p, subst0 t q)
+  | subst0 t (Iff (p, q)) = Iff (subst0 t p, subst0 t q)
+  | subst0 t (Closed p) = Closed p
+  | subst0 t (NClosed p) = NClosed p;
+
+type 'a plus = {plus : 'a -> 'a -> 'a};
+val plus = #plus : 'a plus -> 'a -> 'a -> 'a;
+
+type 'a semigroup_add = {plus_semigroup_add : 'a plus};
+val plus_semigroup_add = #plus_semigroup_add : 'a semigroup_add -> 'a plus;
+
+type 'a cancel_semigroup_add =
+  {semigroup_add_cancel_semigroup_add : 'a semigroup_add};
+val semigroup_add_cancel_semigroup_add = #semigroup_add_cancel_semigroup_add :
+  'a cancel_semigroup_add -> 'a semigroup_add;
+
+type 'a ab_semigroup_add = {semigroup_add_ab_semigroup_add : 'a semigroup_add};
+val semigroup_add_ab_semigroup_add = #semigroup_add_ab_semigroup_add :
+  'a ab_semigroup_add -> 'a semigroup_add;
+
+type 'a cancel_ab_semigroup_add =
+  {ab_semigroup_add_cancel_ab_semigroup_add : 'a ab_semigroup_add,
+    cancel_semigroup_add_cancel_ab_semigroup_add : 'a cancel_semigroup_add};
+val ab_semigroup_add_cancel_ab_semigroup_add =
+  #ab_semigroup_add_cancel_ab_semigroup_add :
+  'a cancel_ab_semigroup_add -> 'a ab_semigroup_add;
+val cancel_semigroup_add_cancel_ab_semigroup_add =
+  #cancel_semigroup_add_cancel_ab_semigroup_add :
+  'a cancel_ab_semigroup_add -> 'a cancel_semigroup_add;
+
+type 'a zero = {zero : 'a};
+val zero = #zero : 'a zero -> 'a;
+
+type 'a monoid_add =
+  {semigroup_add_monoid_add : 'a semigroup_add, zero_monoid_add : 'a zero};
+val semigroup_add_monoid_add = #semigroup_add_monoid_add :
+  'a monoid_add -> 'a semigroup_add;
+val zero_monoid_add = #zero_monoid_add : 'a monoid_add -> 'a zero;
+
+type 'a comm_monoid_add =
+  {ab_semigroup_add_comm_monoid_add : 'a ab_semigroup_add,
+    monoid_add_comm_monoid_add : 'a monoid_add};
+val ab_semigroup_add_comm_monoid_add = #ab_semigroup_add_comm_monoid_add :
+  'a comm_monoid_add -> 'a ab_semigroup_add;
+val monoid_add_comm_monoid_add = #monoid_add_comm_monoid_add :
+  'a comm_monoid_add -> 'a monoid_add;
+
+type 'a cancel_comm_monoid_add =
+  {cancel_ab_semigroup_add_cancel_comm_monoid_add : 'a cancel_ab_semigroup_add,
+    comm_monoid_add_cancel_comm_monoid_add : 'a comm_monoid_add};
+val cancel_ab_semigroup_add_cancel_comm_monoid_add =
+  #cancel_ab_semigroup_add_cancel_comm_monoid_add :
+  'a cancel_comm_monoid_add -> 'a cancel_ab_semigroup_add;
+val comm_monoid_add_cancel_comm_monoid_add =
+  #comm_monoid_add_cancel_comm_monoid_add :
+  'a cancel_comm_monoid_add -> 'a comm_monoid_add;
+
+type 'a times = {times : 'a -> 'a -> 'a};
+val times = #times : 'a times -> 'a -> 'a -> 'a;
+
+type 'a mult_zero = {times_mult_zero : 'a times, zero_mult_zero : 'a zero};
+val times_mult_zero = #times_mult_zero : 'a mult_zero -> 'a times;
+val zero_mult_zero = #zero_mult_zero : 'a mult_zero -> 'a zero;
+
+type 'a semigroup_mult = {times_semigroup_mult : 'a times};
+val times_semigroup_mult = #times_semigroup_mult :
+  'a semigroup_mult -> 'a times;
+
+type 'a semiring =
+  {ab_semigroup_add_semiring : 'a ab_semigroup_add,
+    semigroup_mult_semiring : 'a semigroup_mult};
+val ab_semigroup_add_semiring = #ab_semigroup_add_semiring :
+  'a semiring -> 'a ab_semigroup_add;
+val semigroup_mult_semiring = #semigroup_mult_semiring :
+  'a semiring -> 'a semigroup_mult;
+
+type 'a semiring_0 =
+  {comm_monoid_add_semiring_0 : 'a comm_monoid_add,
+    mult_zero_semiring_0 : 'a mult_zero, semiring_semiring_0 : 'a semiring};
+val comm_monoid_add_semiring_0 = #comm_monoid_add_semiring_0 :
+  'a semiring_0 -> 'a comm_monoid_add;
+val mult_zero_semiring_0 = #mult_zero_semiring_0 :
+  'a semiring_0 -> 'a mult_zero;
+val semiring_semiring_0 = #semiring_semiring_0 : 'a semiring_0 -> 'a semiring;
+
+type 'a semiring_0_cancel =
+  {cancel_comm_monoid_add_semiring_0_cancel : 'a cancel_comm_monoid_add,
+    semiring_0_semiring_0_cancel : 'a semiring_0};
+val cancel_comm_monoid_add_semiring_0_cancel =
+  #cancel_comm_monoid_add_semiring_0_cancel :
+  'a semiring_0_cancel -> 'a cancel_comm_monoid_add;
+val semiring_0_semiring_0_cancel = #semiring_0_semiring_0_cancel :
+  'a semiring_0_cancel -> 'a semiring_0;
+
+type 'a ab_semigroup_mult =
+  {semigroup_mult_ab_semigroup_mult : 'a semigroup_mult};
+val semigroup_mult_ab_semigroup_mult = #semigroup_mult_ab_semigroup_mult :
+  'a ab_semigroup_mult -> 'a semigroup_mult;
+
+type 'a comm_semiring =
+  {ab_semigroup_mult_comm_semiring : 'a ab_semigroup_mult,
+    semiring_comm_semiring : 'a semiring};
+val ab_semigroup_mult_comm_semiring = #ab_semigroup_mult_comm_semiring :
+  'a comm_semiring -> 'a ab_semigroup_mult;
+val semiring_comm_semiring = #semiring_comm_semiring :
+  'a comm_semiring -> 'a semiring;
+
+type 'a comm_semiring_0 =
+  {comm_semiring_comm_semiring_0 : 'a comm_semiring,
+    semiring_0_comm_semiring_0 : 'a semiring_0};
+val comm_semiring_comm_semiring_0 = #comm_semiring_comm_semiring_0 :
+  'a comm_semiring_0 -> 'a comm_semiring;
+val semiring_0_comm_semiring_0 = #semiring_0_comm_semiring_0 :
+  'a comm_semiring_0 -> 'a semiring_0;
+
+type 'a comm_semiring_0_cancel =
+  {comm_semiring_0_comm_semiring_0_cancel : 'a comm_semiring_0,
+    semiring_0_cancel_comm_semiring_0_cancel : 'a semiring_0_cancel};
+val comm_semiring_0_comm_semiring_0_cancel =
+  #comm_semiring_0_comm_semiring_0_cancel :
+  'a comm_semiring_0_cancel -> 'a comm_semiring_0;
+val semiring_0_cancel_comm_semiring_0_cancel =
+  #semiring_0_cancel_comm_semiring_0_cancel :
+  'a comm_semiring_0_cancel -> 'a semiring_0_cancel;
+
+type 'a one = {one : 'a};
+val one = #one : 'a one -> 'a;
+
+type 'a power = {one_power : 'a one, times_power : 'a times};
+val one_power = #one_power : 'a power -> 'a one;
+val times_power = #times_power : 'a power -> 'a times;
+
+type 'a monoid_mult =
+  {semigroup_mult_monoid_mult : 'a semigroup_mult,
+    power_monoid_mult : 'a power};
+val semigroup_mult_monoid_mult = #semigroup_mult_monoid_mult :
+  'a monoid_mult -> 'a semigroup_mult;
+val power_monoid_mult = #power_monoid_mult : 'a monoid_mult -> 'a power;
+
+type 'a numeral =
+  {one_numeral : 'a one, semigroup_add_numeral : 'a semigroup_add};
+val one_numeral = #one_numeral : 'a numeral -> 'a one;
+val semigroup_add_numeral = #semigroup_add_numeral :
+  'a numeral -> 'a semigroup_add;
+
+type 'a semiring_numeral =
+  {monoid_mult_semiring_numeral : 'a monoid_mult,
+    numeral_semiring_numeral : 'a numeral,
+    semiring_semiring_numeral : 'a semiring};
+val monoid_mult_semiring_numeral = #monoid_mult_semiring_numeral :
+  'a semiring_numeral -> 'a monoid_mult;
+val numeral_semiring_numeral = #numeral_semiring_numeral :
+  'a semiring_numeral -> 'a numeral;
+val semiring_semiring_numeral = #semiring_semiring_numeral :
+  'a semiring_numeral -> 'a semiring;
+
+type 'a zero_neq_one = {one_zero_neq_one : 'a one, zero_zero_neq_one : 'a zero};
+val one_zero_neq_one = #one_zero_neq_one : 'a zero_neq_one -> 'a one;
+val zero_zero_neq_one = #zero_zero_neq_one : 'a zero_neq_one -> 'a zero;
+
+type 'a semiring_1 =
+  {semiring_numeral_semiring_1 : 'a semiring_numeral,
+    semiring_0_semiring_1 : 'a semiring_0,
+    zero_neq_one_semiring_1 : 'a zero_neq_one};
+val semiring_numeral_semiring_1 = #semiring_numeral_semiring_1 :
+  'a semiring_1 -> 'a semiring_numeral;
+val semiring_0_semiring_1 = #semiring_0_semiring_1 :
+  'a semiring_1 -> 'a semiring_0;
+val zero_neq_one_semiring_1 = #zero_neq_one_semiring_1 :
+  'a semiring_1 -> 'a zero_neq_one;
+
+type 'a semiring_1_cancel =
+  {semiring_0_cancel_semiring_1_cancel : 'a semiring_0_cancel,
+    semiring_1_semiring_1_cancel : 'a semiring_1};
+val semiring_0_cancel_semiring_1_cancel = #semiring_0_cancel_semiring_1_cancel :
+  'a semiring_1_cancel -> 'a semiring_0_cancel;
+val semiring_1_semiring_1_cancel = #semiring_1_semiring_1_cancel :
+  'a semiring_1_cancel -> 'a semiring_1;
+
+type 'a comm_monoid_mult =
+  {ab_semigroup_mult_comm_monoid_mult : 'a ab_semigroup_mult,
+    monoid_mult_comm_monoid_mult : 'a monoid_mult};
+val ab_semigroup_mult_comm_monoid_mult = #ab_semigroup_mult_comm_monoid_mult :
+  'a comm_monoid_mult -> 'a ab_semigroup_mult;
+val monoid_mult_comm_monoid_mult = #monoid_mult_comm_monoid_mult :
+  'a comm_monoid_mult -> 'a monoid_mult;
+
+type 'a dvd = {times_dvd : 'a times};
+val times_dvd = #times_dvd : 'a dvd -> 'a times;
+
+type 'a comm_semiring_1 =
+  {comm_monoid_mult_comm_semiring_1 : 'a comm_monoid_mult,
+    comm_semiring_0_comm_semiring_1 : 'a comm_semiring_0,
+    dvd_comm_semiring_1 : 'a dvd, semiring_1_comm_semiring_1 : 'a semiring_1};
+val comm_monoid_mult_comm_semiring_1 = #comm_monoid_mult_comm_semiring_1 :
+  'a comm_semiring_1 -> 'a comm_monoid_mult;
+val comm_semiring_0_comm_semiring_1 = #comm_semiring_0_comm_semiring_1 :
+  'a comm_semiring_1 -> 'a comm_semiring_0;
+val dvd_comm_semiring_1 = #dvd_comm_semiring_1 : 'a comm_semiring_1 -> 'a dvd;
+val semiring_1_comm_semiring_1 = #semiring_1_comm_semiring_1 :
+  'a comm_semiring_1 -> 'a semiring_1;
+
+type 'a comm_semiring_1_cancel =
+  {comm_semiring_0_cancel_comm_semiring_1_cancel : 'a comm_semiring_0_cancel,
+    comm_semiring_1_comm_semiring_1_cancel : 'a comm_semiring_1,
+    semiring_1_cancel_comm_semiring_1_cancel : 'a semiring_1_cancel};
+val comm_semiring_0_cancel_comm_semiring_1_cancel =
+  #comm_semiring_0_cancel_comm_semiring_1_cancel :
+  'a comm_semiring_1_cancel -> 'a comm_semiring_0_cancel;
+val comm_semiring_1_comm_semiring_1_cancel =
+  #comm_semiring_1_comm_semiring_1_cancel :
+  'a comm_semiring_1_cancel -> 'a comm_semiring_1;
+val semiring_1_cancel_comm_semiring_1_cancel =
+  #semiring_1_cancel_comm_semiring_1_cancel :
+  'a comm_semiring_1_cancel -> 'a semiring_1_cancel;
+
+type 'a no_zero_divisors =
+  {times_no_zero_divisors : 'a times, zero_no_zero_divisors : 'a zero};
+val times_no_zero_divisors = #times_no_zero_divisors :
+  'a no_zero_divisors -> 'a times;
+val zero_no_zero_divisors = #zero_no_zero_divisors :
+  'a no_zero_divisors -> 'a zero;
+
+type 'a diva = {dvd_div : 'a dvd, diva : 'a -> 'a -> 'a, moda : 'a -> 'a -> 'a};
+val dvd_div = #dvd_div : 'a diva -> 'a dvd;
+val diva = #diva : 'a diva -> 'a -> 'a -> 'a;
+val moda = #moda : 'a diva -> 'a -> 'a -> 'a;
+
+type 'a semiring_div =
+  {div_semiring_div : 'a diva,
+    comm_semiring_1_cancel_semiring_div : 'a comm_semiring_1_cancel,
+    no_zero_divisors_semiring_div : 'a no_zero_divisors};
+val div_semiring_div = #div_semiring_div : 'a semiring_div -> 'a diva;
+val comm_semiring_1_cancel_semiring_div = #comm_semiring_1_cancel_semiring_div :
+  'a semiring_div -> 'a comm_semiring_1_cancel;
+val no_zero_divisors_semiring_div = #no_zero_divisors_semiring_div :
+  'a semiring_div -> 'a no_zero_divisors;
+
+fun plus_inta k l = Int_of_integer (integer_of_int k + integer_of_int l);
+
+val plus_int = {plus = plus_inta} : inta plus;
+
+val semigroup_add_int = {plus_semigroup_add = plus_int} : inta semigroup_add;
+
+val cancel_semigroup_add_int =
+  {semigroup_add_cancel_semigroup_add = semigroup_add_int} :
+  inta cancel_semigroup_add;
+
+val ab_semigroup_add_int = {semigroup_add_ab_semigroup_add = semigroup_add_int}
+  : inta ab_semigroup_add;
+
+val cancel_ab_semigroup_add_int =
+  {ab_semigroup_add_cancel_ab_semigroup_add = ab_semigroup_add_int,
+    cancel_semigroup_add_cancel_ab_semigroup_add = cancel_semigroup_add_int}
+  : inta cancel_ab_semigroup_add;
+
+val zero_inta : inta = Int_of_integer 0;
+
+val zero_int = {zero = zero_inta} : inta zero;
+
+val monoid_add_int =
+  {semigroup_add_monoid_add = semigroup_add_int, zero_monoid_add = zero_int} :
+  inta monoid_add;
+
+val comm_monoid_add_int =
+  {ab_semigroup_add_comm_monoid_add = ab_semigroup_add_int,
+    monoid_add_comm_monoid_add = monoid_add_int}
+  : inta comm_monoid_add;
+
+val cancel_comm_monoid_add_int =
+  {cancel_ab_semigroup_add_cancel_comm_monoid_add = cancel_ab_semigroup_add_int,
+    comm_monoid_add_cancel_comm_monoid_add = comm_monoid_add_int}
+  : inta cancel_comm_monoid_add;
+
+fun times_inta k l = Int_of_integer (integer_of_int k * integer_of_int l);
+
+val times_int = {times = times_inta} : inta times;
+
+val mult_zero_int = {times_mult_zero = times_int, zero_mult_zero = zero_int} :
+  inta mult_zero;
+
+val semigroup_mult_int = {times_semigroup_mult = times_int} :
+  inta semigroup_mult;
+
+val semiring_int =
+  {ab_semigroup_add_semiring = ab_semigroup_add_int,
+    semigroup_mult_semiring = semigroup_mult_int}
+  : inta semiring;
+
+val semiring_0_int =
+  {comm_monoid_add_semiring_0 = comm_monoid_add_int,
+    mult_zero_semiring_0 = mult_zero_int, semiring_semiring_0 = semiring_int}
+  : inta semiring_0;
+
+val semiring_0_cancel_int =
+  {cancel_comm_monoid_add_semiring_0_cancel = cancel_comm_monoid_add_int,
+    semiring_0_semiring_0_cancel = semiring_0_int}
+  : inta semiring_0_cancel;
+
+val ab_semigroup_mult_int =
+  {semigroup_mult_ab_semigroup_mult = semigroup_mult_int} :
+  inta ab_semigroup_mult;
+
+val comm_semiring_int =
+  {ab_semigroup_mult_comm_semiring = ab_semigroup_mult_int,
+    semiring_comm_semiring = semiring_int}
+  : inta comm_semiring;
+
+val comm_semiring_0_int =
+  {comm_semiring_comm_semiring_0 = comm_semiring_int,
+    semiring_0_comm_semiring_0 = semiring_0_int}
+  : inta comm_semiring_0;
+
+val comm_semiring_0_cancel_int =
+  {comm_semiring_0_comm_semiring_0_cancel = comm_semiring_0_int,
+    semiring_0_cancel_comm_semiring_0_cancel = semiring_0_cancel_int}
+  : inta comm_semiring_0_cancel;
+
+val one_inta : inta = Int_of_integer (1 : IntInf.int);
+
+val one_int = {one = one_inta} : inta one;
+
+val power_int = {one_power = one_int, times_power = times_int} : inta power;
+
+val monoid_mult_int =
+  {semigroup_mult_monoid_mult = semigroup_mult_int,
+    power_monoid_mult = power_int}
+  : inta monoid_mult;
+
+val numeral_int =
+  {one_numeral = one_int, semigroup_add_numeral = semigroup_add_int} :
+  inta numeral;
+
+val semiring_numeral_int =
+  {monoid_mult_semiring_numeral = monoid_mult_int,
+    numeral_semiring_numeral = numeral_int,
+    semiring_semiring_numeral = semiring_int}
+  : inta semiring_numeral;
+
+val zero_neq_one_int =
+  {one_zero_neq_one = one_int, zero_zero_neq_one = zero_int} :
+  inta zero_neq_one;
+
+val semiring_1_int =
+  {semiring_numeral_semiring_1 = semiring_numeral_int,
+    semiring_0_semiring_1 = semiring_0_int,
+    zero_neq_one_semiring_1 = zero_neq_one_int}
+  : inta semiring_1;
+
+val semiring_1_cancel_int =
+  {semiring_0_cancel_semiring_1_cancel = semiring_0_cancel_int,
+    semiring_1_semiring_1_cancel = semiring_1_int}
+  : inta semiring_1_cancel;
+
+val comm_monoid_mult_int =
+  {ab_semigroup_mult_comm_monoid_mult = ab_semigroup_mult_int,
+    monoid_mult_comm_monoid_mult = monoid_mult_int}
+  : inta comm_monoid_mult;
+
+val dvd_int = {times_dvd = times_int} : inta dvd;
+
+val comm_semiring_1_int =
+  {comm_monoid_mult_comm_semiring_1 = comm_monoid_mult_int,
+    comm_semiring_0_comm_semiring_1 = comm_semiring_0_int,
+    dvd_comm_semiring_1 = dvd_int, semiring_1_comm_semiring_1 = semiring_1_int}
+  : inta comm_semiring_1;
+
+val comm_semiring_1_cancel_int =
+  {comm_semiring_0_cancel_comm_semiring_1_cancel = comm_semiring_0_cancel_int,
+    comm_semiring_1_comm_semiring_1_cancel = comm_semiring_1_int,
+    semiring_1_cancel_comm_semiring_1_cancel = semiring_1_cancel_int}
+  : inta comm_semiring_1_cancel;
+
+val no_zero_divisors_int =
+  {times_no_zero_divisors = times_int, zero_no_zero_divisors = zero_int} :
+  inta no_zero_divisors;
+
+fun sgn_integer k =
+  (if k = 0 then 0
+    else (if k < 0 then ~ (1 : IntInf.int) else (1 : IntInf.int)));
+
+fun abs_integer k = (if k < 0 then ~ k else k);
+
+fun apsnd f (x, y) = (x, f y);
+
+fun divmod_integer k l =
+  (if k = 0 then (0, 0)
+    else (if l = 0 then (0, k)
+           else (apsnd o (fn a => fn b => a * b) o sgn_integer) l
+                  (if sgn_integer k = sgn_integer l
+                    then Integer.div_mod (abs k) (abs l)
+                    else let
+                           val (r, s) = Integer.div_mod (abs k) (abs l);
+                         in
+                           (if s = 0 then (~ r, 0)
+                             else (~ r - (1 : IntInf.int), abs_integer l - s))
+                         end)));
+
+fun snd (a, b) = b;
+
+fun mod_integer k l = snd (divmod_integer k l);
+
+fun mod_int k l =
+  Int_of_integer (mod_integer (integer_of_int k) (integer_of_int l));
+
+fun fst (a, b) = a;
+
+fun div_integer k l = fst (divmod_integer k l);
+
+fun div_inta k l =
+  Int_of_integer (div_integer (integer_of_int k) (integer_of_int l));
+
+val div_int = {dvd_div = dvd_int, diva = div_inta, moda = mod_int} : inta diva;
+
+val semiring_div_int =
+  {div_semiring_div = div_int,
+    comm_semiring_1_cancel_semiring_div = comm_semiring_1_cancel_int,
+    no_zero_divisors_semiring_div = no_zero_divisors_int}
+  : inta semiring_div;
+
+fun less_eq_int k l = integer_of_int k <= integer_of_int l;
+
+fun uminus_int k = Int_of_integer (~ (integer_of_int k));
+
+fun nummul i (C j) = C (times_inta i j)
+  | nummul i (Cn (n, c, t)) = Cn (n, times_inta c i, nummul i t)
+  | nummul i (Bound v) = Mul (i, Bound v)
+  | nummul i (Neg v) = Mul (i, Neg v)
+  | nummul i (Add (v, va)) = Mul (i, Add (v, va))
+  | nummul i (Sub (v, va)) = Mul (i, Sub (v, va))
+  | nummul i (Mul (v, va)) = Mul (i, Mul (v, va));
+
+fun numneg t = nummul (uminus_int (Int_of_integer (1 : IntInf.int))) t;
+
+fun less_eq_nat m n = integer_of_nat m <= integer_of_nat n;
+
+fun numadd (Cn (n1, c1, r1), Cn (n2, c2, r2)) =
+  (if equal_nat n1 n2
+    then let
+           val c = plus_inta c1 c2;
+         in
+           (if equal_inta c zero_inta then numadd (r1, r2)
+             else Cn (n1, c, numadd (r1, r2)))
+         end
+    else (if less_eq_nat n1 n2
+           then Cn (n1, c1, numadd (r1, Add (Mul (c2, Bound n2), r2)))
+           else Cn (n2, c2, numadd (Add (Mul (c1, Bound n1), r1), r2))))
+  | numadd (Cn (n1, c1, r1), C dd) = Cn (n1, c1, numadd (r1, C dd))
+  | numadd (Cn (n1, c1, r1), Bound de) = Cn (n1, c1, numadd (r1, Bound de))
+  | numadd (Cn (n1, c1, r1), Neg di) = Cn (n1, c1, numadd (r1, Neg di))
+  | numadd (Cn (n1, c1, r1), Add (dj, dk)) =
+    Cn (n1, c1, numadd (r1, Add (dj, dk)))
+  | numadd (Cn (n1, c1, r1), Sub (dl, dm)) =
+    Cn (n1, c1, numadd (r1, Sub (dl, dm)))
+  | numadd (Cn (n1, c1, r1), Mul (dn, doa)) =
+    Cn (n1, c1, numadd (r1, Mul (dn, doa)))
+  | numadd (C w, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (C w, r2))
+  | numadd (Bound x, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (Bound x, r2))
+  | numadd (Neg ac, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (Neg ac, r2))
+  | numadd (Add (ad, ae), Cn (n2, c2, r2)) =
+    Cn (n2, c2, numadd (Add (ad, ae), r2))
+  | numadd (Sub (af, ag), Cn (n2, c2, r2)) =
+    Cn (n2, c2, numadd (Sub (af, ag), r2))
+  | numadd (Mul (ah, ai), Cn (n2, c2, r2)) =
+    Cn (n2, c2, numadd (Mul (ah, ai), r2))
+  | numadd (C b1, C b2) = C (plus_inta b1 b2)
+  | numadd (C aj, Bound bi) = Add (C aj, Bound bi)
+  | numadd (C aj, Neg bm) = Add (C aj, Neg bm)
+  | numadd (C aj, Add (bn, bo)) = Add (C aj, Add (bn, bo))
+  | numadd (C aj, Sub (bp, bq)) = Add (C aj, Sub (bp, bq))
+  | numadd (C aj, Mul (br, bs)) = Add (C aj, Mul (br, bs))
+  | numadd (Bound ak, C cf) = Add (Bound ak, C cf)
+  | numadd (Bound ak, Bound cg) = Add (Bound ak, Bound cg)
+  | numadd (Bound ak, Neg ck) = Add (Bound ak, Neg ck)
+  | numadd (Bound ak, Add (cl, cm)) = Add (Bound ak, Add (cl, cm))
+  | numadd (Bound ak, Sub (cn, co)) = Add (Bound ak, Sub (cn, co))
+  | numadd (Bound ak, Mul (cp, cq)) = Add (Bound ak, Mul (cp, cq))
+  | numadd (Neg ao, C en) = Add (Neg ao, C en)
+  | numadd (Neg ao, Bound eo) = Add (Neg ao, Bound eo)
+  | numadd (Neg ao, Neg et) = Add (Neg ao, Neg et)
+  | numadd (Neg ao, Add (eu, ev)) = Add (Neg ao, Add (eu, ev))
+  | numadd (Neg ao, Sub (ew, ex)) = Add (Neg ao, Sub (ew, ex))
+  | numadd (Neg ao, Mul (ey, ez)) = Add (Neg ao, Mul (ey, ez))
+  | numadd (Add (ap, aq), C fm) = Add (Add (ap, aq), C fm)
+  | numadd (Add (ap, aq), Bound fna) = Add (Add (ap, aq), Bound fna)
+  | numadd (Add (ap, aq), Neg fr) = Add (Add (ap, aq), Neg fr)
+  | numadd (Add (ap, aq), Add (fs, ft)) = Add (Add (ap, aq), Add (fs, ft))
+  | numadd (Add (ap, aq), Sub (fu, fv)) = Add (Add (ap, aq), Sub (fu, fv))
+  | numadd (Add (ap, aq), Mul (fw, fx)) = Add (Add (ap, aq), Mul (fw, fx))
+  | numadd (Sub (ar, asa), C gk) = Add (Sub (ar, asa), C gk)
+  | numadd (Sub (ar, asa), Bound gl) = Add (Sub (ar, asa), Bound gl)
+  | numadd (Sub (ar, asa), Neg gp) = Add (Sub (ar, asa), Neg gp)
+  | numadd (Sub (ar, asa), Add (gq, gr)) = Add (Sub (ar, asa), Add (gq, gr))
+  | numadd (Sub (ar, asa), Sub (gs, gt)) = Add (Sub (ar, asa), Sub (gs, gt))
+  | numadd (Sub (ar, asa), Mul (gu, gv)) = Add (Sub (ar, asa), Mul (gu, gv))
+  | numadd (Mul (at, au), C hi) = Add (Mul (at, au), C hi)
+  | numadd (Mul (at, au), Bound hj) = Add (Mul (at, au), Bound hj)
+  | numadd (Mul (at, au), Neg hn) = Add (Mul (at, au), Neg hn)
+  | numadd (Mul (at, au), Add (ho, hp)) = Add (Mul (at, au), Add (ho, hp))
+  | numadd (Mul (at, au), Sub (hq, hr)) = Add (Mul (at, au), Sub (hq, hr))
+  | numadd (Mul (at, au), Mul (hs, ht)) = Add (Mul (at, au), Mul (hs, ht));
+
+fun numsub s t = (if equal_numa s t then C zero_inta else numadd (s, numneg t));
+
+fun simpnum (C j) = C j
+  | simpnum (Bound n) = Cn (n, Int_of_integer (1 : IntInf.int), C zero_inta)
+  | simpnum (Neg t) = numneg (simpnum t)
+  | simpnum (Add (t, s)) = numadd (simpnum t, simpnum s)
+  | simpnum (Sub (t, s)) = numsub (simpnum t) (simpnum s)
+  | simpnum (Mul (i, t)) =
+    (if equal_inta i zero_inta then C zero_inta else nummul i (simpnum t))
+  | simpnum (Cn (v, va, vb)) = Cn (v, va, vb);
+
+fun less_int k l = integer_of_int k < integer_of_int l;
+
+val equal_int = {equal = equal_inta} : inta equal;
+
+fun abs_int i = (if less_int i zero_inta then uminus_int i else i);
+
+fun nota (Not p) = p
+  | nota T = F
+  | nota F = T
+  | nota (Lt v) = Not (Lt v)
+  | nota (Le v) = Not (Le v)
+  | nota (Gt v) = Not (Gt v)
+  | nota (Ge v) = Not (Ge v)
+  | nota (Eq v) = Not (Eq v)
+  | nota (NEq v) = Not (NEq v)
+  | nota (Dvd (v, va)) = Not (Dvd (v, va))
+  | nota (NDvd (v, va)) = Not (NDvd (v, va))
+  | nota (And (v, va)) = Not (And (v, va))
+  | nota (Or (v, va)) = Not (Or (v, va))
+  | nota (Imp (v, va)) = Not (Imp (v, va))
+  | nota (Iff (v, va)) = Not (Iff (v, va))
+  | nota (E v) = Not (E v)
+  | nota (A v) = Not (A v)
+  | nota (Closed v) = Not (Closed v)
+  | nota (NClosed v) = Not (NClosed v);
+
+fun impa p q =
+  (if equal_fm p F orelse equal_fm q T then T
+    else (if equal_fm p T then q
+           else (if equal_fm q F then nota p else Imp (p, q))));
+
+fun iffa p q =
+  (if equal_fm p q then T
+    else (if equal_fm p (nota q) orelse equal_fm (nota p) q then F
+           else (if equal_fm p F then nota q
+                  else (if equal_fm q F then nota p
+                         else (if equal_fm p T then q
+                                else (if equal_fm q T then p
+                                       else Iff (p, q)))))));
+
+fun disj p q =
+  (if equal_fm p T orelse equal_fm q T then T
+    else (if equal_fm p F then q else (if equal_fm q F then p else Or (p, q))));
+
+fun conj p q =
+  (if equal_fm p F orelse equal_fm q F then F
+    else (if equal_fm p T then q
+           else (if equal_fm q T then p else And (p, q))));
+
+fun dvd (A1_, A2_) a b =
+  eq A2_ (moda (div_semiring_div A1_) b a)
+    (zero ((zero_mult_zero o mult_zero_semiring_0 o semiring_0_semiring_1 o
+             semiring_1_comm_semiring_1 o
+             comm_semiring_1_comm_semiring_1_cancel o
+             comm_semiring_1_cancel_semiring_div)
+            A1_));
 
-fun dj f p = evaldjf f (disjuncts p);
+fun simpfm (And (p, q)) = conj (simpfm p) (simpfm q)
+  | simpfm (Or (p, q)) = disj (simpfm p) (simpfm q)
+  | simpfm (Imp (p, q)) = impa (simpfm p) (simpfm q)
+  | simpfm (Iff (p, q)) = iffa (simpfm p) (simpfm q)
+  | simpfm (Not p) = nota (simpfm p)
+  | simpfm (Lt a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if less_int v zero_inta then T else F)
+        | Bound _ => Lt aa | Cn (_, _, _) => Lt aa | Neg _ => Lt aa
+        | Add (_, _) => Lt aa | Sub (_, _) => Lt aa | Mul (_, _) => Lt aa)
+    end
+  | simpfm (Le a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if less_eq_int v zero_inta then T else F)
+        | Bound _ => Le aa | Cn (_, _, _) => Le aa | Neg _ => Le aa
+        | Add (_, _) => Le aa | Sub (_, _) => Le aa | Mul (_, _) => Le aa)
+    end
+  | simpfm (Gt a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if less_int zero_inta v then T else F)
+        | Bound _ => Gt aa | Cn (_, _, _) => Gt aa | Neg _ => Gt aa
+        | Add (_, _) => Gt aa | Sub (_, _) => Gt aa | Mul (_, _) => Gt aa)
+    end
+  | simpfm (Ge a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if less_eq_int zero_inta v then T else F)
+        | Bound _ => Ge aa | Cn (_, _, _) => Ge aa | Neg _ => Ge aa
+        | Add (_, _) => Ge aa | Sub (_, _) => Ge aa | Mul (_, _) => Ge aa)
+    end
+  | simpfm (Eq a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if equal_inta v zero_inta then T else F)
+        | Bound _ => Eq aa | Cn (_, _, _) => Eq aa | Neg _ => Eq aa
+        | Add (_, _) => Eq aa | Sub (_, _) => Eq aa | Mul (_, _) => Eq aa)
+    end
+  | simpfm (NEq a) =
+    let
+      val aa = simpnum a;
+    in
+      (case aa of C v => (if not (equal_inta v zero_inta) then T else F)
+        | Bound _ => NEq aa | Cn (_, _, _) => NEq aa | Neg _ => NEq aa
+        | Add (_, _) => NEq aa | Sub (_, _) => NEq aa | Mul (_, _) => NEq aa)
+    end
+  | simpfm (Dvd (i, a)) =
+    (if equal_inta i zero_inta then simpfm (Eq a)
+      else (if equal_inta (abs_int i) (Int_of_integer (1 : IntInf.int)) then T
+             else let
+                    val aa = simpnum a;
+                  in
+                    (case aa
+                      of C v =>
+                        (if dvd (semiring_div_int, equal_int) i v then T else F)
+                      | Bound _ => Dvd (i, aa) | Cn (_, _, _) => Dvd (i, aa)
+                      | Neg _ => Dvd (i, aa) | Add (_, _) => Dvd (i, aa)
+                      | Sub (_, _) => Dvd (i, aa) | Mul (_, _) => Dvd (i, aa))
+                  end))
+  | simpfm (NDvd (i, a)) =
+    (if equal_inta i zero_inta then simpfm (NEq a)
+      else (if equal_inta (abs_int i) (Int_of_integer (1 : IntInf.int)) then F
+             else let
+                    val aa = simpnum a;
+                  in
+                    (case aa
+                      of C v =>
+                        (if not (dvd (semiring_div_int, equal_int) i v) then T
+                          else F)
+                      | Bound _ => NDvd (i, aa) | Cn (_, _, _) => NDvd (i, aa)
+                      | Neg _ => NDvd (i, aa) | Add (_, _) => NDvd (i, aa)
+                      | Sub (_, _) => NDvd (i, aa) | Mul (_, _) => NDvd (i, aa))
+                  end))
+  | simpfm T = T
+  | simpfm F = F
+  | simpfm (E v) = E v
+  | simpfm (A v) = A v
+  | simpfm (Closed v) = Closed v
+  | simpfm (NClosed v) = NClosed v;
+
+val equal_num = {equal = equal_numa} : numa equal;
+
+fun gen_length n (x :: xs) = gen_length (suc n) xs
+  | gen_length n [] = n;
+
+fun size_list x = gen_length zero_nat x;
+
+fun mirror (And (p, q)) = And (mirror p, mirror q)
+  | mirror (Or (p, q)) = Or (mirror p, mirror q)
+  | mirror T = T
+  | mirror F = F
+  | mirror (Lt (C bo)) = Lt (C bo)
+  | mirror (Lt (Bound bp)) = Lt (Bound bp)
+  | mirror (Lt (Neg bt)) = Lt (Neg bt)
+  | mirror (Lt (Add (bu, bv))) = Lt (Add (bu, bv))
+  | mirror (Lt (Sub (bw, bx))) = Lt (Sub (bw, bx))
+  | mirror (Lt (Mul (by, bz))) = Lt (Mul (by, bz))
+  | mirror (Le (C co)) = Le (C co)
+  | mirror (Le (Bound cp)) = Le (Bound cp)
+  | mirror (Le (Neg ct)) = Le (Neg ct)
+  | mirror (Le (Add (cu, cv))) = Le (Add (cu, cv))
+  | mirror (Le (Sub (cw, cx))) = Le (Sub (cw, cx))
+  | mirror (Le (Mul (cy, cz))) = Le (Mul (cy, cz))
+  | mirror (Gt (C doa)) = Gt (C doa)
+  | mirror (Gt (Bound dp)) = Gt (Bound dp)
+  | mirror (Gt (Neg dt)) = Gt (Neg dt)
+  | mirror (Gt (Add (du, dv))) = Gt (Add (du, dv))
+  | mirror (Gt (Sub (dw, dx))) = Gt (Sub (dw, dx))
+  | mirror (Gt (Mul (dy, dz))) = Gt (Mul (dy, dz))
+  | mirror (Ge (C eo)) = Ge (C eo)
+  | mirror (Ge (Bound ep)) = Ge (Bound ep)
+  | mirror (Ge (Neg et)) = Ge (Neg et)
+  | mirror (Ge (Add (eu, ev))) = Ge (Add (eu, ev))
+  | mirror (Ge (Sub (ew, ex))) = Ge (Sub (ew, ex))
+  | mirror (Ge (Mul (ey, ez))) = Ge (Mul (ey, ez))
+  | mirror (Eq (C fo)) = Eq (C fo)
+  | mirror (Eq (Bound fp)) = Eq (Bound fp)
+  | mirror (Eq (Neg ft)) = Eq (Neg ft)
+  | mirror (Eq (Add (fu, fv))) = Eq (Add (fu, fv))
+  | mirror (Eq (Sub (fw, fx))) = Eq (Sub (fw, fx))
+  | mirror (Eq (Mul (fy, fz))) = Eq (Mul (fy, fz))
+  | mirror (NEq (C go)) = NEq (C go)
+  | mirror (NEq (Bound gp)) = NEq (Bound gp)
+  | mirror (NEq (Neg gt)) = NEq (Neg gt)
+  | mirror (NEq (Add (gu, gv))) = NEq (Add (gu, gv))
+  | mirror (NEq (Sub (gw, gx))) = NEq (Sub (gw, gx))
+  | mirror (NEq (Mul (gy, gz))) = NEq (Mul (gy, gz))
+  | mirror (Dvd (aa, C ho)) = Dvd (aa, C ho)
+  | mirror (Dvd (aa, Bound hp)) = Dvd (aa, Bound hp)
+  | mirror (Dvd (aa, Neg ht)) = Dvd (aa, Neg ht)
+  | mirror (Dvd (aa, Add (hu, hv))) = Dvd (aa, Add (hu, hv))
+  | mirror (Dvd (aa, Sub (hw, hx))) = Dvd (aa, Sub (hw, hx))
+  | mirror (Dvd (aa, Mul (hy, hz))) = Dvd (aa, Mul (hy, hz))
+  | mirror (NDvd (ac, C io)) = NDvd (ac, C io)
+  | mirror (NDvd (ac, Bound ip)) = NDvd (ac, Bound ip)
+  | mirror (NDvd (ac, Neg it)) = NDvd (ac, Neg it)
+  | mirror (NDvd (ac, Add (iu, iv))) = NDvd (ac, Add (iu, iv))
+  | mirror (NDvd (ac, Sub (iw, ix))) = NDvd (ac, Sub (iw, ix))
+  | mirror (NDvd (ac, Mul (iy, iz))) = NDvd (ac, Mul (iy, iz))
+  | mirror (Not ae) = Not ae
+  | mirror (Imp (aj, ak)) = Imp (aj, ak)
+  | mirror (Iff (al, am)) = Iff (al, am)
+  | mirror (E an) = E an
+  | mirror (A ao) = A ao
+  | mirror (Closed ap) = Closed ap
+  | mirror (NClosed aq) = NClosed aq
+  | mirror (Lt (Cn (cm, c, e))) =
+    (if equal_nat cm zero_nat then Gt (Cn (zero_nat, c, Neg e))
+      else Lt (Cn (suc (minus_nat cm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | mirror (Le (Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat then Ge (Cn (zero_nat, c, Neg e))
+      else Le (Cn (suc (minus_nat dm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | mirror (Gt (Cn (em, c, e))) =
+    (if equal_nat em zero_nat then Lt (Cn (zero_nat, c, Neg e))
+      else Gt (Cn (suc (minus_nat em (nat_of_integer (1 : IntInf.int))), c, e)))
+  | mirror (Ge (Cn (fm, c, e))) =
+    (if equal_nat fm zero_nat then Le (Cn (zero_nat, c, Neg e))
+      else Ge (Cn (suc (minus_nat fm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | mirror (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat then Eq (Cn (zero_nat, c, Neg e))
+      else Eq (Cn (suc (minus_nat gm (nat_of_integer (1 : IntInf.int))), c, e)))
+  | mirror (NEq (Cn (hm, c, e))) =
+    (if equal_nat hm zero_nat then NEq (Cn (zero_nat, c, Neg e))
+      else NEq (Cn (suc (minus_nat hm (nat_of_integer (1 : IntInf.int))), c,
+                     e)))
+  | mirror (Dvd (i, Cn (im, c, e))) =
+    (if equal_nat im zero_nat then Dvd (i, Cn (zero_nat, c, Neg e))
+      else Dvd (i, Cn (suc (minus_nat im (nat_of_integer (1 : IntInf.int))), c,
+                        e)))
+  | mirror (NDvd (i, Cn (jm, c, e))) =
+    (if equal_nat jm zero_nat then NDvd (i, Cn (zero_nat, c, Neg e))
+      else NDvd (i, Cn (suc (minus_nat jm (nat_of_integer (1 : IntInf.int))), c,
+                         e)));
+
+fun a_beta (And (p, q)) = (fn k => And (a_beta p k, a_beta q k))
+  | a_beta (Or (p, q)) = (fn k => Or (a_beta p k, a_beta q k))
+  | a_beta T = (fn _ => T)
+  | a_beta F = (fn _ => F)
+  | a_beta (Lt (C bo)) = (fn _ => Lt (C bo))
+  | a_beta (Lt (Bound bp)) = (fn _ => Lt (Bound bp))
+  | a_beta (Lt (Neg bt)) = (fn _ => Lt (Neg bt))
+  | a_beta (Lt (Add (bu, bv))) = (fn _ => Lt (Add (bu, bv)))
+  | a_beta (Lt (Sub (bw, bx))) = (fn _ => Lt (Sub (bw, bx)))
+  | a_beta (Lt (Mul (by, bz))) = (fn _ => Lt (Mul (by, bz)))
+  | a_beta (Le (C co)) = (fn _ => Le (C co))
+  | a_beta (Le (Bound cp)) = (fn _ => Le (Bound cp))
+  | a_beta (Le (Neg ct)) = (fn _ => Le (Neg ct))
+  | a_beta (Le (Add (cu, cv))) = (fn _ => Le (Add (cu, cv)))
+  | a_beta (Le (Sub (cw, cx))) = (fn _ => Le (Sub (cw, cx)))
+  | a_beta (Le (Mul (cy, cz))) = (fn _ => Le (Mul (cy, cz)))
+  | a_beta (Gt (C doa)) = (fn _ => Gt (C doa))
+  | a_beta (Gt (Bound dp)) = (fn _ => Gt (Bound dp))
+  | a_beta (Gt (Neg dt)) = (fn _ => Gt (Neg dt))
+  | a_beta (Gt (Add (du, dv))) = (fn _ => Gt (Add (du, dv)))
+  | a_beta (Gt (Sub (dw, dx))) = (fn _ => Gt (Sub (dw, dx)))
+  | a_beta (Gt (Mul (dy, dz))) = (fn _ => Gt (Mul (dy, dz)))
+  | a_beta (Ge (C eo)) = (fn _ => Ge (C eo))
+  | a_beta (Ge (Bound ep)) = (fn _ => Ge (Bound ep))
+  | a_beta (Ge (Neg et)) = (fn _ => Ge (Neg et))
+  | a_beta (Ge (Add (eu, ev))) = (fn _ => Ge (Add (eu, ev)))
+  | a_beta (Ge (Sub (ew, ex))) = (fn _ => Ge (Sub (ew, ex)))
+  | a_beta (Ge (Mul (ey, ez))) = (fn _ => Ge (Mul (ey, ez)))
+  | a_beta (Eq (C fo)) = (fn _ => Eq (C fo))
+  | a_beta (Eq (Bound fp)) = (fn _ => Eq (Bound fp))
+  | a_beta (Eq (Neg ft)) = (fn _ => Eq (Neg ft))
+  | a_beta (Eq (Add (fu, fv))) = (fn _ => Eq (Add (fu, fv)))
+  | a_beta (Eq (Sub (fw, fx))) = (fn _ => Eq (Sub (fw, fx)))
+  | a_beta (Eq (Mul (fy, fz))) = (fn _ => Eq (Mul (fy, fz)))
+  | a_beta (NEq (C go)) = (fn _ => NEq (C go))
+  | a_beta (NEq (Bound gp)) = (fn _ => NEq (Bound gp))
+  | a_beta (NEq (Neg gt)) = (fn _ => NEq (Neg gt))
+  | a_beta (NEq (Add (gu, gv))) = (fn _ => NEq (Add (gu, gv)))
+  | a_beta (NEq (Sub (gw, gx))) = (fn _ => NEq (Sub (gw, gx)))
+  | a_beta (NEq (Mul (gy, gz))) = (fn _ => NEq (Mul (gy, gz)))
+  | a_beta (Dvd (aa, C ho)) = (fn _ => Dvd (aa, C ho))
+  | a_beta (Dvd (aa, Bound hp)) = (fn _ => Dvd (aa, Bound hp))
+  | a_beta (Dvd (aa, Neg ht)) = (fn _ => Dvd (aa, Neg ht))
+  | a_beta (Dvd (aa, Add (hu, hv))) = (fn _ => Dvd (aa, Add (hu, hv)))
+  | a_beta (Dvd (aa, Sub (hw, hx))) = (fn _ => Dvd (aa, Sub (hw, hx)))
+  | a_beta (Dvd (aa, Mul (hy, hz))) = (fn _ => Dvd (aa, Mul (hy, hz)))
+  | a_beta (NDvd (ac, C io)) = (fn _ => NDvd (ac, C io))
+  | a_beta (NDvd (ac, Bound ip)) = (fn _ => NDvd (ac, Bound ip))
+  | a_beta (NDvd (ac, Neg it)) = (fn _ => NDvd (ac, Neg it))
+  | a_beta (NDvd (ac, Add (iu, iv))) = (fn _ => NDvd (ac, Add (iu, iv)))
+  | a_beta (NDvd (ac, Sub (iw, ix))) = (fn _ => NDvd (ac, Sub (iw, ix)))
+  | a_beta (NDvd (ac, Mul (iy, iz))) = (fn _ => NDvd (ac, Mul (iy, iz)))
+  | a_beta (Not ae) = (fn _ => Not ae)
+  | a_beta (Imp (aj, ak)) = (fn _ => Imp (aj, ak))
+  | a_beta (Iff (al, am)) = (fn _ => Iff (al, am))
+  | a_beta (E an) = (fn _ => E an)
+  | a_beta (A ao) = (fn _ => A ao)
+  | a_beta (Closed ap) = (fn _ => Closed ap)
+  | a_beta (NClosed aq) = (fn _ => NClosed aq)
+  | a_beta (Lt (Cn (cm, c, e))) =
+    (if equal_nat cm zero_nat
+      then (fn k =>
+             Lt (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                      Mul (div_inta k c, e))))
+      else (fn _ =>
+             Lt (Cn (suc (minus_nat cm (nat_of_integer (1 : IntInf.int))), c,
+                      e))))
+  | a_beta (Le (Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat
+      then (fn k =>
+             Le (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                      Mul (div_inta k c, e))))
+      else (fn _ =>
+             Le (Cn (suc (minus_nat dm (nat_of_integer (1 : IntInf.int))), c,
+                      e))))
+  | a_beta (Gt (Cn (em, c, e))) =
+    (if equal_nat em zero_nat
+      then (fn k =>
+             Gt (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                      Mul (div_inta k c, e))))
+      else (fn _ =>
+             Gt (Cn (suc (minus_nat em (nat_of_integer (1 : IntInf.int))), c,
+                      e))))
+  | a_beta (Ge (Cn (fm, c, e))) =
+    (if equal_nat fm zero_nat
+      then (fn k =>
+             Ge (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                      Mul (div_inta k c, e))))
+      else (fn _ =>
+             Ge (Cn (suc (minus_nat fm (nat_of_integer (1 : IntInf.int))), c,
+                      e))))
+  | a_beta (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat
+      then (fn k =>
+             Eq (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                      Mul (div_inta k c, e))))
+      else (fn _ =>
+             Eq (Cn (suc (minus_nat gm (nat_of_integer (1 : IntInf.int))), c,
+                      e))))
+  | a_beta (NEq (Cn (hm, c, e))) =
+    (if equal_nat hm zero_nat
+      then (fn k =>
+             NEq (Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                       Mul (div_inta k c, e))))
+      else (fn _ =>
+             NEq (Cn (suc (minus_nat hm (nat_of_integer (1 : IntInf.int))), c,
+                       e))))
+  | a_beta (Dvd (i, Cn (im, c, e))) =
+    (if equal_nat im zero_nat
+      then (fn k =>
+             Dvd (times_inta (div_inta k c) i,
+                   Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                        Mul (div_inta k c, e))))
+      else (fn _ =>
+             Dvd (i, Cn (suc (minus_nat im (nat_of_integer (1 : IntInf.int))),
+                          c, e))))
+  | a_beta (NDvd (i, Cn (jm, c, e))) =
+    (if equal_nat jm zero_nat
+      then (fn k =>
+             NDvd (times_inta (div_inta k c) i,
+                    Cn (zero_nat, Int_of_integer (1 : IntInf.int),
+                         Mul (div_inta k c, e))))
+      else (fn _ =>
+             NDvd (i, Cn (suc (minus_nat jm (nat_of_integer (1 : IntInf.int))),
+                           c, e))));
+
+fun member A_ [] y = false
+  | member A_ (x :: xs) y = eq A_ x y orelse member A_ xs y;
+
+fun remdups A_ [] = []
+  | remdups A_ (x :: xs) =
+    (if member A_ xs x then remdups A_ xs else x :: remdups A_ xs);
+
+fun gcd_int k l =
+  abs_int
+    (if equal_inta l zero_inta then k
+      else gcd_int l (mod_int (abs_int k) (abs_int l)));
+
+fun lcm_int a b = div_inta (times_inta (abs_int a) (abs_int b)) (gcd_int a b);
+
+fun delta (And (p, q)) = lcm_int (delta p) (delta q)
+  | delta (Or (p, q)) = lcm_int (delta p) (delta q)
+  | delta T = Int_of_integer (1 : IntInf.int)
+  | delta F = Int_of_integer (1 : IntInf.int)
+  | delta (Lt u) = Int_of_integer (1 : IntInf.int)
+  | delta (Le v) = Int_of_integer (1 : IntInf.int)
+  | delta (Gt w) = Int_of_integer (1 : IntInf.int)
+  | delta (Ge x) = Int_of_integer (1 : IntInf.int)
+  | delta (Eq y) = Int_of_integer (1 : IntInf.int)
+  | delta (NEq z) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, C bo)) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, Bound bp)) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, Neg bt)) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, Add (bu, bv))) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, Sub (bw, bx))) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (aa, Mul (by, bz))) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, C co)) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, Bound cp)) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, Neg ct)) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, Add (cu, cv))) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, Sub (cw, cx))) = Int_of_integer (1 : IntInf.int)
+  | delta (NDvd (ac, Mul (cy, cz))) = Int_of_integer (1 : IntInf.int)
+  | delta (Not ae) = Int_of_integer (1 : IntInf.int)
+  | delta (Imp (aj, ak)) = Int_of_integer (1 : IntInf.int)
+  | delta (Iff (al, am)) = Int_of_integer (1 : IntInf.int)
+  | delta (E an) = Int_of_integer (1 : IntInf.int)
+  | delta (A ao) = Int_of_integer (1 : IntInf.int)
+  | delta (Closed ap) = Int_of_integer (1 : IntInf.int)
+  | delta (NClosed aq) = Int_of_integer (1 : IntInf.int)
+  | delta (Dvd (i, Cn (cm, c, e))) =
+    (if equal_nat cm zero_nat then i else Int_of_integer (1 : IntInf.int))
+  | delta (NDvd (i, Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat then i else Int_of_integer (1 : IntInf.int));
+
+fun alpha (And (p, q)) = alpha p @ alpha q
+  | alpha (Or (p, q)) = alpha p @ alpha q
+  | alpha T = []
+  | alpha F = []
+  | alpha (Lt (C bo)) = []
+  | alpha (Lt (Bound bp)) = []
+  | alpha (Lt (Neg bt)) = []
+  | alpha (Lt (Add (bu, bv))) = []
+  | alpha (Lt (Sub (bw, bx))) = []
+  | alpha (Lt (Mul (by, bz))) = []
+  | alpha (Le (C co)) = []
+  | alpha (Le (Bound cp)) = []
+  | alpha (Le (Neg ct)) = []
+  | alpha (Le (Add (cu, cv))) = []
+  | alpha (Le (Sub (cw, cx))) = []
+  | alpha (Le (Mul (cy, cz))) = []
+  | alpha (Gt (C doa)) = []
+  | alpha (Gt (Bound dp)) = []
+  | alpha (Gt (Neg dt)) = []
+  | alpha (Gt (Add (du, dv))) = []
+  | alpha (Gt (Sub (dw, dx))) = []
+  | alpha (Gt (Mul (dy, dz))) = []
+  | alpha (Ge (C eo)) = []
+  | alpha (Ge (Bound ep)) = []
+  | alpha (Ge (Neg et)) = []
+  | alpha (Ge (Add (eu, ev))) = []
+  | alpha (Ge (Sub (ew, ex))) = []
+  | alpha (Ge (Mul (ey, ez))) = []
+  | alpha (Eq (C fo)) = []
+  | alpha (Eq (Bound fp)) = []
+  | alpha (Eq (Neg ft)) = []
+  | alpha (Eq (Add (fu, fv))) = []
+  | alpha (Eq (Sub (fw, fx))) = []
+  | alpha (Eq (Mul (fy, fz))) = []
+  | alpha (NEq (C go)) = []
+  | alpha (NEq (Bound gp)) = []
+  | alpha (NEq (Neg gt)) = []
+  | alpha (NEq (Add (gu, gv))) = []
+  | alpha (NEq (Sub (gw, gx))) = []
+  | alpha (NEq (Mul (gy, gz))) = []
+  | alpha (Dvd (aa, ab)) = []
+  | alpha (NDvd (ac, ad)) = []
+  | alpha (Not ae) = []
+  | alpha (Imp (aj, ak)) = []
+  | alpha (Iff (al, am)) = []
+  | alpha (E an) = []
+  | alpha (A ao) = []
+  | alpha (Closed ap) = []
+  | alpha (NClosed aq) = []
+  | alpha (Lt (Cn (cm, c, e))) = (if equal_nat cm zero_nat then [e] else [])
+  | alpha (Le (Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat
+      then [Add (C (Int_of_integer (~1 : IntInf.int)), e)] else [])
+  | alpha (Gt (Cn (em, c, e))) = (if equal_nat em zero_nat then [] else [])
+  | alpha (Ge (Cn (fm, c, e))) = (if equal_nat fm zero_nat then [] else [])
+  | alpha (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat
+      then [Add (C (Int_of_integer (~1 : IntInf.int)), e)] else [])
+  | alpha (NEq (Cn (hm, c, e))) = (if equal_nat hm zero_nat then [e] else []);
+
+fun minus_int k l = Int_of_integer (integer_of_int k - integer_of_int l);
+
+fun zsplit0 (C c) = (zero_inta, C c)
+  | zsplit0 (Bound n) =
+    (if equal_nat n zero_nat then (Int_of_integer (1 : IntInf.int), C zero_inta)
+      else (zero_inta, Bound n))
+  | zsplit0 (Cn (n, i, a)) =
+    let
+      val (ia, aa) = zsplit0 a;
+    in
+      (if equal_nat n zero_nat then (plus_inta i ia, aa)
+        else (ia, Cn (n, i, aa)))
+    end
+  | zsplit0 (Neg a) = let
+                        val (i, aa) = zsplit0 a;
+                      in
+                        (uminus_int i, Neg aa)
+                      end
+  | zsplit0 (Add (a, b)) =
+    let
+      val (ia, aa) = zsplit0 a;
+      val (ib, ba) = zsplit0 b;
+    in
+      (plus_inta ia ib, Add (aa, ba))
+    end
+  | zsplit0 (Sub (a, b)) =
+    let
+      val (ia, aa) = zsplit0 a;
+      val (ib, ba) = zsplit0 b;
+    in
+      (minus_int ia ib, Sub (aa, ba))
+    end
+  | zsplit0 (Mul (i, a)) =
+    let
+      val (ia, aa) = zsplit0 a;
+    in
+      (times_inta i ia, Mul (i, aa))
+    end;
+
+fun zlfm (And (p, q)) = And (zlfm p, zlfm q)
+  | zlfm (Or (p, q)) = Or (zlfm p, zlfm q)
+  | zlfm (Imp (p, q)) = Or (zlfm (Not p), zlfm q)
+  | zlfm (Iff (p, q)) =
+    Or (And (zlfm p, zlfm q), And (zlfm (Not p), zlfm (Not q)))
+  | zlfm (Lt a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then Lt r
+        else (if less_int zero_inta c then Lt (Cn (zero_nat, c, r))
+               else Gt (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (Le a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then Le r
+        else (if less_int zero_inta c then Le (Cn (zero_nat, c, r))
+               else Ge (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (Gt a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then Gt r
+        else (if less_int zero_inta c then Gt (Cn (zero_nat, c, r))
+               else Lt (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (Ge a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then Ge r
+        else (if less_int zero_inta c then Ge (Cn (zero_nat, c, r))
+               else Le (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (Eq a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then Eq r
+        else (if less_int zero_inta c then Eq (Cn (zero_nat, c, r))
+               else Eq (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (NEq a) =
+    let
+      val (c, r) = zsplit0 a;
+    in
+      (if equal_inta c zero_inta then NEq r
+        else (if less_int zero_inta c then NEq (Cn (zero_nat, c, r))
+               else NEq (Cn (zero_nat, uminus_int c, Neg r))))
+    end
+  | zlfm (Dvd (i, a)) =
+    (if equal_inta i zero_inta then zlfm (Eq a)
+      else let
+             val (c, r) = zsplit0 a;
+           in
+             (if equal_inta c zero_inta then Dvd (abs_int i, r)
+               else (if less_int zero_inta c
+                      then Dvd (abs_int i, Cn (zero_nat, c, r))
+                      else Dvd (abs_int i, Cn (zero_nat, uminus_int c, Neg r))))
+           end)
+  | zlfm (NDvd (i, a)) =
+    (if equal_inta i zero_inta then zlfm (NEq a)
+      else let
+             val (c, r) = zsplit0 a;
+           in
+             (if equal_inta c zero_inta then NDvd (abs_int i, r)
+               else (if less_int zero_inta c
+                      then NDvd (abs_int i, Cn (zero_nat, c, r))
+                      else NDvd (abs_int i,
+                                  Cn (zero_nat, uminus_int c, Neg r))))
+           end)
+  | zlfm (Not (And (p, q))) = Or (zlfm (Not p), zlfm (Not q))
+  | zlfm (Not (Or (p, q))) = And (zlfm (Not p), zlfm (Not q))
+  | zlfm (Not (Imp (p, q))) = And (zlfm p, zlfm (Not q))
+  | zlfm (Not (Iff (p, q))) =
+    Or (And (zlfm p, zlfm (Not q)), And (zlfm (Not p), zlfm q))
+  | zlfm (Not (Not p)) = zlfm p
+  | zlfm (Not T) = F
+  | zlfm (Not F) = T
+  | zlfm (Not (Lt a)) = zlfm (Ge a)
+  | zlfm (Not (Le a)) = zlfm (Gt a)
+  | zlfm (Not (Gt a)) = zlfm (Le a)
+  | zlfm (Not (Ge a)) = zlfm (Lt a)
+  | zlfm (Not (Eq a)) = zlfm (NEq a)
+  | zlfm (Not (NEq a)) = zlfm (Eq a)
+  | zlfm (Not (Dvd (i, a))) = zlfm (NDvd (i, a))
+  | zlfm (Not (NDvd (i, a))) = zlfm (Dvd (i, a))
+  | zlfm (Not (Closed p)) = NClosed p
+  | zlfm (Not (NClosed p)) = Closed p
+  | zlfm T = T
+  | zlfm F = F
+  | zlfm (Not (E ci)) = Not (E ci)
+  | zlfm (Not (A cj)) = Not (A cj)
+  | zlfm (E ao) = E ao
+  | zlfm (A ap) = A ap
+  | zlfm (Closed aq) = Closed aq
+  | zlfm (NClosed ar) = NClosed ar;
+
+fun zeta (And (p, q)) = lcm_int (zeta p) (zeta q)
+  | zeta (Or (p, q)) = lcm_int (zeta p) (zeta q)
+  | zeta T = Int_of_integer (1 : IntInf.int)
+  | zeta F = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (C bo)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Bound bp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Neg bt)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Add (bu, bv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Sub (bw, bx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Mul (by, bz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (C co)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (Bound cp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (Neg ct)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (Add (cu, cv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (Sub (cw, cx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Le (Mul (cy, cz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (C doa)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (Bound dp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (Neg dt)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (Add (du, dv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (Sub (dw, dx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Gt (Mul (dy, dz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (C eo)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (Bound ep)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (Neg et)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (Add (eu, ev))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (Sub (ew, ex))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Ge (Mul (ey, ez))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (C fo)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (Bound fp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (Neg ft)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (Add (fu, fv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (Sub (fw, fx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Eq (Mul (fy, fz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (C go)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (Bound gp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (Neg gt)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (Add (gu, gv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (Sub (gw, gx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NEq (Mul (gy, gz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, C ho)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, Bound hp)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, Neg ht)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, Add (hu, hv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, Sub (hw, hx))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Dvd (aa, Mul (hy, hz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, C io)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, Bound ip)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, Neg it)) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, Add (iu, iv))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, Sub (iw, ix))) = Int_of_integer (1 : IntInf.int)
+  | zeta (NDvd (ac, Mul (iy, iz))) = Int_of_integer (1 : IntInf.int)
+  | zeta (Not ae) = Int_of_integer (1 : IntInf.int)
+  | zeta (Imp (aj, ak)) = Int_of_integer (1 : IntInf.int)
+  | zeta (Iff (al, am)) = Int_of_integer (1 : IntInf.int)
+  | zeta (E an) = Int_of_integer (1 : IntInf.int)
+  | zeta (A ao) = Int_of_integer (1 : IntInf.int)
+  | zeta (Closed ap) = Int_of_integer (1 : IntInf.int)
+  | zeta (NClosed aq) = Int_of_integer (1 : IntInf.int)
+  | zeta (Lt (Cn (cm, c, e))) =
+    (if equal_nat cm zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (Le (Cn (dm, c, e))) =
+    (if equal_nat dm zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (Gt (Cn (em, c, e))) =
+    (if equal_nat em zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (Ge (Cn (fm, c, e))) =
+    (if equal_nat fm zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (NEq (Cn (hm, c, e))) =
+    (if equal_nat hm zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (Dvd (i, Cn (im, c, e))) =
+    (if equal_nat im zero_nat then c else Int_of_integer (1 : IntInf.int))
+  | zeta (NDvd (i, Cn (jm, c, e))) =
+    (if equal_nat jm zero_nat then c else Int_of_integer (1 : IntInf.int));
+
+fun beta (And (p, q)) = beta p @ beta q
+  | beta (Or (p, q)) = beta p @ beta q
+  | beta T = []
+  | beta F = []
+  | beta (Lt (C bo)) = []
+  | beta (Lt (Bound bp)) = []
+  | beta (Lt (Neg bt)) = []
+  | beta (Lt (Add (bu, bv))) = []
+  | beta (Lt (Sub (bw, bx))) = []
+  | beta (Lt (Mul (by, bz))) = []
+  | beta (Le (C co)) = []
+  | beta (Le (Bound cp)) = []
+  | beta (Le (Neg ct)) = []
+  | beta (Le (Add (cu, cv))) = []
+  | beta (Le (Sub (cw, cx))) = []
+  | beta (Le (Mul (cy, cz))) = []
+  | beta (Gt (C doa)) = []
+  | beta (Gt (Bound dp)) = []
+  | beta (Gt (Neg dt)) = []
+  | beta (Gt (Add (du, dv))) = []
+  | beta (Gt (Sub (dw, dx))) = []
+  | beta (Gt (Mul (dy, dz))) = []
+  | beta (Ge (C eo)) = []
+  | beta (Ge (Bound ep)) = []
+  | beta (Ge (Neg et)) = []
+  | beta (Ge (Add (eu, ev))) = []
+  | beta (Ge (Sub (ew, ex))) = []
+  | beta (Ge (Mul (ey, ez))) = []
+  | beta (Eq (C fo)) = []
+  | beta (Eq (Bound fp)) = []
+  | beta (Eq (Neg ft)) = []
+  | beta (Eq (Add (fu, fv))) = []
+  | beta (Eq (Sub (fw, fx))) = []
+  | beta (Eq (Mul (fy, fz))) = []
+  | beta (NEq (C go)) = []
+  | beta (NEq (Bound gp)) = []
+  | beta (NEq (Neg gt)) = []
+  | beta (NEq (Add (gu, gv))) = []
+  | beta (NEq (Sub (gw, gx))) = []
+  | beta (NEq (Mul (gy, gz))) = []
+  | beta (Dvd (aa, ab)) = []
+  | beta (NDvd (ac, ad)) = []
+  | beta (Not ae) = []
+  | beta (Imp (aj, ak)) = []
+  | beta (Iff (al, am)) = []
+  | beta (E an) = []
+  | beta (A ao) = []
+  | beta (Closed ap) = []
+  | beta (NClosed aq) = []
+  | beta (Lt (Cn (cm, c, e))) = (if equal_nat cm zero_nat then [] else [])
+  | beta (Le (Cn (dm, c, e))) = (if equal_nat dm zero_nat then [] else [])
+  | beta (Gt (Cn (em, c, e))) = (if equal_nat em zero_nat then [Neg e] else [])
+  | beta (Ge (Cn (fm, c, e))) =
+    (if equal_nat fm zero_nat
+      then [Sub (C (Int_of_integer (~1 : IntInf.int)), e)] else [])
+  | beta (Eq (Cn (gm, c, e))) =
+    (if equal_nat gm zero_nat
+      then [Sub (C (Int_of_integer (~1 : IntInf.int)), e)] else [])
+  | beta (NEq (Cn (hm, c, e))) =
+    (if equal_nat hm zero_nat then [Neg e] else []);
+
+fun unita p =
+  let
+    val pa = zlfm p;
+    val l = zeta pa;
+    val q =
+      And (Dvd (l, Cn (zero_nat, Int_of_integer (1 : IntInf.int), C zero_inta)),
+            a_beta pa l);
+    val d = delta q;
+    val b = remdups equal_num (map simpnum (beta q));
+    val a = remdups equal_num (map simpnum (alpha q));
+  in
+    (if less_eq_nat (size_list b) (size_list a) then (q, (b, d))
+      else (mirror q, (a, d)))
+  end;
+
+fun decrnum (Bound n) = Bound (minus_nat n (nat_of_integer (1 : IntInf.int)))
+  | decrnum (Neg a) = Neg (decrnum a)
+  | decrnum (Add (a, b)) = Add (decrnum a, decrnum b)
+  | decrnum (Sub (a, b)) = Sub (decrnum a, decrnum b)
+  | decrnum (Mul (c, a)) = Mul (c, decrnum a)
+  | decrnum (Cn (n, i, a)) =
+    Cn (minus_nat n (nat_of_integer (1 : IntInf.int)), i, decrnum a)
+  | decrnum (C v) = C v;
+
+fun decr (Lt a) = Lt (decrnum a)
+  | decr (Le a) = Le (decrnum a)
+  | decr (Gt a) = Gt (decrnum a)
+  | decr (Ge a) = Ge (decrnum a)
+  | decr (Eq a) = Eq (decrnum a)
+  | decr (NEq a) = NEq (decrnum a)
+  | decr (Dvd (i, a)) = Dvd (i, decrnum a)
+  | decr (NDvd (i, a)) = NDvd (i, decrnum a)
+  | decr (Not p) = Not (decr p)
+  | decr (And (p, q)) = And (decr p, decr q)
+  | decr (Or (p, q)) = Or (decr p, decr q)
+  | decr (Imp (p, q)) = Imp (decr p, decr q)
+  | decr (Iff (p, q)) = Iff (decr p, decr q)
+  | decr T = T
+  | decr F = F
+  | decr (E v) = E v
+  | decr (A v) = A v
+  | decr (Closed v) = Closed v
+  | decr (NClosed v) = NClosed v;
+
+fun uptoa i j =
+  (if less_eq_int i j
+    then i :: uptoa (plus_inta i (Int_of_integer (1 : IntInf.int))) j else []);
+
+fun maps f [] = []
+  | maps f (x :: xs) = f x @ maps f xs;
+
+fun cooper p =
+  let
+    val (q, (b, d)) = unita p;
+    val js = uptoa (Int_of_integer (1 : IntInf.int)) d;
+    val mq = simpfm (minusinf q);
+    val md = evaldjf (fn j => simpfm (subst0 (C j) mq)) js;
+  in
+    (if equal_fm md T then T
+      else let
+             val qd =
+               evaldjf (fn (ba, j) => simpfm (subst0 (Add (ba, C j)) q))
+                 (maps (fn ba => map (fn a => (ba, a)) js) b);
+           in
+             decr (disj md qd)
+           end)
+  end;
+
+fun qelim (E p) = (fn qe => dj qe (qelim p qe))
+  | qelim (A p) = (fn qe => nota (qe (qelim (Not p) qe)))
+  | qelim (Not p) = (fn qe => nota (qelim p qe))
+  | qelim (And (p, q)) = (fn qe => conj (qelim p qe) (qelim q qe))
+  | qelim (Or (p, q)) = (fn qe => disj (qelim p qe) (qelim q qe))
+  | qelim (Imp (p, q)) = (fn qe => impa (qelim p qe) (qelim q qe))
+  | qelim (Iff (p, q)) = (fn qe => iffa (qelim p qe) (qelim q qe))
+  | qelim T = (fn _ => simpfm T)
+  | qelim F = (fn _ => simpfm F)
+  | qelim (Lt v) = (fn _ => simpfm (Lt v))
+  | qelim (Le v) = (fn _ => simpfm (Le v))
+  | qelim (Gt v) = (fn _ => simpfm (Gt v))
+  | qelim (Ge v) = (fn _ => simpfm (Ge v))
+  | qelim (Eq v) = (fn _ => simpfm (Eq v))
+  | qelim (NEq v) = (fn _ => simpfm (NEq v))
+  | qelim (Dvd (v, va)) = (fn _ => simpfm (Dvd (v, va)))
+  | qelim (NDvd (v, va)) = (fn _ => simpfm (NDvd (v, va)))
+  | qelim (Closed v) = (fn _ => simpfm (Closed v))
+  | qelim (NClosed v) = (fn _ => simpfm (NClosed v));
 
 fun prep (E T) = T
   | prep (E F) = F
@@ -787,1369 +2302,6 @@
   | prep (Closed ap) = Closed ap
   | prep (NClosed aq) = NClosed aq;
 
-fun conj p q =
-  (if equal_fm p F orelse equal_fm q F then F
-    else (if equal_fm p T then q
-           else (if equal_fm q T then p else And (p, q))));
-
-fun disj p q =
-  (if equal_fm p T orelse equal_fm q T then T
-    else (if equal_fm p F then q else (if equal_fm q F then p else Or (p, q))));
-
-fun nota (Not p) = p
-  | nota T = F
-  | nota F = T
-  | nota (Lt v) = Not (Lt v)
-  | nota (Le v) = Not (Le v)
-  | nota (Gt v) = Not (Gt v)
-  | nota (Ge v) = Not (Ge v)
-  | nota (Eq v) = Not (Eq v)
-  | nota (NEq v) = Not (NEq v)
-  | nota (Dvd (v, va)) = Not (Dvd (v, va))
-  | nota (NDvd (v, va)) = Not (NDvd (v, va))
-  | nota (And (v, va)) = Not (And (v, va))
-  | nota (Or (v, va)) = Not (Or (v, va))
-  | nota (Imp (v, va)) = Not (Imp (v, va))
-  | nota (Iff (v, va)) = Not (Iff (v, va))
-  | nota (E v) = Not (E v)
-  | nota (A v) = Not (A v)
-  | nota (Closed v) = Not (Closed v)
-  | nota (NClosed v) = Not (NClosed v);
-
-fun iffa p q =
-  (if equal_fm p q then T
-    else (if equal_fm p (nota q) orelse equal_fm (nota p) q then F
-           else (if equal_fm p F then nota q
-                  else (if equal_fm q F then nota p
-                         else (if equal_fm p T then q
-                                else (if equal_fm q T then p
-                                       else Iff (p, q)))))));
-
-fun impa p q =
-  (if equal_fm p F orelse equal_fm q T then T
-    else (if equal_fm p T then q
-           else (if equal_fm q F then nota p else Imp (p, q))));
-
-type 'a times = {times : 'a -> 'a -> 'a};
-val times = #times : 'a times -> 'a -> 'a -> 'a;
-
-type 'a dvd = {times_dvd : 'a times};
-val times_dvd = #times_dvd : 'a dvd -> 'a times;
-
-type 'a diva = {dvd_div : 'a dvd, diva : 'a -> 'a -> 'a, moda : 'a -> 'a -> 'a};
-val dvd_div = #dvd_div : 'a diva -> 'a dvd;
-val diva = #diva : 'a diva -> 'a -> 'a -> 'a;
-val moda = #moda : 'a diva -> 'a -> 'a -> 'a;
-
-type 'a zero = {zero : 'a};
-val zero = #zero : 'a zero -> 'a;
-
-type 'a no_zero_divisors =
-  {times_no_zero_divisors : 'a times, zero_no_zero_divisors : 'a zero};
-val times_no_zero_divisors = #times_no_zero_divisors :
-  'a no_zero_divisors -> 'a times;
-val zero_no_zero_divisors = #zero_no_zero_divisors :
-  'a no_zero_divisors -> 'a zero;
-
-type 'a semigroup_mult = {times_semigroup_mult : 'a times};
-val times_semigroup_mult = #times_semigroup_mult :
-  'a semigroup_mult -> 'a times;
-
-type 'a plus = {plus : 'a -> 'a -> 'a};
-val plus = #plus : 'a plus -> 'a -> 'a -> 'a;
-
-type 'a semigroup_add = {plus_semigroup_add : 'a plus};
-val plus_semigroup_add = #plus_semigroup_add : 'a semigroup_add -> 'a plus;
-
-type 'a ab_semigroup_add = {semigroup_add_ab_semigroup_add : 'a semigroup_add};
-val semigroup_add_ab_semigroup_add = #semigroup_add_ab_semigroup_add :
-  'a ab_semigroup_add -> 'a semigroup_add;
-
-type 'a semiring =
-  {ab_semigroup_add_semiring : 'a ab_semigroup_add,
-    semigroup_mult_semiring : 'a semigroup_mult};
-val ab_semigroup_add_semiring = #ab_semigroup_add_semiring :
-  'a semiring -> 'a ab_semigroup_add;
-val semigroup_mult_semiring = #semigroup_mult_semiring :
-  'a semiring -> 'a semigroup_mult;
-
-type 'a mult_zero = {times_mult_zero : 'a times, zero_mult_zero : 'a zero};
-val times_mult_zero = #times_mult_zero : 'a mult_zero -> 'a times;
-val zero_mult_zero = #zero_mult_zero : 'a mult_zero -> 'a zero;
-
-type 'a monoid_add =
-  {semigroup_add_monoid_add : 'a semigroup_add, zero_monoid_add : 'a zero};
-val semigroup_add_monoid_add = #semigroup_add_monoid_add :
-  'a monoid_add -> 'a semigroup_add;
-val zero_monoid_add = #zero_monoid_add : 'a monoid_add -> 'a zero;
-
-type 'a comm_monoid_add =
-  {ab_semigroup_add_comm_monoid_add : 'a ab_semigroup_add,
-    monoid_add_comm_monoid_add : 'a monoid_add};
-val ab_semigroup_add_comm_monoid_add = #ab_semigroup_add_comm_monoid_add :
-  'a comm_monoid_add -> 'a ab_semigroup_add;
-val monoid_add_comm_monoid_add = #monoid_add_comm_monoid_add :
-  'a comm_monoid_add -> 'a monoid_add;
-
-type 'a semiring_0 =
-  {comm_monoid_add_semiring_0 : 'a comm_monoid_add,
-    mult_zero_semiring_0 : 'a mult_zero, semiring_semiring_0 : 'a semiring};
-val comm_monoid_add_semiring_0 = #comm_monoid_add_semiring_0 :
-  'a semiring_0 -> 'a comm_monoid_add;
-val mult_zero_semiring_0 = #mult_zero_semiring_0 :
-  'a semiring_0 -> 'a mult_zero;
-val semiring_semiring_0 = #semiring_semiring_0 : 'a semiring_0 -> 'a semiring;
-
-type 'a one = {one : 'a};
-val one = #one : 'a one -> 'a;
-
-type 'a power = {one_power : 'a one, times_power : 'a times};
-val one_power = #one_power : 'a power -> 'a one;
-val times_power = #times_power : 'a power -> 'a times;
-
-type 'a monoid_mult =
-  {semigroup_mult_monoid_mult : 'a semigroup_mult,
-    power_monoid_mult : 'a power};
-val semigroup_mult_monoid_mult = #semigroup_mult_monoid_mult :
-  'a monoid_mult -> 'a semigroup_mult;
-val power_monoid_mult = #power_monoid_mult : 'a monoid_mult -> 'a power;
-
-type 'a zero_neq_one = {one_zero_neq_one : 'a one, zero_zero_neq_one : 'a zero};
-val one_zero_neq_one = #one_zero_neq_one : 'a zero_neq_one -> 'a one;
-val zero_zero_neq_one = #zero_zero_neq_one : 'a zero_neq_one -> 'a zero;
-
-type 'a semiring_1 =
-  {monoid_mult_semiring_1 : 'a monoid_mult,
-    semiring_0_semiring_1 : 'a semiring_0,
-    zero_neq_one_semiring_1 : 'a zero_neq_one};
-val monoid_mult_semiring_1 = #monoid_mult_semiring_1 :
-  'a semiring_1 -> 'a monoid_mult;
-val semiring_0_semiring_1 = #semiring_0_semiring_1 :
-  'a semiring_1 -> 'a semiring_0;
-val zero_neq_one_semiring_1 = #zero_neq_one_semiring_1 :
-  'a semiring_1 -> 'a zero_neq_one;
-
-type 'a ab_semigroup_mult =
-  {semigroup_mult_ab_semigroup_mult : 'a semigroup_mult};
-val semigroup_mult_ab_semigroup_mult = #semigroup_mult_ab_semigroup_mult :
-  'a ab_semigroup_mult -> 'a semigroup_mult;
-
-type 'a comm_semiring =
-  {ab_semigroup_mult_comm_semiring : 'a ab_semigroup_mult,
-    semiring_comm_semiring : 'a semiring};
-val ab_semigroup_mult_comm_semiring = #ab_semigroup_mult_comm_semiring :
-  'a comm_semiring -> 'a ab_semigroup_mult;
-val semiring_comm_semiring = #semiring_comm_semiring :
-  'a comm_semiring -> 'a semiring;
-
-type 'a comm_semiring_0 =
-  {comm_semiring_comm_semiring_0 : 'a comm_semiring,
-    semiring_0_comm_semiring_0 : 'a semiring_0};
-val comm_semiring_comm_semiring_0 = #comm_semiring_comm_semiring_0 :
-  'a comm_semiring_0 -> 'a comm_semiring;
-val semiring_0_comm_semiring_0 = #semiring_0_comm_semiring_0 :
-  'a comm_semiring_0 -> 'a semiring_0;
-
-type 'a comm_monoid_mult =
-  {ab_semigroup_mult_comm_monoid_mult : 'a ab_semigroup_mult,
-    monoid_mult_comm_monoid_mult : 'a monoid_mult};
-val ab_semigroup_mult_comm_monoid_mult = #ab_semigroup_mult_comm_monoid_mult :
-  'a comm_monoid_mult -> 'a ab_semigroup_mult;
-val monoid_mult_comm_monoid_mult = #monoid_mult_comm_monoid_mult :
-  'a comm_monoid_mult -> 'a monoid_mult;
-
-type 'a comm_semiring_1 =
-  {comm_monoid_mult_comm_semiring_1 : 'a comm_monoid_mult,
-    comm_semiring_0_comm_semiring_1 : 'a comm_semiring_0,
-    dvd_comm_semiring_1 : 'a dvd, semiring_1_comm_semiring_1 : 'a semiring_1};
-val comm_monoid_mult_comm_semiring_1 = #comm_monoid_mult_comm_semiring_1 :
-  'a comm_semiring_1 -> 'a comm_monoid_mult;
-val comm_semiring_0_comm_semiring_1 = #comm_semiring_0_comm_semiring_1 :
-  'a comm_semiring_1 -> 'a comm_semiring_0;
-val dvd_comm_semiring_1 = #dvd_comm_semiring_1 : 'a comm_semiring_1 -> 'a dvd;
-val semiring_1_comm_semiring_1 = #semiring_1_comm_semiring_1 :
-  'a comm_semiring_1 -> 'a semiring_1;
-
-type 'a cancel_semigroup_add =
-  {semigroup_add_cancel_semigroup_add : 'a semigroup_add};
-val semigroup_add_cancel_semigroup_add = #semigroup_add_cancel_semigroup_add :
-  'a cancel_semigroup_add -> 'a semigroup_add;
-
-type 'a cancel_ab_semigroup_add =
-  {ab_semigroup_add_cancel_ab_semigroup_add : 'a ab_semigroup_add,
-    cancel_semigroup_add_cancel_ab_semigroup_add : 'a cancel_semigroup_add};
-val ab_semigroup_add_cancel_ab_semigroup_add =
-  #ab_semigroup_add_cancel_ab_semigroup_add :
-  'a cancel_ab_semigroup_add -> 'a ab_semigroup_add;
-val cancel_semigroup_add_cancel_ab_semigroup_add =
-  #cancel_semigroup_add_cancel_ab_semigroup_add :
-  'a cancel_ab_semigroup_add -> 'a cancel_semigroup_add;
-
-type 'a cancel_comm_monoid_add =
-  {cancel_ab_semigroup_add_cancel_comm_monoid_add : 'a cancel_ab_semigroup_add,
-    comm_monoid_add_cancel_comm_monoid_add : 'a comm_monoid_add};
-val cancel_ab_semigroup_add_cancel_comm_monoid_add =
-  #cancel_ab_semigroup_add_cancel_comm_monoid_add :
-  'a cancel_comm_monoid_add -> 'a cancel_ab_semigroup_add;
-val comm_monoid_add_cancel_comm_monoid_add =
-  #comm_monoid_add_cancel_comm_monoid_add :
-  'a cancel_comm_monoid_add -> 'a comm_monoid_add;
-
-type 'a semiring_0_cancel =
-  {cancel_comm_monoid_add_semiring_0_cancel : 'a cancel_comm_monoid_add,
-    semiring_0_semiring_0_cancel : 'a semiring_0};
-val cancel_comm_monoid_add_semiring_0_cancel =
-  #cancel_comm_monoid_add_semiring_0_cancel :
-  'a semiring_0_cancel -> 'a cancel_comm_monoid_add;
-val semiring_0_semiring_0_cancel = #semiring_0_semiring_0_cancel :
-  'a semiring_0_cancel -> 'a semiring_0;
-
-type 'a semiring_1_cancel =
-  {semiring_0_cancel_semiring_1_cancel : 'a semiring_0_cancel,
-    semiring_1_semiring_1_cancel : 'a semiring_1};
-val semiring_0_cancel_semiring_1_cancel = #semiring_0_cancel_semiring_1_cancel :
-  'a semiring_1_cancel -> 'a semiring_0_cancel;
-val semiring_1_semiring_1_cancel = #semiring_1_semiring_1_cancel :
-  'a semiring_1_cancel -> 'a semiring_1;
-
-type 'a comm_semiring_0_cancel =
-  {comm_semiring_0_comm_semiring_0_cancel : 'a comm_semiring_0,
-    semiring_0_cancel_comm_semiring_0_cancel : 'a semiring_0_cancel};
-val comm_semiring_0_comm_semiring_0_cancel =
-  #comm_semiring_0_comm_semiring_0_cancel :
-  'a comm_semiring_0_cancel -> 'a comm_semiring_0;
-val semiring_0_cancel_comm_semiring_0_cancel =
-  #semiring_0_cancel_comm_semiring_0_cancel :
-  'a comm_semiring_0_cancel -> 'a semiring_0_cancel;
-
-type 'a comm_semiring_1_cancel =
-  {comm_semiring_0_cancel_comm_semiring_1_cancel : 'a comm_semiring_0_cancel,
-    comm_semiring_1_comm_semiring_1_cancel : 'a comm_semiring_1,
-    semiring_1_cancel_comm_semiring_1_cancel : 'a semiring_1_cancel};
-val comm_semiring_0_cancel_comm_semiring_1_cancel =
-  #comm_semiring_0_cancel_comm_semiring_1_cancel :
-  'a comm_semiring_1_cancel -> 'a comm_semiring_0_cancel;
-val comm_semiring_1_comm_semiring_1_cancel =
-  #comm_semiring_1_comm_semiring_1_cancel :
-  'a comm_semiring_1_cancel -> 'a comm_semiring_1;
-val semiring_1_cancel_comm_semiring_1_cancel =
-  #semiring_1_cancel_comm_semiring_1_cancel :
-  'a comm_semiring_1_cancel -> 'a semiring_1_cancel;
-
-type 'a semiring_div =
-  {div_semiring_div : 'a diva,
-    comm_semiring_1_cancel_semiring_div : 'a comm_semiring_1_cancel,
-    no_zero_divisors_semiring_div : 'a no_zero_divisors};
-val div_semiring_div = #div_semiring_div : 'a semiring_div -> 'a diva;
-val comm_semiring_1_cancel_semiring_div = #comm_semiring_1_cancel_semiring_div :
-  'a semiring_div -> 'a comm_semiring_1_cancel;
-val no_zero_divisors_semiring_div = #no_zero_divisors_semiring_div :
-  'a semiring_div -> 'a no_zero_divisors;
-
-fun dvd (A1_, A2_) a b =
-  eq A2_ (moda (div_semiring_div A1_) b a)
-    (zero ((zero_mult_zero o mult_zero_semiring_0 o semiring_0_semiring_1 o
-             semiring_1_comm_semiring_1 o
-             comm_semiring_1_comm_semiring_1_cancel o
-             comm_semiring_1_cancel_semiring_div)
-            A1_));
-
-fun abs_int i = (if i < (0 : IntInf.int) then ~ i else i);
-
-val equal_int = {equal = (fn a => fn b => a = b)} : int equal;
-
-fun numadd (Cn (n1, c1, r1), Cn (n2, c2, r2)) =
-  (if n1 = n2
-    then let
-           val c = c1 + c2;
-         in
-           (if c = (0 : IntInf.int) then numadd (r1, r2)
-             else Cn (n1, c, numadd (r1, r2)))
-         end
-    else (if n1 <= n2
-           then Cn (n1, c1, numadd (r1, Add (Mul (c2, Bound n2), r2)))
-           else Cn (n2, c2, numadd (Add (Mul (c1, Bound n1), r1), r2))))
-  | numadd (Cn (n1, c1, r1), C dd) = Cn (n1, c1, numadd (r1, C dd))
-  | numadd (Cn (n1, c1, r1), Bound de) = Cn (n1, c1, numadd (r1, Bound de))
-  | numadd (Cn (n1, c1, r1), Neg di) = Cn (n1, c1, numadd (r1, Neg di))
-  | numadd (Cn (n1, c1, r1), Add (dj, dk)) =
-    Cn (n1, c1, numadd (r1, Add (dj, dk)))
-  | numadd (Cn (n1, c1, r1), Sub (dl, dm)) =
-    Cn (n1, c1, numadd (r1, Sub (dl, dm)))
-  | numadd (Cn (n1, c1, r1), Mul (dn, doa)) =
-    Cn (n1, c1, numadd (r1, Mul (dn, doa)))
-  | numadd (C w, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (C w, r2))
-  | numadd (Bound x, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (Bound x, r2))
-  | numadd (Neg ac, Cn (n2, c2, r2)) = Cn (n2, c2, numadd (Neg ac, r2))
-  | numadd (Add (ad, ae), Cn (n2, c2, r2)) =
-    Cn (n2, c2, numadd (Add (ad, ae), r2))
-  | numadd (Sub (af, ag), Cn (n2, c2, r2)) =
-    Cn (n2, c2, numadd (Sub (af, ag), r2))
-  | numadd (Mul (ah, ai), Cn (n2, c2, r2)) =
-    Cn (n2, c2, numadd (Mul (ah, ai), r2))
-  | numadd (C b1, C b2) = C (b1 + b2)
-  | numadd (C aj, Bound bi) = Add (C aj, Bound bi)
-  | numadd (C aj, Neg bm) = Add (C aj, Neg bm)
-  | numadd (C aj, Add (bn, bo)) = Add (C aj, Add (bn, bo))
-  | numadd (C aj, Sub (bp, bq)) = Add (C aj, Sub (bp, bq))
-  | numadd (C aj, Mul (br, bs)) = Add (C aj, Mul (br, bs))
-  | numadd (Bound ak, C cf) = Add (Bound ak, C cf)
-  | numadd (Bound ak, Bound cg) = Add (Bound ak, Bound cg)
-  | numadd (Bound ak, Neg ck) = Add (Bound ak, Neg ck)
-  | numadd (Bound ak, Add (cl, cm)) = Add (Bound ak, Add (cl, cm))
-  | numadd (Bound ak, Sub (cn, co)) = Add (Bound ak, Sub (cn, co))
-  | numadd (Bound ak, Mul (cp, cq)) = Add (Bound ak, Mul (cp, cq))
-  | numadd (Neg ao, C en) = Add (Neg ao, C en)
-  | numadd (Neg ao, Bound eo) = Add (Neg ao, Bound eo)
-  | numadd (Neg ao, Neg et) = Add (Neg ao, Neg et)
-  | numadd (Neg ao, Add (eu, ev)) = Add (Neg ao, Add (eu, ev))
-  | numadd (Neg ao, Sub (ew, ex)) = Add (Neg ao, Sub (ew, ex))
-  | numadd (Neg ao, Mul (ey, ez)) = Add (Neg ao, Mul (ey, ez))
-  | numadd (Add (ap, aq), C fm) = Add (Add (ap, aq), C fm)
-  | numadd (Add (ap, aq), Bound fna) = Add (Add (ap, aq), Bound fna)
-  | numadd (Add (ap, aq), Neg fr) = Add (Add (ap, aq), Neg fr)
-  | numadd (Add (ap, aq), Add (fs, ft)) = Add (Add (ap, aq), Add (fs, ft))
-  | numadd (Add (ap, aq), Sub (fu, fv)) = Add (Add (ap, aq), Sub (fu, fv))
-  | numadd (Add (ap, aq), Mul (fw, fx)) = Add (Add (ap, aq), Mul (fw, fx))
-  | numadd (Sub (ar, asa), C gk) = Add (Sub (ar, asa), C gk)
-  | numadd (Sub (ar, asa), Bound gl) = Add (Sub (ar, asa), Bound gl)
-  | numadd (Sub (ar, asa), Neg gp) = Add (Sub (ar, asa), Neg gp)
-  | numadd (Sub (ar, asa), Add (gq, gr)) = Add (Sub (ar, asa), Add (gq, gr))
-  | numadd (Sub (ar, asa), Sub (gs, gt)) = Add (Sub (ar, asa), Sub (gs, gt))
-  | numadd (Sub (ar, asa), Mul (gu, gv)) = Add (Sub (ar, asa), Mul (gu, gv))
-  | numadd (Mul (at, au), C hi) = Add (Mul (at, au), C hi)
-  | numadd (Mul (at, au), Bound hj) = Add (Mul (at, au), Bound hj)
-  | numadd (Mul (at, au), Neg hn) = Add (Mul (at, au), Neg hn)
-  | numadd (Mul (at, au), Add (ho, hp)) = Add (Mul (at, au), Add (ho, hp))
-  | numadd (Mul (at, au), Sub (hq, hr)) = Add (Mul (at, au), Sub (hq, hr))
-  | numadd (Mul (at, au), Mul (hs, ht)) = Add (Mul (at, au), Mul (hs, ht));
-
-fun nummul i (C j) = C (i * j)
-  | nummul i (Cn (n, c, t)) = Cn (n, c * i, nummul i t)
-  | nummul i (Bound v) = Mul (i, Bound v)
-  | nummul i (Neg v) = Mul (i, Neg v)
-  | nummul i (Add (v, va)) = Mul (i, Add (v, va))
-  | nummul i (Sub (v, va)) = Mul (i, Sub (v, va))
-  | nummul i (Mul (v, va)) = Mul (i, Mul (v, va));
-
-fun numneg t = nummul (~ (1 : IntInf.int)) t;
-
-fun numsub s t =
-  (if equal_numa s t then C (0 : IntInf.int) else numadd (s, numneg t));
-
-fun simpnum (C j) = C j
-  | simpnum (Bound n) = Cn (n, (1 : IntInf.int), C (0 : IntInf.int))
-  | simpnum (Neg t) = numneg (simpnum t)
-  | simpnum (Add (t, s)) = numadd (simpnum t, simpnum s)
-  | simpnum (Sub (t, s)) = numsub (simpnum t) (simpnum s)
-  | simpnum (Mul (i, t)) =
-    (if i = (0 : IntInf.int) then C (0 : IntInf.int) else nummul i (simpnum t))
-  | simpnum (Cn (v, va, vb)) = Cn (v, va, vb);
-
-val one_inta : int = (1 : IntInf.int);
-
-val zero_inta : int = (0 : IntInf.int);
-
-val times_int = {times = (fn a => fn b => a * b)} : int times;
-
-val dvd_int = {times_dvd = times_int} : int dvd;
-
-fun fst (a, b) = a;
-
-fun sgn_int i =
-  (if i = (0 : IntInf.int) then (0 : IntInf.int)
-    else (if (0 : IntInf.int) < i then (1 : IntInf.int)
-           else ~ (1 : IntInf.int)));
-
-fun apsnd f (x, y) = (x, f y);
-
-fun divmod_int k l =
-  (if k = (0 : IntInf.int) then ((0 : IntInf.int), (0 : IntInf.int))
-    else (if l = (0 : IntInf.int) then ((0 : IntInf.int), k)
-           else apsnd (fn a => sgn_int l * a)
-                  (if sgn_int k = sgn_int l then Integer.div_mod (abs k) (abs l)
-                    else let
-                           val (r, s) = Integer.div_mod (abs k) (abs l);
-                         in
-                           (if s = (0 : IntInf.int) then (~ r, (0 : IntInf.int))
-                             else (~ r - (1 : IntInf.int), abs_int l - s))
-                         end)));
-
-fun div_inta a b = fst (divmod_int a b);
-
-fun snd (a, b) = b;
-
-fun mod_int a b = snd (divmod_int a b);
-
-val div_int = {dvd_div = dvd_int, diva = div_inta, moda = mod_int} : int diva;
-
-val zero_int = {zero = zero_inta} : int zero;
-
-val no_zero_divisors_int =
-  {times_no_zero_divisors = times_int, zero_no_zero_divisors = zero_int} :
-  int no_zero_divisors;
-
-val semigroup_mult_int = {times_semigroup_mult = times_int} :
-  int semigroup_mult;
-
-val plus_int = {plus = (fn a => fn b => a + b)} : int plus;
-
-val semigroup_add_int = {plus_semigroup_add = plus_int} : int semigroup_add;
-
-val ab_semigroup_add_int = {semigroup_add_ab_semigroup_add = semigroup_add_int}
-  : int ab_semigroup_add;
-
-val semiring_int =
-  {ab_semigroup_add_semiring = ab_semigroup_add_int,
-    semigroup_mult_semiring = semigroup_mult_int}
-  : int semiring;
-
-val mult_zero_int = {times_mult_zero = times_int, zero_mult_zero = zero_int} :
-  int mult_zero;
-
-val monoid_add_int =
-  {semigroup_add_monoid_add = semigroup_add_int, zero_monoid_add = zero_int} :
-  int monoid_add;
-
-val comm_monoid_add_int =
-  {ab_semigroup_add_comm_monoid_add = ab_semigroup_add_int,
-    monoid_add_comm_monoid_add = monoid_add_int}
-  : int comm_monoid_add;
-
-val semiring_0_int =
-  {comm_monoid_add_semiring_0 = comm_monoid_add_int,
-    mult_zero_semiring_0 = mult_zero_int, semiring_semiring_0 = semiring_int}
-  : int semiring_0;
-
-val one_int = {one = one_inta} : int one;
-
-val power_int = {one_power = one_int, times_power = times_int} : int power;
-
-val monoid_mult_int =
-  {semigroup_mult_monoid_mult = semigroup_mult_int,
-    power_monoid_mult = power_int}
-  : int monoid_mult;
-
-val zero_neq_one_int =
-  {one_zero_neq_one = one_int, zero_zero_neq_one = zero_int} : int zero_neq_one;
-
-val semiring_1_int =
-  {monoid_mult_semiring_1 = monoid_mult_int,
-    semiring_0_semiring_1 = semiring_0_int,
-    zero_neq_one_semiring_1 = zero_neq_one_int}
-  : int semiring_1;
-
-val ab_semigroup_mult_int =
-  {semigroup_mult_ab_semigroup_mult = semigroup_mult_int} :
-  int ab_semigroup_mult;
-
-val comm_semiring_int =
-  {ab_semigroup_mult_comm_semiring = ab_semigroup_mult_int,
-    semiring_comm_semiring = semiring_int}
-  : int comm_semiring;
-
-val comm_semiring_0_int =
-  {comm_semiring_comm_semiring_0 = comm_semiring_int,
-    semiring_0_comm_semiring_0 = semiring_0_int}
-  : int comm_semiring_0;
-
-val comm_monoid_mult_int =
-  {ab_semigroup_mult_comm_monoid_mult = ab_semigroup_mult_int,
-    monoid_mult_comm_monoid_mult = monoid_mult_int}
-  : int comm_monoid_mult;
-
-val comm_semiring_1_int =
-  {comm_monoid_mult_comm_semiring_1 = comm_monoid_mult_int,
-    comm_semiring_0_comm_semiring_1 = comm_semiring_0_int,
-    dvd_comm_semiring_1 = dvd_int, semiring_1_comm_semiring_1 = semiring_1_int}
-  : int comm_semiring_1;
-
-val cancel_semigroup_add_int =
-  {semigroup_add_cancel_semigroup_add = semigroup_add_int} :
-  int cancel_semigroup_add;
-
-val cancel_ab_semigroup_add_int =
-  {ab_semigroup_add_cancel_ab_semigroup_add = ab_semigroup_add_int,
-    cancel_semigroup_add_cancel_ab_semigroup_add = cancel_semigroup_add_int}
-  : int cancel_ab_semigroup_add;
-
-val cancel_comm_monoid_add_int =
-  {cancel_ab_semigroup_add_cancel_comm_monoid_add = cancel_ab_semigroup_add_int,
-    comm_monoid_add_cancel_comm_monoid_add = comm_monoid_add_int}
-  : int cancel_comm_monoid_add;
-
-val semiring_0_cancel_int =
-  {cancel_comm_monoid_add_semiring_0_cancel = cancel_comm_monoid_add_int,
-    semiring_0_semiring_0_cancel = semiring_0_int}
-  : int semiring_0_cancel;
-
-val semiring_1_cancel_int =
-  {semiring_0_cancel_semiring_1_cancel = semiring_0_cancel_int,
-    semiring_1_semiring_1_cancel = semiring_1_int}
-  : int semiring_1_cancel;
-
-val comm_semiring_0_cancel_int =
-  {comm_semiring_0_comm_semiring_0_cancel = comm_semiring_0_int,
-    semiring_0_cancel_comm_semiring_0_cancel = semiring_0_cancel_int}
-  : int comm_semiring_0_cancel;
-
-val comm_semiring_1_cancel_int =
-  {comm_semiring_0_cancel_comm_semiring_1_cancel = comm_semiring_0_cancel_int,
-    comm_semiring_1_comm_semiring_1_cancel = comm_semiring_1_int,
-    semiring_1_cancel_comm_semiring_1_cancel = semiring_1_cancel_int}
-  : int comm_semiring_1_cancel;
-
-val semiring_div_int =
-  {div_semiring_div = div_int,
-    comm_semiring_1_cancel_semiring_div = comm_semiring_1_cancel_int,
-    no_zero_divisors_semiring_div = no_zero_divisors_int}
-  : int semiring_div;
-
-fun simpfm (And (p, q)) = conj (simpfm p) (simpfm q)
-  | simpfm (Or (p, q)) = disj (simpfm p) (simpfm q)
-  | simpfm (Imp (p, q)) = impa (simpfm p) (simpfm q)
-  | simpfm (Iff (p, q)) = iffa (simpfm p) (simpfm q)
-  | simpfm (Not p) = nota (simpfm p)
-  | simpfm (Lt a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if v < (0 : IntInf.int) then T else F)
-        | Bound _ => Lt aa | Cn (_, _, _) => Lt aa | Neg _ => Lt aa
-        | Add (_, _) => Lt aa | Sub (_, _) => Lt aa | Mul (_, _) => Lt aa)
-    end
-  | simpfm (Le a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if v <= (0 : IntInf.int) then T else F)
-        | Bound _ => Le aa | Cn (_, _, _) => Le aa | Neg _ => Le aa
-        | Add (_, _) => Le aa | Sub (_, _) => Le aa | Mul (_, _) => Le aa)
-    end
-  | simpfm (Gt a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if (0 : IntInf.int) < v then T else F)
-        | Bound _ => Gt aa | Cn (_, _, _) => Gt aa | Neg _ => Gt aa
-        | Add (_, _) => Gt aa | Sub (_, _) => Gt aa | Mul (_, _) => Gt aa)
-    end
-  | simpfm (Ge a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if (0 : IntInf.int) <= v then T else F)
-        | Bound _ => Ge aa | Cn (_, _, _) => Ge aa | Neg _ => Ge aa
-        | Add (_, _) => Ge aa | Sub (_, _) => Ge aa | Mul (_, _) => Ge aa)
-    end
-  | simpfm (Eq a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if v = (0 : IntInf.int) then T else F)
-        | Bound _ => Eq aa | Cn (_, _, _) => Eq aa | Neg _ => Eq aa
-        | Add (_, _) => Eq aa | Sub (_, _) => Eq aa | Mul (_, _) => Eq aa)
-    end
-  | simpfm (NEq a) =
-    let
-      val aa = simpnum a;
-    in
-      (case aa of C v => (if not (v = (0 : IntInf.int)) then T else F)
-        | Bound _ => NEq aa | Cn (_, _, _) => NEq aa | Neg _ => NEq aa
-        | Add (_, _) => NEq aa | Sub (_, _) => NEq aa | Mul (_, _) => NEq aa)
-    end
-  | simpfm (Dvd (i, a)) =
-    (if i = (0 : IntInf.int) then simpfm (Eq a)
-      else (if abs_int i = (1 : IntInf.int) then T
-             else let
-                    val aa = simpnum a;
-                  in
-                    (case aa
-                      of C v =>
-                        (if dvd (semiring_div_int, equal_int) i v then T else F)
-                      | Bound _ => Dvd (i, aa) | Cn (_, _, _) => Dvd (i, aa)
-                      | Neg _ => Dvd (i, aa) | Add (_, _) => Dvd (i, aa)
-                      | Sub (_, _) => Dvd (i, aa) | Mul (_, _) => Dvd (i, aa))
-                  end))
-  | simpfm (NDvd (i, a)) =
-    (if i = (0 : IntInf.int) then simpfm (NEq a)
-      else (if abs_int i = (1 : IntInf.int) then F
-             else let
-                    val aa = simpnum a;
-                  in
-                    (case aa
-                      of C v =>
-                        (if not (dvd (semiring_div_int, equal_int) i v) then T
-                          else F)
-                      | Bound _ => NDvd (i, aa) | Cn (_, _, _) => NDvd (i, aa)
-                      | Neg _ => NDvd (i, aa) | Add (_, _) => NDvd (i, aa)
-                      | Sub (_, _) => NDvd (i, aa) | Mul (_, _) => NDvd (i, aa))
-                  end))
-  | simpfm T = T
-  | simpfm F = F
-  | simpfm (E v) = E v
-  | simpfm (A v) = A v
-  | simpfm (Closed v) = Closed v
-  | simpfm (NClosed v) = NClosed v;
-
-fun qelim (E p) = (fn qe => dj qe (qelim p qe))
-  | qelim (A p) = (fn qe => nota (qe (qelim (Not p) qe)))
-  | qelim (Not p) = (fn qe => nota (qelim p qe))
-  | qelim (And (p, q)) = (fn qe => conj (qelim p qe) (qelim q qe))
-  | qelim (Or (p, q)) = (fn qe => disj (qelim p qe) (qelim q qe))
-  | qelim (Imp (p, q)) = (fn qe => impa (qelim p qe) (qelim q qe))
-  | qelim (Iff (p, q)) = (fn qe => iffa (qelim p qe) (qelim q qe))
-  | qelim T = (fn _ => simpfm T)
-  | qelim F = (fn _ => simpfm F)
-  | qelim (Lt v) = (fn _ => simpfm (Lt v))
-  | qelim (Le v) = (fn _ => simpfm (Le v))
-  | qelim (Gt v) = (fn _ => simpfm (Gt v))
-  | qelim (Ge v) = (fn _ => simpfm (Ge v))
-  | qelim (Eq v) = (fn _ => simpfm (Eq v))
-  | qelim (NEq v) = (fn _ => simpfm (NEq v))
-  | qelim (Dvd (v, va)) = (fn _ => simpfm (Dvd (v, va)))
-  | qelim (NDvd (v, va)) = (fn _ => simpfm (NDvd (v, va)))
-  | qelim (Closed v) = (fn _ => simpfm (Closed v))
-  | qelim (NClosed v) = (fn _ => simpfm (NClosed v));
-
-fun maps f [] = []
-  | maps f (x :: xs) = f x @ maps f xs;
-
-fun uptoa i j = (if i <= j then i :: uptoa (i + (1 : IntInf.int)) j else []);
-
-fun minus_nat n m = Integer.max (n - m) 0;
-
-fun decrnum (Bound n) = Bound (minus_nat n (1 : IntInf.int))
-  | decrnum (Neg a) = Neg (decrnum a)
-  | decrnum (Add (a, b)) = Add (decrnum a, decrnum b)
-  | decrnum (Sub (a, b)) = Sub (decrnum a, decrnum b)
-  | decrnum (Mul (c, a)) = Mul (c, decrnum a)
-  | decrnum (Cn (n, i, a)) = Cn (minus_nat n (1 : IntInf.int), i, decrnum a)
-  | decrnum (C v) = C v;
-
-fun decr (Lt a) = Lt (decrnum a)
-  | decr (Le a) = Le (decrnum a)
-  | decr (Gt a) = Gt (decrnum a)
-  | decr (Ge a) = Ge (decrnum a)
-  | decr (Eq a) = Eq (decrnum a)
-  | decr (NEq a) = NEq (decrnum a)
-  | decr (Dvd (i, a)) = Dvd (i, decrnum a)
-  | decr (NDvd (i, a)) = NDvd (i, decrnum a)
-  | decr (Not p) = Not (decr p)
-  | decr (And (p, q)) = And (decr p, decr q)
-  | decr (Or (p, q)) = Or (decr p, decr q)
-  | decr (Imp (p, q)) = Imp (decr p, decr q)
-  | decr (Iff (p, q)) = Iff (decr p, decr q)
-  | decr T = T
-  | decr F = F
-  | decr (E v) = E v
-  | decr (A v) = A v
-  | decr (Closed v) = Closed v
-  | decr (NClosed v) = NClosed v;
-
-fun beta (And (p, q)) = beta p @ beta q
-  | beta (Or (p, q)) = beta p @ beta q
-  | beta T = []
-  | beta F = []
-  | beta (Lt (C bo)) = []
-  | beta (Lt (Bound bp)) = []
-  | beta (Lt (Neg bt)) = []
-  | beta (Lt (Add (bu, bv))) = []
-  | beta (Lt (Sub (bw, bx))) = []
-  | beta (Lt (Mul (by, bz))) = []
-  | beta (Le (C co)) = []
-  | beta (Le (Bound cp)) = []
-  | beta (Le (Neg ct)) = []
-  | beta (Le (Add (cu, cv))) = []
-  | beta (Le (Sub (cw, cx))) = []
-  | beta (Le (Mul (cy, cz))) = []
-  | beta (Gt (C doa)) = []
-  | beta (Gt (Bound dp)) = []
-  | beta (Gt (Neg dt)) = []
-  | beta (Gt (Add (du, dv))) = []
-  | beta (Gt (Sub (dw, dx))) = []
-  | beta (Gt (Mul (dy, dz))) = []
-  | beta (Ge (C eo)) = []
-  | beta (Ge (Bound ep)) = []
-  | beta (Ge (Neg et)) = []
-  | beta (Ge (Add (eu, ev))) = []
-  | beta (Ge (Sub (ew, ex))) = []
-  | beta (Ge (Mul (ey, ez))) = []
-  | beta (Eq (C fo)) = []
-  | beta (Eq (Bound fp)) = []
-  | beta (Eq (Neg ft)) = []
-  | beta (Eq (Add (fu, fv))) = []
-  | beta (Eq (Sub (fw, fx))) = []
-  | beta (Eq (Mul (fy, fz))) = []
-  | beta (NEq (C go)) = []
-  | beta (NEq (Bound gp)) = []
-  | beta (NEq (Neg gt)) = []
-  | beta (NEq (Add (gu, gv))) = []
-  | beta (NEq (Sub (gw, gx))) = []
-  | beta (NEq (Mul (gy, gz))) = []
-  | beta (Dvd (aa, ab)) = []
-  | beta (NDvd (ac, ad)) = []
-  | beta (Not ae) = []
-  | beta (Imp (aj, ak)) = []
-  | beta (Iff (al, am)) = []
-  | beta (E an) = []
-  | beta (A ao) = []
-  | beta (Closed ap) = []
-  | beta (NClosed aq) = []
-  | beta (Lt (Cn (cm, c, e))) = (if cm = (0 : IntInf.int) then [] else [])
-  | beta (Le (Cn (dm, c, e))) = (if dm = (0 : IntInf.int) then [] else [])
-  | beta (Gt (Cn (em, c, e))) = (if em = (0 : IntInf.int) then [Neg e] else [])
-  | beta (Ge (Cn (fm, c, e))) =
-    (if fm = (0 : IntInf.int) then [Sub (C (~1 : IntInf.int), e)] else [])
-  | beta (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int) then [Sub (C (~1 : IntInf.int), e)] else [])
-  | beta (NEq (Cn (hm, c, e))) =
-    (if hm = (0 : IntInf.int) then [Neg e] else []);
-
-fun gcd_int k l =
-  abs_int
-    (if l = (0 : IntInf.int) then k
-      else gcd_int l (mod_int (abs_int k) (abs_int l)));
-
-fun lcm_int a b = div_inta (abs_int a * abs_int b) (gcd_int a b);
-
-fun zeta (And (p, q)) = lcm_int (zeta p) (zeta q)
-  | zeta (Or (p, q)) = lcm_int (zeta p) (zeta q)
-  | zeta T = (1 : IntInf.int)
-  | zeta F = (1 : IntInf.int)
-  | zeta (Lt (C bo)) = (1 : IntInf.int)
-  | zeta (Lt (Bound bp)) = (1 : IntInf.int)
-  | zeta (Lt (Neg bt)) = (1 : IntInf.int)
-  | zeta (Lt (Add (bu, bv))) = (1 : IntInf.int)
-  | zeta (Lt (Sub (bw, bx))) = (1 : IntInf.int)
-  | zeta (Lt (Mul (by, bz))) = (1 : IntInf.int)
-  | zeta (Le (C co)) = (1 : IntInf.int)
-  | zeta (Le (Bound cp)) = (1 : IntInf.int)
-  | zeta (Le (Neg ct)) = (1 : IntInf.int)
-  | zeta (Le (Add (cu, cv))) = (1 : IntInf.int)
-  | zeta (Le (Sub (cw, cx))) = (1 : IntInf.int)
-  | zeta (Le (Mul (cy, cz))) = (1 : IntInf.int)
-  | zeta (Gt (C doa)) = (1 : IntInf.int)
-  | zeta (Gt (Bound dp)) = (1 : IntInf.int)
-  | zeta (Gt (Neg dt)) = (1 : IntInf.int)
-  | zeta (Gt (Add (du, dv))) = (1 : IntInf.int)
-  | zeta (Gt (Sub (dw, dx))) = (1 : IntInf.int)
-  | zeta (Gt (Mul (dy, dz))) = (1 : IntInf.int)
-  | zeta (Ge (C eo)) = (1 : IntInf.int)
-  | zeta (Ge (Bound ep)) = (1 : IntInf.int)
-  | zeta (Ge (Neg et)) = (1 : IntInf.int)
-  | zeta (Ge (Add (eu, ev))) = (1 : IntInf.int)
-  | zeta (Ge (Sub (ew, ex))) = (1 : IntInf.int)
-  | zeta (Ge (Mul (ey, ez))) = (1 : IntInf.int)
-  | zeta (Eq (C fo)) = (1 : IntInf.int)
-  | zeta (Eq (Bound fp)) = (1 : IntInf.int)
-  | zeta (Eq (Neg ft)) = (1 : IntInf.int)
-  | zeta (Eq (Add (fu, fv))) = (1 : IntInf.int)
-  | zeta (Eq (Sub (fw, fx))) = (1 : IntInf.int)
-  | zeta (Eq (Mul (fy, fz))) = (1 : IntInf.int)
-  | zeta (NEq (C go)) = (1 : IntInf.int)
-  | zeta (NEq (Bound gp)) = (1 : IntInf.int)
-  | zeta (NEq (Neg gt)) = (1 : IntInf.int)
-  | zeta (NEq (Add (gu, gv))) = (1 : IntInf.int)
-  | zeta (NEq (Sub (gw, gx))) = (1 : IntInf.int)
-  | zeta (NEq (Mul (gy, gz))) = (1 : IntInf.int)
-  | zeta (Dvd (aa, C ho)) = (1 : IntInf.int)
-  | zeta (Dvd (aa, Bound hp)) = (1 : IntInf.int)
-  | zeta (Dvd (aa, Neg ht)) = (1 : IntInf.int)
-  | zeta (Dvd (aa, Add (hu, hv))) = (1 : IntInf.int)
-  | zeta (Dvd (aa, Sub (hw, hx))) = (1 : IntInf.int)
-  | zeta (Dvd (aa, Mul (hy, hz))) = (1 : IntInf.int)
-  | zeta (NDvd (ac, C io)) = (1 : IntInf.int)
-  | zeta (NDvd (ac, Bound ip)) = (1 : IntInf.int)
-  | zeta (NDvd (ac, Neg it)) = (1 : IntInf.int)
-  | zeta (NDvd (ac, Add (iu, iv))) = (1 : IntInf.int)
-  | zeta (NDvd (ac, Sub (iw, ix))) = (1 : IntInf.int)
-  | zeta (NDvd (ac, Mul (iy, iz))) = (1 : IntInf.int)
-  | zeta (Not ae) = (1 : IntInf.int)
-  | zeta (Imp (aj, ak)) = (1 : IntInf.int)
-  | zeta (Iff (al, am)) = (1 : IntInf.int)
-  | zeta (E an) = (1 : IntInf.int)
-  | zeta (A ao) = (1 : IntInf.int)
-  | zeta (Closed ap) = (1 : IntInf.int)
-  | zeta (NClosed aq) = (1 : IntInf.int)
-  | zeta (Lt (Cn (cm, c, e))) =
-    (if cm = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (Le (Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (Gt (Cn (em, c, e))) =
-    (if em = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (Ge (Cn (fm, c, e))) =
-    (if fm = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (NEq (Cn (hm, c, e))) =
-    (if hm = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (Dvd (i, Cn (im, c, e))) =
-    (if im = (0 : IntInf.int) then c else (1 : IntInf.int))
-  | zeta (NDvd (i, Cn (jm, c, e))) =
-    (if jm = (0 : IntInf.int) then c else (1 : IntInf.int));
-
-fun zsplit0 (C c) = ((0 : IntInf.int), C c)
-  | zsplit0 (Bound n) =
-    (if n = (0 : IntInf.int) then ((1 : IntInf.int), C (0 : IntInf.int))
-      else ((0 : IntInf.int), Bound n))
-  | zsplit0 (Cn (n, i, a)) =
-    let
-      val (ia, aa) = zsplit0 a;
-    in
-      (if n = (0 : IntInf.int) then (i + ia, aa) else (ia, Cn (n, i, aa)))
-    end
-  | zsplit0 (Neg a) = let
-                        val (i, aa) = zsplit0 a;
-                      in
-                        (~ i, Neg aa)
-                      end
-  | zsplit0 (Add (a, b)) =
-    let
-      val (ia, aa) = zsplit0 a;
-      val (ib, ba) = zsplit0 b;
-    in
-      (ia + ib, Add (aa, ba))
-    end
-  | zsplit0 (Sub (a, b)) =
-    let
-      val (ia, aa) = zsplit0 a;
-      val (ib, ba) = zsplit0 b;
-    in
-      (ia - ib, Sub (aa, ba))
-    end
-  | zsplit0 (Mul (i, a)) =
-    let
-      val (ia, aa) = zsplit0 a;
-    in
-      (i * ia, Mul (i, aa))
-    end;
-
-fun zlfm (And (p, q)) = And (zlfm p, zlfm q)
-  | zlfm (Or (p, q)) = Or (zlfm p, zlfm q)
-  | zlfm (Imp (p, q)) = Or (zlfm (Not p), zlfm q)
-  | zlfm (Iff (p, q)) =
-    Or (And (zlfm p, zlfm q), And (zlfm (Not p), zlfm (Not q)))
-  | zlfm (Lt a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then Lt r
-        else (if (0 : IntInf.int) < c then Lt (Cn ((0 : IntInf.int), c, r))
-               else Gt (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (Le a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then Le r
-        else (if (0 : IntInf.int) < c then Le (Cn ((0 : IntInf.int), c, r))
-               else Ge (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (Gt a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then Gt r
-        else (if (0 : IntInf.int) < c then Gt (Cn ((0 : IntInf.int), c, r))
-               else Lt (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (Ge a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then Ge r
-        else (if (0 : IntInf.int) < c then Ge (Cn ((0 : IntInf.int), c, r))
-               else Le (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (Eq a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then Eq r
-        else (if (0 : IntInf.int) < c then Eq (Cn ((0 : IntInf.int), c, r))
-               else Eq (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (NEq a) =
-    let
-      val (c, r) = zsplit0 a;
-    in
-      (if c = (0 : IntInf.int) then NEq r
-        else (if (0 : IntInf.int) < c then NEq (Cn ((0 : IntInf.int), c, r))
-               else NEq (Cn ((0 : IntInf.int), ~ c, Neg r))))
-    end
-  | zlfm (Dvd (i, a)) =
-    (if i = (0 : IntInf.int) then zlfm (Eq a)
-      else let
-             val (c, r) = zsplit0 a;
-           in
-             (if c = (0 : IntInf.int) then Dvd (abs_int i, r)
-               else (if (0 : IntInf.int) < c
-                      then Dvd (abs_int i, Cn ((0 : IntInf.int), c, r))
-                      else Dvd (abs_int i, Cn ((0 : IntInf.int), ~ c, Neg r))))
-           end)
-  | zlfm (NDvd (i, a)) =
-    (if i = (0 : IntInf.int) then zlfm (NEq a)
-      else let
-             val (c, r) = zsplit0 a;
-           in
-             (if c = (0 : IntInf.int) then NDvd (abs_int i, r)
-               else (if (0 : IntInf.int) < c
-                      then NDvd (abs_int i, Cn ((0 : IntInf.int), c, r))
-                      else NDvd (abs_int i, Cn ((0 : IntInf.int), ~ c, Neg r))))
-           end)
-  | zlfm (Not (And (p, q))) = Or (zlfm (Not p), zlfm (Not q))
-  | zlfm (Not (Or (p, q))) = And (zlfm (Not p), zlfm (Not q))
-  | zlfm (Not (Imp (p, q))) = And (zlfm p, zlfm (Not q))
-  | zlfm (Not (Iff (p, q))) =
-    Or (And (zlfm p, zlfm (Not q)), And (zlfm (Not p), zlfm q))
-  | zlfm (Not (Not p)) = zlfm p
-  | zlfm (Not T) = F
-  | zlfm (Not F) = T
-  | zlfm (Not (Lt a)) = zlfm (Ge a)
-  | zlfm (Not (Le a)) = zlfm (Gt a)
-  | zlfm (Not (Gt a)) = zlfm (Le a)
-  | zlfm (Not (Ge a)) = zlfm (Lt a)
-  | zlfm (Not (Eq a)) = zlfm (NEq a)
-  | zlfm (Not (NEq a)) = zlfm (Eq a)
-  | zlfm (Not (Dvd (i, a))) = zlfm (NDvd (i, a))
-  | zlfm (Not (NDvd (i, a))) = zlfm (Dvd (i, a))
-  | zlfm (Not (Closed p)) = NClosed p
-  | zlfm (Not (NClosed p)) = Closed p
-  | zlfm T = T
-  | zlfm F = F
-  | zlfm (Not (E ci)) = Not (E ci)
-  | zlfm (Not (A cj)) = Not (A cj)
-  | zlfm (E ao) = E ao
-  | zlfm (A ap) = A ap
-  | zlfm (Closed aq) = Closed aq
-  | zlfm (NClosed ar) = NClosed ar;
-
-fun alpha (And (p, q)) = alpha p @ alpha q
-  | alpha (Or (p, q)) = alpha p @ alpha q
-  | alpha T = []
-  | alpha F = []
-  | alpha (Lt (C bo)) = []
-  | alpha (Lt (Bound bp)) = []
-  | alpha (Lt (Neg bt)) = []
-  | alpha (Lt (Add (bu, bv))) = []
-  | alpha (Lt (Sub (bw, bx))) = []
-  | alpha (Lt (Mul (by, bz))) = []
-  | alpha (Le (C co)) = []
-  | alpha (Le (Bound cp)) = []
-  | alpha (Le (Neg ct)) = []
-  | alpha (Le (Add (cu, cv))) = []
-  | alpha (Le (Sub (cw, cx))) = []
-  | alpha (Le (Mul (cy, cz))) = []
-  | alpha (Gt (C doa)) = []
-  | alpha (Gt (Bound dp)) = []
-  | alpha (Gt (Neg dt)) = []
-  | alpha (Gt (Add (du, dv))) = []
-  | alpha (Gt (Sub (dw, dx))) = []
-  | alpha (Gt (Mul (dy, dz))) = []
-  | alpha (Ge (C eo)) = []
-  | alpha (Ge (Bound ep)) = []
-  | alpha (Ge (Neg et)) = []
-  | alpha (Ge (Add (eu, ev))) = []
-  | alpha (Ge (Sub (ew, ex))) = []
-  | alpha (Ge (Mul (ey, ez))) = []
-  | alpha (Eq (C fo)) = []
-  | alpha (Eq (Bound fp)) = []
-  | alpha (Eq (Neg ft)) = []
-  | alpha (Eq (Add (fu, fv))) = []
-  | alpha (Eq (Sub (fw, fx))) = []
-  | alpha (Eq (Mul (fy, fz))) = []
-  | alpha (NEq (C go)) = []
-  | alpha (NEq (Bound gp)) = []
-  | alpha (NEq (Neg gt)) = []
-  | alpha (NEq (Add (gu, gv))) = []
-  | alpha (NEq (Sub (gw, gx))) = []
-  | alpha (NEq (Mul (gy, gz))) = []
-  | alpha (Dvd (aa, ab)) = []
-  | alpha (NDvd (ac, ad)) = []
-  | alpha (Not ae) = []
-  | alpha (Imp (aj, ak)) = []
-  | alpha (Iff (al, am)) = []
-  | alpha (E an) = []
-  | alpha (A ao) = []
-  | alpha (Closed ap) = []
-  | alpha (NClosed aq) = []
-  | alpha (Lt (Cn (cm, c, e))) = (if cm = (0 : IntInf.int) then [e] else [])
-  | alpha (Le (Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int) then [Add (C (~1 : IntInf.int), e)] else [])
-  | alpha (Gt (Cn (em, c, e))) = (if em = (0 : IntInf.int) then [] else [])
-  | alpha (Ge (Cn (fm, c, e))) = (if fm = (0 : IntInf.int) then [] else [])
-  | alpha (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int) then [Add (C (~1 : IntInf.int), e)] else [])
-  | alpha (NEq (Cn (hm, c, e))) = (if hm = (0 : IntInf.int) then [e] else []);
-
-fun delta (And (p, q)) = lcm_int (delta p) (delta q)
-  | delta (Or (p, q)) = lcm_int (delta p) (delta q)
-  | delta T = (1 : IntInf.int)
-  | delta F = (1 : IntInf.int)
-  | delta (Lt u) = (1 : IntInf.int)
-  | delta (Le v) = (1 : IntInf.int)
-  | delta (Gt w) = (1 : IntInf.int)
-  | delta (Ge x) = (1 : IntInf.int)
-  | delta (Eq y) = (1 : IntInf.int)
-  | delta (NEq z) = (1 : IntInf.int)
-  | delta (Dvd (aa, C bo)) = (1 : IntInf.int)
-  | delta (Dvd (aa, Bound bp)) = (1 : IntInf.int)
-  | delta (Dvd (aa, Neg bt)) = (1 : IntInf.int)
-  | delta (Dvd (aa, Add (bu, bv))) = (1 : IntInf.int)
-  | delta (Dvd (aa, Sub (bw, bx))) = (1 : IntInf.int)
-  | delta (Dvd (aa, Mul (by, bz))) = (1 : IntInf.int)
-  | delta (NDvd (ac, C co)) = (1 : IntInf.int)
-  | delta (NDvd (ac, Bound cp)) = (1 : IntInf.int)
-  | delta (NDvd (ac, Neg ct)) = (1 : IntInf.int)
-  | delta (NDvd (ac, Add (cu, cv))) = (1 : IntInf.int)
-  | delta (NDvd (ac, Sub (cw, cx))) = (1 : IntInf.int)
-  | delta (NDvd (ac, Mul (cy, cz))) = (1 : IntInf.int)
-  | delta (Not ae) = (1 : IntInf.int)
-  | delta (Imp (aj, ak)) = (1 : IntInf.int)
-  | delta (Iff (al, am)) = (1 : IntInf.int)
-  | delta (E an) = (1 : IntInf.int)
-  | delta (A ao) = (1 : IntInf.int)
-  | delta (Closed ap) = (1 : IntInf.int)
-  | delta (NClosed aq) = (1 : IntInf.int)
-  | delta (Dvd (i, Cn (cm, c, e))) =
-    (if cm = (0 : IntInf.int) then i else (1 : IntInf.int))
-  | delta (NDvd (i, Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int) then i else (1 : IntInf.int));
-
-fun member A_ [] y = false
-  | member A_ (x :: xs) y = eq A_ x y orelse member A_ xs y;
-
-fun remdups A_ [] = []
-  | remdups A_ (x :: xs) =
-    (if member A_ xs x then remdups A_ xs else x :: remdups A_ xs);
-
-fun a_beta (And (p, q)) = (fn k => And (a_beta p k, a_beta q k))
-  | a_beta (Or (p, q)) = (fn k => Or (a_beta p k, a_beta q k))
-  | a_beta T = (fn _ => T)
-  | a_beta F = (fn _ => F)
-  | a_beta (Lt (C bo)) = (fn _ => Lt (C bo))
-  | a_beta (Lt (Bound bp)) = (fn _ => Lt (Bound bp))
-  | a_beta (Lt (Neg bt)) = (fn _ => Lt (Neg bt))
-  | a_beta (Lt (Add (bu, bv))) = (fn _ => Lt (Add (bu, bv)))
-  | a_beta (Lt (Sub (bw, bx))) = (fn _ => Lt (Sub (bw, bx)))
-  | a_beta (Lt (Mul (by, bz))) = (fn _ => Lt (Mul (by, bz)))
-  | a_beta (Le (C co)) = (fn _ => Le (C co))
-  | a_beta (Le (Bound cp)) = (fn _ => Le (Bound cp))
-  | a_beta (Le (Neg ct)) = (fn _ => Le (Neg ct))
-  | a_beta (Le (Add (cu, cv))) = (fn _ => Le (Add (cu, cv)))
-  | a_beta (Le (Sub (cw, cx))) = (fn _ => Le (Sub (cw, cx)))
-  | a_beta (Le (Mul (cy, cz))) = (fn _ => Le (Mul (cy, cz)))
-  | a_beta (Gt (C doa)) = (fn _ => Gt (C doa))
-  | a_beta (Gt (Bound dp)) = (fn _ => Gt (Bound dp))
-  | a_beta (Gt (Neg dt)) = (fn _ => Gt (Neg dt))
-  | a_beta (Gt (Add (du, dv))) = (fn _ => Gt (Add (du, dv)))
-  | a_beta (Gt (Sub (dw, dx))) = (fn _ => Gt (Sub (dw, dx)))
-  | a_beta (Gt (Mul (dy, dz))) = (fn _ => Gt (Mul (dy, dz)))
-  | a_beta (Ge (C eo)) = (fn _ => Ge (C eo))
-  | a_beta (Ge (Bound ep)) = (fn _ => Ge (Bound ep))
-  | a_beta (Ge (Neg et)) = (fn _ => Ge (Neg et))
-  | a_beta (Ge (Add (eu, ev))) = (fn _ => Ge (Add (eu, ev)))
-  | a_beta (Ge (Sub (ew, ex))) = (fn _ => Ge (Sub (ew, ex)))
-  | a_beta (Ge (Mul (ey, ez))) = (fn _ => Ge (Mul (ey, ez)))
-  | a_beta (Eq (C fo)) = (fn _ => Eq (C fo))
-  | a_beta (Eq (Bound fp)) = (fn _ => Eq (Bound fp))
-  | a_beta (Eq (Neg ft)) = (fn _ => Eq (Neg ft))
-  | a_beta (Eq (Add (fu, fv))) = (fn _ => Eq (Add (fu, fv)))
-  | a_beta (Eq (Sub (fw, fx))) = (fn _ => Eq (Sub (fw, fx)))
-  | a_beta (Eq (Mul (fy, fz))) = (fn _ => Eq (Mul (fy, fz)))
-  | a_beta (NEq (C go)) = (fn _ => NEq (C go))
-  | a_beta (NEq (Bound gp)) = (fn _ => NEq (Bound gp))
-  | a_beta (NEq (Neg gt)) = (fn _ => NEq (Neg gt))
-  | a_beta (NEq (Add (gu, gv))) = (fn _ => NEq (Add (gu, gv)))
-  | a_beta (NEq (Sub (gw, gx))) = (fn _ => NEq (Sub (gw, gx)))
-  | a_beta (NEq (Mul (gy, gz))) = (fn _ => NEq (Mul (gy, gz)))
-  | a_beta (Dvd (aa, C ho)) = (fn _ => Dvd (aa, C ho))
-  | a_beta (Dvd (aa, Bound hp)) = (fn _ => Dvd (aa, Bound hp))
-  | a_beta (Dvd (aa, Neg ht)) = (fn _ => Dvd (aa, Neg ht))
-  | a_beta (Dvd (aa, Add (hu, hv))) = (fn _ => Dvd (aa, Add (hu, hv)))
-  | a_beta (Dvd (aa, Sub (hw, hx))) = (fn _ => Dvd (aa, Sub (hw, hx)))
-  | a_beta (Dvd (aa, Mul (hy, hz))) = (fn _ => Dvd (aa, Mul (hy, hz)))
-  | a_beta (NDvd (ac, C io)) = (fn _ => NDvd (ac, C io))
-  | a_beta (NDvd (ac, Bound ip)) = (fn _ => NDvd (ac, Bound ip))
-  | a_beta (NDvd (ac, Neg it)) = (fn _ => NDvd (ac, Neg it))
-  | a_beta (NDvd (ac, Add (iu, iv))) = (fn _ => NDvd (ac, Add (iu, iv)))
-  | a_beta (NDvd (ac, Sub (iw, ix))) = (fn _ => NDvd (ac, Sub (iw, ix)))
-  | a_beta (NDvd (ac, Mul (iy, iz))) = (fn _ => NDvd (ac, Mul (iy, iz)))
-  | a_beta (Not ae) = (fn _ => Not ae)
-  | a_beta (Imp (aj, ak)) = (fn _ => Imp (aj, ak))
-  | a_beta (Iff (al, am)) = (fn _ => Iff (al, am))
-  | a_beta (E an) = (fn _ => E an)
-  | a_beta (A ao) = (fn _ => A ao)
-  | a_beta (Closed ap) = (fn _ => Closed ap)
-  | a_beta (NClosed aq) = (fn _ => NClosed aq)
-  | a_beta (Lt (Cn (cm, c, e))) =
-    (if cm = (0 : IntInf.int)
-      then (fn k =>
-             Lt (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                      Mul (div_inta k c, e))))
-      else (fn _ => Lt (Cn (suc (minus_nat cm (1 : IntInf.int)), c, e))))
-  | a_beta (Le (Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int)
-      then (fn k =>
-             Le (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                      Mul (div_inta k c, e))))
-      else (fn _ => Le (Cn (suc (minus_nat dm (1 : IntInf.int)), c, e))))
-  | a_beta (Gt (Cn (em, c, e))) =
-    (if em = (0 : IntInf.int)
-      then (fn k =>
-             Gt (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                      Mul (div_inta k c, e))))
-      else (fn _ => Gt (Cn (suc (minus_nat em (1 : IntInf.int)), c, e))))
-  | a_beta (Ge (Cn (fm, c, e))) =
-    (if fm = (0 : IntInf.int)
-      then (fn k =>
-             Ge (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                      Mul (div_inta k c, e))))
-      else (fn _ => Ge (Cn (suc (minus_nat fm (1 : IntInf.int)), c, e))))
-  | a_beta (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int)
-      then (fn k =>
-             Eq (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                      Mul (div_inta k c, e))))
-      else (fn _ => Eq (Cn (suc (minus_nat gm (1 : IntInf.int)), c, e))))
-  | a_beta (NEq (Cn (hm, c, e))) =
-    (if hm = (0 : IntInf.int)
-      then (fn k =>
-             NEq (Cn ((0 : IntInf.int), (1 : IntInf.int),
-                       Mul (div_inta k c, e))))
-      else (fn _ => NEq (Cn (suc (minus_nat hm (1 : IntInf.int)), c, e))))
-  | a_beta (Dvd (i, Cn (im, c, e))) =
-    (if im = (0 : IntInf.int)
-      then (fn k =>
-             Dvd (div_inta k c * i,
-                   Cn ((0 : IntInf.int), (1 : IntInf.int),
-                        Mul (div_inta k c, e))))
-      else (fn _ => Dvd (i, Cn (suc (minus_nat im (1 : IntInf.int)), c, e))))
-  | a_beta (NDvd (i, Cn (jm, c, e))) =
-    (if jm = (0 : IntInf.int)
-      then (fn k =>
-             NDvd (div_inta k c * i,
-                    Cn ((0 : IntInf.int), (1 : IntInf.int),
-                         Mul (div_inta k c, e))))
-      else (fn _ => NDvd (i, Cn (suc (minus_nat jm (1 : IntInf.int)), c, e))));
-
-fun mirror (And (p, q)) = And (mirror p, mirror q)
-  | mirror (Or (p, q)) = Or (mirror p, mirror q)
-  | mirror T = T
-  | mirror F = F
-  | mirror (Lt (C bo)) = Lt (C bo)
-  | mirror (Lt (Bound bp)) = Lt (Bound bp)
-  | mirror (Lt (Neg bt)) = Lt (Neg bt)
-  | mirror (Lt (Add (bu, bv))) = Lt (Add (bu, bv))
-  | mirror (Lt (Sub (bw, bx))) = Lt (Sub (bw, bx))
-  | mirror (Lt (Mul (by, bz))) = Lt (Mul (by, bz))
-  | mirror (Le (C co)) = Le (C co)
-  | mirror (Le (Bound cp)) = Le (Bound cp)
-  | mirror (Le (Neg ct)) = Le (Neg ct)
-  | mirror (Le (Add (cu, cv))) = Le (Add (cu, cv))
-  | mirror (Le (Sub (cw, cx))) = Le (Sub (cw, cx))
-  | mirror (Le (Mul (cy, cz))) = Le (Mul (cy, cz))
-  | mirror (Gt (C doa)) = Gt (C doa)
-  | mirror (Gt (Bound dp)) = Gt (Bound dp)
-  | mirror (Gt (Neg dt)) = Gt (Neg dt)
-  | mirror (Gt (Add (du, dv))) = Gt (Add (du, dv))
-  | mirror (Gt (Sub (dw, dx))) = Gt (Sub (dw, dx))
-  | mirror (Gt (Mul (dy, dz))) = Gt (Mul (dy, dz))
-  | mirror (Ge (C eo)) = Ge (C eo)
-  | mirror (Ge (Bound ep)) = Ge (Bound ep)
-  | mirror (Ge (Neg et)) = Ge (Neg et)
-  | mirror (Ge (Add (eu, ev))) = Ge (Add (eu, ev))
-  | mirror (Ge (Sub (ew, ex))) = Ge (Sub (ew, ex))
-  | mirror (Ge (Mul (ey, ez))) = Ge (Mul (ey, ez))
-  | mirror (Eq (C fo)) = Eq (C fo)
-  | mirror (Eq (Bound fp)) = Eq (Bound fp)
-  | mirror (Eq (Neg ft)) = Eq (Neg ft)
-  | mirror (Eq (Add (fu, fv))) = Eq (Add (fu, fv))
-  | mirror (Eq (Sub (fw, fx))) = Eq (Sub (fw, fx))
-  | mirror (Eq (Mul (fy, fz))) = Eq (Mul (fy, fz))
-  | mirror (NEq (C go)) = NEq (C go)
-  | mirror (NEq (Bound gp)) = NEq (Bound gp)
-  | mirror (NEq (Neg gt)) = NEq (Neg gt)
-  | mirror (NEq (Add (gu, gv))) = NEq (Add (gu, gv))
-  | mirror (NEq (Sub (gw, gx))) = NEq (Sub (gw, gx))
-  | mirror (NEq (Mul (gy, gz))) = NEq (Mul (gy, gz))
-  | mirror (Dvd (aa, C ho)) = Dvd (aa, C ho)
-  | mirror (Dvd (aa, Bound hp)) = Dvd (aa, Bound hp)
-  | mirror (Dvd (aa, Neg ht)) = Dvd (aa, Neg ht)
-  | mirror (Dvd (aa, Add (hu, hv))) = Dvd (aa, Add (hu, hv))
-  | mirror (Dvd (aa, Sub (hw, hx))) = Dvd (aa, Sub (hw, hx))
-  | mirror (Dvd (aa, Mul (hy, hz))) = Dvd (aa, Mul (hy, hz))
-  | mirror (NDvd (ac, C io)) = NDvd (ac, C io)
-  | mirror (NDvd (ac, Bound ip)) = NDvd (ac, Bound ip)
-  | mirror (NDvd (ac, Neg it)) = NDvd (ac, Neg it)
-  | mirror (NDvd (ac, Add (iu, iv))) = NDvd (ac, Add (iu, iv))
-  | mirror (NDvd (ac, Sub (iw, ix))) = NDvd (ac, Sub (iw, ix))
-  | mirror (NDvd (ac, Mul (iy, iz))) = NDvd (ac, Mul (iy, iz))
-  | mirror (Not ae) = Not ae
-  | mirror (Imp (aj, ak)) = Imp (aj, ak)
-  | mirror (Iff (al, am)) = Iff (al, am)
-  | mirror (E an) = E an
-  | mirror (A ao) = A ao
-  | mirror (Closed ap) = Closed ap
-  | mirror (NClosed aq) = NClosed aq
-  | mirror (Lt (Cn (cm, c, e))) =
-    (if cm = (0 : IntInf.int) then Gt (Cn ((0 : IntInf.int), c, Neg e))
-      else Lt (Cn (suc (minus_nat cm (1 : IntInf.int)), c, e)))
-  | mirror (Le (Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int) then Ge (Cn ((0 : IntInf.int), c, Neg e))
-      else Le (Cn (suc (minus_nat dm (1 : IntInf.int)), c, e)))
-  | mirror (Gt (Cn (em, c, e))) =
-    (if em = (0 : IntInf.int) then Lt (Cn ((0 : IntInf.int), c, Neg e))
-      else Gt (Cn (suc (minus_nat em (1 : IntInf.int)), c, e)))
-  | mirror (Ge (Cn (fm, c, e))) =
-    (if fm = (0 : IntInf.int) then Le (Cn ((0 : IntInf.int), c, Neg e))
-      else Ge (Cn (suc (minus_nat fm (1 : IntInf.int)), c, e)))
-  | mirror (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int) then Eq (Cn ((0 : IntInf.int), c, Neg e))
-      else Eq (Cn (suc (minus_nat gm (1 : IntInf.int)), c, e)))
-  | mirror (NEq (Cn (hm, c, e))) =
-    (if hm = (0 : IntInf.int) then NEq (Cn ((0 : IntInf.int), c, Neg e))
-      else NEq (Cn (suc (minus_nat hm (1 : IntInf.int)), c, e)))
-  | mirror (Dvd (i, Cn (im, c, e))) =
-    (if im = (0 : IntInf.int) then Dvd (i, Cn ((0 : IntInf.int), c, Neg e))
-      else Dvd (i, Cn (suc (minus_nat im (1 : IntInf.int)), c, e)))
-  | mirror (NDvd (i, Cn (jm, c, e))) =
-    (if jm = (0 : IntInf.int) then NDvd (i, Cn ((0 : IntInf.int), c, Neg e))
-      else NDvd (i, Cn (suc (minus_nat jm (1 : IntInf.int)), c, e)));
-
-fun size_list [] = (0 : IntInf.int)
-  | size_list (a :: lista) = size_list lista + suc (0 : IntInf.int);
-
-val equal_num = {equal = equal_numa} : num equal;
-
-fun unita p =
-  let
-    val pa = zlfm p;
-    val l = zeta pa;
-    val q =
-      And (Dvd (l, Cn ((0 : IntInf.int), (1 : IntInf.int), C (0 : IntInf.int))),
-            a_beta pa l);
-    val d = delta q;
-    val b = remdups equal_num (map simpnum (beta q));
-    val a = remdups equal_num (map simpnum (alpha q));
-  in
-    (if size_list b <= size_list a then (q, (b, d)) else (mirror q, (a, d)))
-  end;
-
-fun numsubst0 t (C c) = C c
-  | numsubst0 t (Bound n) = (if n = (0 : IntInf.int) then t else Bound n)
-  | numsubst0 t (Neg a) = Neg (numsubst0 t a)
-  | numsubst0 t (Add (a, b)) = Add (numsubst0 t a, numsubst0 t b)
-  | numsubst0 t (Sub (a, b)) = Sub (numsubst0 t a, numsubst0 t b)
-  | numsubst0 t (Mul (i, a)) = Mul (i, numsubst0 t a)
-  | numsubst0 t (Cn (v, i, a)) =
-    (if v = (0 : IntInf.int) then Add (Mul (i, t), numsubst0 t a)
-      else Cn (suc (minus_nat v (1 : IntInf.int)), i, numsubst0 t a));
-
-fun subst0 t T = T
-  | subst0 t F = F
-  | subst0 t (Lt a) = Lt (numsubst0 t a)
-  | subst0 t (Le a) = Le (numsubst0 t a)
-  | subst0 t (Gt a) = Gt (numsubst0 t a)
-  | subst0 t (Ge a) = Ge (numsubst0 t a)
-  | subst0 t (Eq a) = Eq (numsubst0 t a)
-  | subst0 t (NEq a) = NEq (numsubst0 t a)
-  | subst0 t (Dvd (i, a)) = Dvd (i, numsubst0 t a)
-  | subst0 t (NDvd (i, a)) = NDvd (i, numsubst0 t a)
-  | subst0 t (Not p) = Not (subst0 t p)
-  | subst0 t (And (p, q)) = And (subst0 t p, subst0 t q)
-  | subst0 t (Or (p, q)) = Or (subst0 t p, subst0 t q)
-  | subst0 t (Imp (p, q)) = Imp (subst0 t p, subst0 t q)
-  | subst0 t (Iff (p, q)) = Iff (subst0 t p, subst0 t q)
-  | subst0 t (Closed p) = Closed p
-  | subst0 t (NClosed p) = NClosed p;
-
-fun minusinf (And (p, q)) = And (minusinf p, minusinf q)
-  | minusinf (Or (p, q)) = Or (minusinf p, minusinf q)
-  | minusinf T = T
-  | minusinf F = F
-  | minusinf (Lt (C bo)) = Lt (C bo)
-  | minusinf (Lt (Bound bp)) = Lt (Bound bp)
-  | minusinf (Lt (Neg bt)) = Lt (Neg bt)
-  | minusinf (Lt (Add (bu, bv))) = Lt (Add (bu, bv))
-  | minusinf (Lt (Sub (bw, bx))) = Lt (Sub (bw, bx))
-  | minusinf (Lt (Mul (by, bz))) = Lt (Mul (by, bz))
-  | minusinf (Le (C co)) = Le (C co)
-  | minusinf (Le (Bound cp)) = Le (Bound cp)
-  | minusinf (Le (Neg ct)) = Le (Neg ct)
-  | minusinf (Le (Add (cu, cv))) = Le (Add (cu, cv))
-  | minusinf (Le (Sub (cw, cx))) = Le (Sub (cw, cx))
-  | minusinf (Le (Mul (cy, cz))) = Le (Mul (cy, cz))
-  | minusinf (Gt (C doa)) = Gt (C doa)
-  | minusinf (Gt (Bound dp)) = Gt (Bound dp)
-  | minusinf (Gt (Neg dt)) = Gt (Neg dt)
-  | minusinf (Gt (Add (du, dv))) = Gt (Add (du, dv))
-  | minusinf (Gt (Sub (dw, dx))) = Gt (Sub (dw, dx))
-  | minusinf (Gt (Mul (dy, dz))) = Gt (Mul (dy, dz))
-  | minusinf (Ge (C eo)) = Ge (C eo)
-  | minusinf (Ge (Bound ep)) = Ge (Bound ep)
-  | minusinf (Ge (Neg et)) = Ge (Neg et)
-  | minusinf (Ge (Add (eu, ev))) = Ge (Add (eu, ev))
-  | minusinf (Ge (Sub (ew, ex))) = Ge (Sub (ew, ex))
-  | minusinf (Ge (Mul (ey, ez))) = Ge (Mul (ey, ez))
-  | minusinf (Eq (C fo)) = Eq (C fo)
-  | minusinf (Eq (Bound fp)) = Eq (Bound fp)
-  | minusinf (Eq (Neg ft)) = Eq (Neg ft)
-  | minusinf (Eq (Add (fu, fv))) = Eq (Add (fu, fv))
-  | minusinf (Eq (Sub (fw, fx))) = Eq (Sub (fw, fx))
-  | minusinf (Eq (Mul (fy, fz))) = Eq (Mul (fy, fz))
-  | minusinf (NEq (C go)) = NEq (C go)
-  | minusinf (NEq (Bound gp)) = NEq (Bound gp)
-  | minusinf (NEq (Neg gt)) = NEq (Neg gt)
-  | minusinf (NEq (Add (gu, gv))) = NEq (Add (gu, gv))
-  | minusinf (NEq (Sub (gw, gx))) = NEq (Sub (gw, gx))
-  | minusinf (NEq (Mul (gy, gz))) = NEq (Mul (gy, gz))
-  | minusinf (Dvd (aa, ab)) = Dvd (aa, ab)
-  | minusinf (NDvd (ac, ad)) = NDvd (ac, ad)
-  | minusinf (Not ae) = Not ae
-  | minusinf (Imp (aj, ak)) = Imp (aj, ak)
-  | minusinf (Iff (al, am)) = Iff (al, am)
-  | minusinf (E an) = E an
-  | minusinf (A ao) = A ao
-  | minusinf (Closed ap) = Closed ap
-  | minusinf (NClosed aq) = NClosed aq
-  | minusinf (Lt (Cn (cm, c, e))) =
-    (if cm = (0 : IntInf.int) then T
-      else Lt (Cn (suc (minus_nat cm (1 : IntInf.int)), c, e)))
-  | minusinf (Le (Cn (dm, c, e))) =
-    (if dm = (0 : IntInf.int) then T
-      else Le (Cn (suc (minus_nat dm (1 : IntInf.int)), c, e)))
-  | minusinf (Gt (Cn (em, c, e))) =
-    (if em = (0 : IntInf.int) then F
-      else Gt (Cn (suc (minus_nat em (1 : IntInf.int)), c, e)))
-  | minusinf (Ge (Cn (fm, c, e))) =
-    (if fm = (0 : IntInf.int) then F
-      else Ge (Cn (suc (minus_nat fm (1 : IntInf.int)), c, e)))
-  | minusinf (Eq (Cn (gm, c, e))) =
-    (if gm = (0 : IntInf.int) then F
-      else Eq (Cn (suc (minus_nat gm (1 : IntInf.int)), c, e)))
-  | minusinf (NEq (Cn (hm, c, e))) =
-    (if hm = (0 : IntInf.int) then T
-      else NEq (Cn (suc (minus_nat hm (1 : IntInf.int)), c, e)));
-
-fun cooper p =
-  let
-    val (q, (b, d)) = unita p;
-    val js = uptoa (1 : IntInf.int) d;
-    val mq = simpfm (minusinf q);
-    val md = evaldjf (fn j => simpfm (subst0 (C j) mq)) js;
-  in
-    (if equal_fm md T then T
-      else let
-             val qd =
-               evaldjf (fn (ba, j) => simpfm (subst0 (Add (ba, C j)) q))
-                 (maps (fn ba => map (fn a => (ba, a)) js) b);
-           in
-             decr (disj md qd)
-           end)
-  end;
-
 fun pa p = qelim (prep p) cooper;
 
 end; (*struct Cooper_Procedure*)
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -10,11 +10,11 @@
     Proof.context -> (term * term list) list -> bool -> int list -> (bool * term list) option * Quickcheck.report option
   val compile_generator_exprs: Proof.context -> term list -> (int -> term list option) list
   val compile_validator_exprs: Proof.context -> term list -> (int -> bool) list
-  val put_counterexample: (unit -> int -> bool -> int -> (bool * term list) option)
+  val put_counterexample: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> (bool * term list) option)
     -> Proof.context -> Proof.context
-  val put_counterexample_batch: (unit -> (int -> term list option) list)
+  val put_counterexample_batch: (unit -> (Code_Numeral.natural -> term list option) list)
     -> Proof.context -> Proof.context
-  val put_validator_batch: (unit -> (int -> bool) list) -> Proof.context -> Proof.context
+  val put_validator_batch: (unit -> (Code_Numeral.natural -> bool) list) -> Proof.context -> Proof.context
   exception Counterexample of term list
   val smart_quantifier : bool Config.T
   val optimise_equality : bool Config.T
@@ -50,9 +50,9 @@
 
 fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => Code_Evaluation.term"});
 
-val size = @{term "i :: code_numeral"}
-val size_pred = @{term "(i :: code_numeral) - 1"}
-val size_ge_zero = @{term "(i :: code_numeral) > 0"}
+val size = @{term "i :: natural"}
+val size_pred = @{term "(i :: natural) - 1"}
+val size_ge_zero = @{term "(i :: natural) > 0"}
 
 fun mk_none_continuation (x, y) =
   let
@@ -81,13 +81,13 @@
 val bounded_forallN = "bounded_forall";
 
 fun fast_exhaustiveT T = (T --> @{typ unit})
-  --> @{typ code_numeral} --> @{typ unit}
+  --> @{typ natural} --> @{typ unit}
 
-fun exhaustiveT T = (T --> resultT) --> @{typ code_numeral} --> resultT
+fun exhaustiveT T = (T --> resultT) --> @{typ natural} --> resultT
 
-fun bounded_forallT T = (T --> @{typ bool}) --> @{typ code_numeral} --> @{typ bool}
+fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool}
 
-fun full_exhaustiveT T = (termifyT T --> resultT) --> @{typ code_numeral} --> resultT
+fun full_exhaustiveT T = (termifyT T --> resultT) --> @{typ natural} --> resultT
 
 fun check_allT T = (termifyT T --> resultT) --> resultT
 
@@ -323,7 +323,7 @@
     val frees = map Free (Term.add_frees t [])
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt'
-    val depth = Free (depth_name, @{typ code_numeral})
+    val depth = Free (depth_name, @{typ natural})
     fun return _ = @{term "throw_Counterexample :: term list => unit"} $
       (HOLogic.mk_list @{typ "term"}
         (map (fn t => HOLogic.mk_term_of (fastype_of t) t) (frees @ eval_terms)))
@@ -353,7 +353,7 @@
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name, genuine_only_name], _) =
       Variable.variant_fixes ["depth", "genuine_only"] ctxt'
-    val depth = Free (depth_name, @{typ code_numeral})
+    val depth = Free (depth_name, @{typ natural})
     val genuine_only = Free (genuine_only_name, @{typ bool}) 
     val return = mk_return (HOLogic.mk_list @{typ "term"}
         (map (fn t => HOLogic.mk_term_of (fastype_of t) t) frees @ map mk_safe_term eval_terms))
@@ -375,7 +375,7 @@
     val ([depth_name, genuine_only_name], ctxt'') =
       Variable.variant_fixes ["depth", "genuine_only"] ctxt'
     val (term_names, _) = Variable.variant_fixes (map (prefix "t_") names) ctxt''
-    val depth = Free (depth_name, @{typ code_numeral})
+    val depth = Free (depth_name, @{typ natural})
     val genuine_only = Free (genuine_only_name, @{typ bool})    
     val term_vars = map (fn n => Free (n, @{typ "unit => term"})) term_names
     fun lookup v = the (AList.lookup (op =) (names ~~ (frees ~~ term_vars)) v)
@@ -402,18 +402,18 @@
 fun mk_parametric_generator_expr mk_generator_expr =
   Quickcheck_Common.gen_mk_parametric_generator_expr 
     ((mk_generator_expr,
-      absdummy @{typ bool} (absdummy @{typ "code_numeral"} (Const (@{const_name "None"}, resultT)))),
-      @{typ bool} --> @{typ "code_numeral"} --> resultT)
+      absdummy @{typ bool} (absdummy @{typ natural} (Const (@{const_name "None"}, resultT)))),
+      @{typ bool} --> @{typ natural} --> resultT)
 
 fun mk_validator_expr ctxt t =
   let
-    fun bounded_forallT T = (T --> @{typ bool}) --> @{typ code_numeral} --> @{typ bool}
+    fun bounded_forallT T = (T --> @{typ bool}) --> @{typ natural} --> @{typ bool}
     val ctxt' = Variable.auto_fixes t ctxt
     val names = Term.add_free_names t []
     val frees = map Free (Term.add_frees t [])
     fun lookup v = the (AList.lookup (op =) (names ~~ frees) v)
     val ([depth_name], _) = Variable.variant_fixes ["depth"] ctxt'
-    val depth = Free (depth_name, @{typ code_numeral})
+    val depth = Free (depth_name, @{typ natural})
     fun mk_bounded_forall (Free (s, T)) t =
       Const (@{const_name "Quickcheck_Exhaustive.bounded_forall_class.bounded_forall"}, bounded_forallT T)
         $ lambda (Free (s, T)) t $ depth
@@ -443,7 +443,7 @@
 
 structure Counterexample = Proof_Data
 (
-  type T = unit -> int -> bool -> int -> (bool * term list) option
+  type T = unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> (bool * term list) option
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -451,7 +451,7 @@
 
 structure Counterexample_Batch = Proof_Data
 (
-  type T = unit -> (int -> term list option) list
+  type T = unit -> (Code_Numeral.natural -> term list option) list
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -459,7 +459,7 @@
 
 structure Validator_Batch = Proof_Data
 (
-  type T = unit -> (int -> bool) list
+  type T = unit -> (Code_Numeral.natural -> bool) list
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -468,7 +468,7 @@
 
 val target = "Quickcheck";
 
-fun compile_generator_expr ctxt ts =
+fun compile_generator_expr_raw ctxt ts =
   let
     val thy = Proof_Context.theory_of ctxt
     val mk_generator_expr = 
@@ -487,7 +487,14 @@
         Option.map (apsnd (map Quickcheck_Common.post_process_term)) else I))
   end;
 
-fun compile_generator_exprs ctxt ts =
+fun compile_generator_expr ctxt ts =
+  let
+    val compiled = compile_generator_expr_raw ctxt ts;
+  in fn genuine_only => fn [card, size] =>
+    compiled genuine_only [Code_Numeral.natural_of_integer card, Code_Numeral.natural_of_integer size]
+  end;
+
+fun compile_generator_exprs_raw ctxt ts =
   let
     val thy = Proof_Context.theory_of ctxt
     val ts' = map (fn t => mk_generator_expr ctxt (t, [])) ts;
@@ -495,22 +502,30 @@
       (Counterexample_Batch.get, put_counterexample_batch,
         "Exhaustive_Generators.put_counterexample_batch")
       thy (SOME target) (fn proc => map (fn g => g #> (Option.map o map) proc))
-      (HOLogic.mk_list @{typ "code_numeral => term list option"} ts') []
+      (HOLogic.mk_list @{typ "natural => term list option"} ts') []
   in
     map (fn compile => fn size => compile size
-      |> Option.map (map Quickcheck_Common.post_process_term)) compiles
+      |> (Option.map o map) Quickcheck_Common.post_process_term) compiles
   end;
 
-fun compile_validator_exprs ctxt ts =
+fun compile_generator_exprs ctxt ts =
+  compile_generator_exprs_raw ctxt ts
+  |> map (fn f => fn size => f (Code_Numeral.natural_of_integer size));
+
+fun compile_validator_exprs_raw ctxt ts =
   let
     val thy = Proof_Context.theory_of ctxt
     val ts' = map (mk_validator_expr ctxt) ts
   in
     Code_Runtime.dynamic_value_strict
       (Validator_Batch.get, put_validator_batch, "Exhaustive_Generators.put_validator_batch")
-      thy (SOME target) (K I) (HOLogic.mk_list @{typ "code_numeral => bool"} ts') []
+      thy (SOME target) (K I) (HOLogic.mk_list @{typ "natural => bool"} ts') []
   end;
 
+fun compile_validator_exprs ctxt ts =
+  compile_validator_exprs_raw ctxt ts
+  |> map (fn f => fn size => f (Code_Numeral.natural_of_integer size));
+
 fun size_matters_for thy Ts = not (forall (fn T => Sign.of_sort thy (T,  @{sort check_all})) Ts)
 
 val test_goals =
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -70,8 +70,8 @@
 fun mk_partial_term_of_eq thy ty (i, (c, (_, tys))) =
   let
     val frees = map Free (Name.invent_names Name.context "a" (map (K @{typ narrowing_term}) tys))
-    val narrowing_term = @{term "Quickcheck_Narrowing.Narrowing_constructor"} $ HOLogic.mk_number @{typ code_int} i
-      $ (HOLogic.mk_list @{typ narrowing_term} (rev frees))
+    val narrowing_term = @{term Quickcheck_Narrowing.Narrowing_constructor} $ HOLogic.mk_number @{typ integer} i
+      $ HOLogic.mk_list @{typ narrowing_term} (rev frees)
     val rhs = fold (fn u => fn t => @{term "Code_Evaluation.App"} $ t $ u)
         (map mk_partial_term_of (frees ~~ tys))
         (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty))
@@ -124,7 +124,7 @@
 val narrowingN = "narrowing";
 
 fun narrowingT T =
-  @{typ Quickcheck_Narrowing.code_int} --> Type (@{type_name Quickcheck_Narrowing.narrowing_cons}, [T])
+  @{typ integer} --> Type (@{type_name Quickcheck_Narrowing.narrowing_cons}, [T])
 
 fun mk_cons c T = Const (@{const_name Quickcheck_Narrowing.cons}, T --> narrowingT T) $ Const (c, T)
 
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -429,10 +429,10 @@
   let
     val if_t = Const (@{const_name "If"}, @{typ bool} --> T --> T --> T)
     fun mk_if (index, (t, eval_terms)) else_t = if_t $
-        (HOLogic.eq_const @{typ code_numeral} $ Bound 0 $ HOLogic.mk_number @{typ code_numeral} index) $
+        (HOLogic.eq_const @{typ natural} $ Bound 0 $ HOLogic.mk_number @{typ natural} index) $
         (mk_generator_expr ctxt (t, eval_terms)) $ else_t
   in
-    absdummy @{typ "code_numeral"} (fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds)
+    absdummy @{typ natural} (fold_rev mk_if (1 upto (length ts) ~~ ts) out_of_bounds)
   end
 
 (** post-processing of function terms **)
--- a/src/HOL/Tools/Quickcheck/random_generators.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -12,9 +12,9 @@
     -> seed -> (('a -> 'b) * (unit -> term)) * seed
   val compile_generator_expr:
     Proof.context -> (term * term list) list -> bool -> int list -> (bool * term list) option * Quickcheck.report option
-  val put_counterexample: (unit -> int -> bool -> int -> seed -> (bool * term list) option * seed)
+  val put_counterexample: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> (bool * term list) option * seed)
     -> Proof.context -> Proof.context
-  val put_counterexample_report: (unit -> int -> bool -> int -> seed -> ((bool * term list) option * (bool list * bool)) * seed)
+  val put_counterexample_report: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed)
     -> Proof.context -> Proof.context
   val instantiate_random_datatype : Datatype_Aux.config -> Datatype_Aux.descr ->
     (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
@@ -27,9 +27,9 @@
 (** abstract syntax **)
 
 fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"})
-val size = @{term "i::code_numeral"};
-val size_pred = @{term "(i::code_numeral) - 1"};
-val size' = @{term "j::code_numeral"};
+val size = @{term "i::natural"};
+val size_pred = @{term "(i::natural) - 1"};
+val size' = @{term "j::natural"};
 val seed = @{term "s::Random.seed"};
 
 val resultT =  @{typ "(bool * term list) option"};
@@ -76,8 +76,8 @@
 val pt_rhs = eq |> Thm.dest_arg |> Thm.dest_fun;
 val aT = pt_random_aux |> Thm.ctyp_of_term |> dest_ctyp_nth 1;
 
-val rew_thms = map mk_meta_eq [@{thm code_numeral_zero_minus_one},
-  @{thm Suc_code_numeral_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
+val rew_thms = map mk_meta_eq [@{thm natural_zero_minus_one},
+  @{thm Suc_natural_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
 val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms;
 val rew_ss = HOL_ss addsimps rew_thms;
 
@@ -94,7 +94,7 @@
     val inst = Thm.instantiate_cterm ([(aT, icT)], []);
     fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t);
     val t_rhs = lambda t_k proto_t_rhs;
-    val eqs0 = [subst_v @{term "0::code_numeral"} eq,
+    val eqs0 = [subst_v @{term "0::natural"} eq,
       subst_v (@{const Code_Numeral.Suc} $ t_k) eq];
     val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0;
     val ((_, (_, eqs2)), lthy') = Primrec.add_primrec_simple
@@ -189,7 +189,7 @@
     val rTs = Ts @ Us;
     fun random_resultT T = @{typ Random.seed}
       --> HOLogic.mk_prodT (termifyT T,@{typ Random.seed});
-    fun sizeT T = @{typ code_numeral} --> @{typ code_numeral} --> T;
+    fun sizeT T = @{typ natural} --> @{typ natural} --> T;
     val random_auxT = sizeT o random_resultT;
     val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT))
       random_auxsN rTs;
@@ -222,9 +222,9 @@
             tc @{typ Random.seed} (SOME T, @{typ Random.seed});
         val tk = if is_rec
           then if k = 0 then size
-            else @{term "Quickcheck_Random.beyond :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
-             $ HOLogic.mk_number @{typ code_numeral} k $ size
-          else @{term "1::code_numeral"}
+            else @{term "Quickcheck_Random.beyond :: natural \<Rightarrow> natural \<Rightarrow> natural"}
+             $ HOLogic.mk_number @{typ natural} k $ size
+          else @{term "1::natural"}
       in (is_rec, HOLogic.mk_prod (tk, t)) end;
     fun sort_rec xs =
       map_filter (fn (true, t) => SOME t | _ =>  NONE) xs
@@ -235,7 +235,7 @@
     fun mk_select (rT, xs) =
       mk_const @{const_name Quickcheck_Random.collapse} [@{typ Random.seed}, termifyT rT]
       $ (mk_const @{const_name Random.select_weight} [random_resultT rT]
-        $ HOLogic.mk_list (HOLogic.mk_prodT (@{typ code_numeral}, random_resultT rT)) xs)
+        $ HOLogic.mk_list (HOLogic.mk_prodT (@{typ natural}, random_resultT rT)) xs)
           $ seed;
     val auxs_lhss = map (fn t => t $ size $ size' $ seed) random_auxs;
     val auxs_rhss = map mk_select gen_exprss;
@@ -247,8 +247,8 @@
     val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
     fun mk_size_arg k = case Datatype_Aux.find_shortest_path descr k
      of SOME (_, l) => if l = 0 then size
-          else @{term "max :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
-            $ HOLogic.mk_number @{typ code_numeral} l $ size
+          else @{term "max :: natural \<Rightarrow> natural \<Rightarrow> natural"}
+            $ HOLogic.mk_number @{typ natural} l $ size
       | NONE => size;
     val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq
       (mk_random_aux_eqs thy descr vs (names, auxnames) (Ts, Us));
@@ -272,7 +272,7 @@
 
 structure Counterexample = Proof_Data
 (
-  type T = unit -> int -> bool -> int -> int * int -> (bool * term list) option * (int * int)
+  type T = unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> (bool * term list) option * seed
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample"
 );
@@ -280,7 +280,7 @@
 
 structure Counterexample_Report = Proof_Data
 (
-  type T = unit -> int -> bool -> int -> seed -> ((bool * term list) option * (bool list * bool)) * seed
+  type T = unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed
   (* FIXME avoid user error with non-user text *)
   fun init _ () = error "Counterexample_Report"
 );
@@ -318,7 +318,7 @@
       (Sign.mk_const thy (@{const_name Quickcheck_Random.random}, [T]) $ Bound i);
   in
     lambda genuine_only
-      (Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check true)))
+      (Abs ("n", @{typ natural}, fold_rev mk_bindclause bounds (return $ check true)))
   end;
 
 fun mk_reporting_generator_expr ctxt (t, _) =
@@ -364,21 +364,21 @@
       (Sign.mk_const thy (@{const_name Quickcheck_Random.random}, [T]) $ Bound i);
   in
     lambda genuine_only
-      (Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check true)))
+      (Abs ("n", @{typ natural}, fold_rev mk_bindclause bounds (return $ check true)))
   end
 
 val mk_parametric_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr 
   ((mk_generator_expr, 
-    absdummy @{typ bool} (absdummy @{typ code_numeral}
+    absdummy @{typ bool} (absdummy @{typ natural}
       @{term "Pair None :: Random.seed => (bool * term list) option * Random.seed"})),
-    @{typ "bool => code_numeral => Random.seed => (bool * term list) option * Random.seed"})
+    @{typ "bool => natural => Random.seed => (bool * term list) option * Random.seed"})
 
 val mk_parametric_reporting_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr 
   ((mk_reporting_generator_expr,
-    absdummy @{typ bool} (absdummy @{typ code_numeral}
+    absdummy @{typ bool} (absdummy @{typ natural}
       @{term "Pair (None, ([], False)) :: Random.seed =>
         ((bool * term list) option * (bool list * bool)) * Random.seed"})),
-    @{typ "bool => code_numeral => Random.seed => ((bool * term list) option * (bool list * bool)) * Random.seed"})
+    @{typ "bool => natural => Random.seed => ((bool * term list) option * (bool list * bool)) * Random.seed"})
     
     
 (* single quickcheck report *)
@@ -402,7 +402,7 @@
 val empty_report = Quickcheck.Report { iterations = 0, raised_match_errors = 0,
   satisfied_assms = [], positive_concl_tests = 0 }
     
-fun compile_generator_expr ctxt ts =
+fun compile_generator_expr_raw ctxt ts =
   let
     val thy = Proof_Context.theory_of ctxt
     val iterations = Config.get ctxt Quickcheck.iterations
@@ -449,6 +449,13 @@
       end
   end;
 
+fun compile_generator_expr ctxt ts =
+  let
+    val compiled = compile_generator_expr_raw ctxt ts
+  in fn genuine_only => fn [card, size] =>
+    compiled genuine_only [Code_Numeral.natural_of_integer card, Code_Numeral.natural_of_integer size]
+  end;
+
 val size_types = [@{type_name Enum.finite_1}, @{type_name Enum.finite_2},
   @{type_name Enum.finite_3}, @{type_name Enum.finite_4}, @{type_name Enum.finite_5}];
 
--- a/src/HOL/Tools/hologic.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/hologic.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -91,7 +91,6 @@
   val dest_nat: term -> int
   val class_size: string
   val size_const: typ -> term
-  val code_numeralT: typ
   val intT: typ
   val one_const: term
   val bit0_const: term
@@ -105,6 +104,8 @@
   val add_numerals: term -> (term * typ) list -> (term * typ) list
   val mk_number: typ -> int -> term
   val dest_number: term -> typ * int
+  val code_integerT: typ
+  val code_naturalT: typ
   val realT: typ
   val nibbleT: typ
   val mk_nibble: int -> term
@@ -487,11 +488,6 @@
 fun size_const T = Const ("Nat.size_class.size", T --> natT);
 
 
-(* code numeral *)
-
-val code_numeralT = Type ("Code_Numeral.code_numeral", []);
-
-
 (* binary numerals and int *)
 
 val numT = Type ("Num.num", []);
@@ -543,6 +539,13 @@
   | dest_number t = raise TERM ("dest_number", [t]);
 
 
+(* code target numerals *)
+
+val code_integerT = Type ("Code_Numeral.integer", []);
+
+val code_naturalT = Type ("Code_Numeral.natural", []);
+
+
 (* real *)
 
 val realT = Type ("RealDef.real", []);
@@ -684,9 +687,9 @@
 
 (* random seeds *)
 
-val random_seedT = mk_prodT (code_numeralT, code_numeralT);
+val random_seedT = mk_prodT (code_naturalT, code_naturalT);
 
-fun mk_random T t = Const ("Quickcheck_Random.random_class.random", code_numeralT
+fun mk_random T t = Const ("Quickcheck_Random.random_class.random", code_naturalT
   --> random_seedT --> mk_prodT (mk_prodT (T, unitT --> termT), random_seedT)) $ t;
 
 end;
--- a/src/HOL/Tools/record.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Tools/record.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -1668,7 +1668,7 @@
 fun mk_random_eq tyco vs extN Ts =
   let
     (* FIXME local i etc. *)
-    val size = @{term "i::code_numeral"};
+    val size = @{term "i::natural"};
     fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"});
     val T = Type (tyco, map TFree vs);
     val Tm = termifyT T;
@@ -1688,14 +1688,14 @@
 fun mk_full_exhaustive_eq tyco vs extN Ts =
   let
     (* FIXME local i etc. *)
-    val size = @{term "i::code_numeral"};
+    val size = @{term "i::natural"};
     fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"});
     val T = Type (tyco, map TFree vs);
     val test_function = Free ("f", termifyT T --> @{typ "(bool * term list) option"});
     val params = Name.invent_names Name.context "x" Ts;
     fun full_exhaustiveT T =
       (termifyT T --> @{typ "(bool * Code_Evaluation.term list) option"}) -->
-        @{typ code_numeral} --> @{typ "(bool * Code_Evaluation.term list) option"};
+        @{typ natural} --> @{typ "(bool * Code_Evaluation.term list) option"};
     fun mk_full_exhaustive T =
       Const (@{const_name "Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive"},
         full_exhaustiveT T);
--- a/src/HOL/Word/Word.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Word/Word.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -86,7 +86,7 @@
 
 definition
   "random_word i = Random.range i \<circ>\<rightarrow> (\<lambda>k. Pair (
-     let j = word_of_int (Code_Numeral.int_of k) :: 'a word
+     let j = word_of_int (int_of_integer (integer_of_natural k)) :: 'a word
      in (j, \<lambda>_::unit. Code_Evaluation.term_of j)))"
 
 instance ..
@@ -4651,3 +4651,4 @@
 hide_const (open) Word
 
 end
+
--- a/src/HOL/ex/Code_Binary_Nat_examples.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/ex/Code_Binary_Nat_examples.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -2,10 +2,10 @@
     Author:     Florian Haftmann, TU Muenchen
 *)
 
-header {* Simple examples for natural numbers implemented in binary representation theory. *}
+header {* Simple examples for natural numbers implemented in binary representation. *}
 
 theory Code_Binary_Nat_examples
-imports Complex_Main "~~/src/HOL/Library/Efficient_Nat"
+imports Complex_Main "~~/src/HOL/Library/Code_Binary_Nat"
 begin
 
 fun to_n :: "nat \<Rightarrow> nat list"
--- a/src/HOL/ex/IArray_Examples.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/ex/IArray_Examples.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -1,5 +1,5 @@
 theory IArray_Examples
-imports "~~/src/HOL/Library/IArray" "~~/src/HOL/Library/Efficient_Nat"
+imports "~~/src/HOL/Library/IArray" "~~/src/HOL/Library/Code_Target_Numeral"
 begin
 
 lemma "IArray [True,False] !! 1 = False"
--- a/src/Tools/Code/code_haskell.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Tools/Code/code_haskell.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -418,9 +418,6 @@
   literal_char = Library.enclose "'" "'" o char_haskell,
   literal_string = quote o translate_string char_haskell,
   literal_numeral = numeral_haskell,
-  literal_positive_numeral = numeral_haskell,
-  literal_alternative_numeral = numeral_haskell,
-  literal_naive_numeral = numeral_haskell,
   literal_list = enum "," "[" "]",
   infix_cons = (5, ":")
 } end;
--- a/src/Tools/Code/code_ml.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Tools/Code/code_ml.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -345,9 +345,6 @@
   literal_char = prefix "#" o quote o ML_Syntax.print_char,
   literal_string = quote o translate_string ML_Syntax.print_char,
   literal_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)",
-  literal_positive_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)",
-  literal_alternative_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)",
-  literal_naive_numeral = string_of_int,
   literal_list = enum "," "[" "]",
   infix_cons = (7, "::")
 };
@@ -693,9 +690,6 @@
   literal_char = Library.enclose "'" "'" o char_ocaml,
   literal_string = quote o translate_string char_ocaml,
   literal_numeral = numeral_ocaml,
-  literal_positive_numeral = numeral_ocaml,
-  literal_alternative_numeral = numeral_ocaml,
-  literal_naive_numeral = numeral_ocaml,
   literal_list = enum ";" "[" "]",
   infix_cons = (6, "::")
 } end;
--- a/src/Tools/Code/code_printer.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Tools/Code/code_printer.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -41,17 +41,11 @@
   type literals
   val Literals: { literal_char: string -> string, literal_string: string -> string,
         literal_numeral: int -> string,
-        literal_positive_numeral: int -> string,
-        literal_alternative_numeral: int -> string,
-        literal_naive_numeral: int -> string,
         literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }
     -> literals
   val literal_char: literals -> string -> string
   val literal_string: literals -> string -> string
   val literal_numeral: literals -> int -> string
-  val literal_positive_numeral: literals -> int -> string
-  val literal_alternative_numeral: literals -> int -> string
-  val literal_naive_numeral: literals -> int -> string
   val literal_list: literals -> Pretty.T list -> Pretty.T
   val infix_cons: literals -> int * string
 
@@ -207,9 +201,6 @@
   literal_char: string -> string,
   literal_string: string -> string,
   literal_numeral: int -> string,
-  literal_positive_numeral: int -> string,
-  literal_alternative_numeral: int -> string,
-  literal_naive_numeral: int -> string,
   literal_list: Pretty.T list -> Pretty.T,
   infix_cons: int * string
 };
@@ -219,9 +210,6 @@
 val literal_char = #literal_char o dest_Literals;
 val literal_string = #literal_string o dest_Literals;
 val literal_numeral = #literal_numeral o dest_Literals;
-val literal_positive_numeral = #literal_positive_numeral o dest_Literals;
-val literal_alternative_numeral = #literal_alternative_numeral o dest_Literals;
-val literal_naive_numeral = #literal_naive_numeral o dest_Literals;
 val literal_list = #literal_list o dest_Literals;
 val infix_cons = #infix_cons o dest_Literals;
 
--- a/src/Tools/Code/code_scala.ML	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/Tools/Code/code_scala.ML	Fri Feb 15 08:31:31 2013 +0100
@@ -399,9 +399,6 @@
   literal_char = Library.enclose "'" "'" o char_scala,
   literal_string = quote o translate_string char_scala,
   literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
-  literal_positive_numeral = fn k => "Nat(" ^ numeral_scala k ^ ")",
-  literal_alternative_numeral = fn k => "Natural(" ^ numeral_scala k ^ ")",
-  literal_naive_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
   literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
   infix_cons = (6, "::")
 } end;