merged
authorhuffman
Thu, 29 Apr 2010 11:42:34 -0700
changeset 36594 56ea7385916d
parent 36593 fb69c8cd27bd (current diff)
parent 36540 e31f9ac000dd (diff)
child 36595 c0486affbd9b
merged
src/HOL/IsaMakefile
src/HOL/Quotient_Examples/LarryDatatype.thy
src/HOL/Quotient_Examples/LarryInt.thy
--- a/NEWS	Thu Apr 29 11:41:04 2010 -0700
+++ b/NEWS	Thu Apr 29 11:42:34 2010 -0700
@@ -64,6 +64,11 @@
 
 * Type constructors admit general mixfix syntax, not just infix.
 
+* Concrete syntax may be attached to local entities without a proof
+body, too.  This works via regular mixfix annotations for 'fix',
+'def', 'obtain' etc. or via the explicit 'write' command, which is
+similar to the 'notation' command in theory specifications.
+
 * Use of cumulative prems via "!" in some proof methods has been
 discontinued (legacy feature).
 
@@ -84,6 +89,10 @@
 
 *** Pure ***
 
+* 'code_reflect' allows to incorporate generated ML code into
+runtime environment;  replaces immature code_datatype antiquotation.
+INCOMPATIBILITY.
+
 * Empty class specifications observe default sort.  INCOMPATIBILITY.
 
 * Old 'axclass' has been discontinued.  Use 'class' instead.  INCOMPATIBILITY.
--- a/doc-src/IsarRef/Thy/Inner_Syntax.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -365,6 +365,7 @@
     @{command_def "no_type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
     @{command_def "notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
     @{command_def "no_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+    @{command_def "write"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
   \end{matharray}
 
   \begin{rail}
@@ -372,6 +373,8 @@
     ;
     ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
+    'write' mode? (nameref structmixfix + 'and')
+    ;
   \end{rail}
 
   \begin{description}
@@ -392,12 +395,14 @@
   but removes the specified syntax annotation from the present
   context.
 
+  \item @{command "write"} is similar to @{command "notation"}, but
+  works within an Isar proof body.
+
   \end{description}
 
-  Compared to the underlying @{command "syntax"} and @{command
-  "no_syntax"} primitives (\secref{sec:syn-trans}), the above commands
-  provide explicit checking wrt.\ the logical context, and work within
-  general local theory targets, not just the global theory.
+  Note that the more primitive commands @{command "syntax"} and
+  @{command "no_syntax"} (\secref{sec:syn-trans}) provide raw access
+  to the syntax tables of a global theory.
 *}
 
 
--- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Thu Apr 29 11:41:04 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Thu Apr 29 11:42:34 2010 -0700
@@ -388,6 +388,7 @@
     \indexdef{}{command}{no\_type\_notation}\hypertarget{command.no-type-notation}{\hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}type{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
     \indexdef{}{command}{notation}\hypertarget{command.notation}{\hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
     \indexdef{}{command}{no\_notation}\hypertarget{command.no-notation}{\hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
+    \indexdef{}{command}{write}\hypertarget{command.write}{\hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isasymrightarrow}\ proof{\isacharparenleft}state{\isacharparenright}{\isachardoublequote}} \\
   \end{matharray}
 
   \begin{rail}
@@ -395,6 +396,8 @@
     ;
     ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
+    'write' mode? (nameref structmixfix + 'and')
+    ;
   \end{rail}
 
   \begin{description}
@@ -414,11 +417,14 @@
   but removes the specified syntax annotation from the present
   context.
 
+  \item \hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}, but
+  works within an Isar proof body.
+
   \end{description}
 
-  Compared to the underlying \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} and \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}} primitives (\secref{sec:syn-trans}), the above commands
-  provide explicit checking wrt.\ the logical context, and work within
-  general local theory targets, not just the global theory.%
+  Note that the more primitive commands \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} and
+  \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}} (\secref{sec:syn-trans}) provide raw access
+  to the syntax tables of a global theory.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/etc/isar-keywords-ZF.el	Thu Apr 29 11:41:04 2010 -0700
+++ b/etc/isar-keywords-ZF.el	Thu Apr 29 11:42:34 2010 -0700
@@ -210,6 +210,7 @@
     "using"
     "welcome"
     "with"
+    "write"
     "{"
     "}"))
 
@@ -486,7 +487,8 @@
     "txt"
     "txt_raw"
     "unfolding"
-    "using"))
+    "using"
+    "write"))
 
 (defconst isar-keywords-proof-asm
   '("assume"
--- a/etc/isar-keywords.el	Thu Apr 29 11:41:04 2010 -0700
+++ b/etc/isar-keywords.el	Thu Apr 29 11:42:34 2010 -0700
@@ -273,6 +273,7 @@
     "values"
     "welcome"
     "with"
+    "write"
     "{"
     "}"))
 
@@ -628,7 +629,8 @@
     "txt"
     "txt_raw"
     "unfolding"
-    "using"))
+    "using"
+    "write"))
 
 (defconst isar-keywords-proof-asm
   '("assume"
--- a/src/HOL/Decision_Procs/Approximation.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Decision_Procs/Approximation.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -3209,47 +3209,12 @@
   interpret_floatarith_divide interpret_floatarith_diff interpret_floatarith_tan interpret_floatarith_powr interpret_floatarith_log
   interpret_floatarith_sin
 
-ML {*
-structure Float_Arith =
-struct
-
-@{code_datatype float = Float}
-@{code_datatype floatarith = Add | Minus | Mult | Inverse | Cos | Arctan
-                           | Abs | Max | Min | Pi | Sqrt | Exp | Ln | Power | Var | Num }
-@{code_datatype form = Bound | Assign | Less | LessEqual | AtLeastAtMost}
-
-val approx_form = @{code approx_form}
-val approx_tse_form = @{code approx_tse_form}
-val approx' = @{code approx'}
-val approx_form_eval = @{code approx_form_eval}
-
-end
-*}
-
-code_reserved Eval Float_Arith
-
-code_type float (Eval "Float'_Arith.float")
-code_const Float (Eval "Float'_Arith.Float/ (_,/ _)")
-
-code_type floatarith (Eval "Float'_Arith.floatarith")
-code_const Add and Minus and Mult and Inverse and Cos and Arctan and Abs and Max and Min and
-           Pi and Sqrt  and Exp and Ln and Power and Var and Num
-  (Eval "Float'_Arith.Add/ (_,/ _)" and "Float'_Arith.Minus" and "Float'_Arith.Mult/ (_,/ _)" and
-        "Float'_Arith.Inverse" and "Float'_Arith.Cos" and
-        "Float'_Arith.Arctan" and "Float'_Arith.Abs" and "Float'_Arith.Max/ (_,/ _)" and
-        "Float'_Arith.Min/ (_,/ _)" and "Float'_Arith.Pi" and "Float'_Arith.Sqrt" and
-        "Float'_Arith.Exp" and "Float'_Arith.Ln" and "Float'_Arith.Power/ (_,/ _)" and
-        "Float'_Arith.Var" and "Float'_Arith.Num")
-
-code_type form (Eval "Float'_Arith.form")
-code_const Bound and Assign and Less and LessEqual and AtLeastAtMost
-      (Eval "Float'_Arith.Bound/ (_,/ _,/ _,/ _)" and "Float'_Arith.Assign/ (_,/ _,/ _)" and
-            "Float'_Arith.Less/ (_,/ _)" and "Float'_Arith.LessEqual/ (_,/ _)"  and
-            "Float'_Arith.AtLeastAtMost/ (_,/ _,/ _)")
-
-code_const approx_form (Eval "Float'_Arith.approx'_form")
-code_const approx_tse_form (Eval "Float'_Arith.approx'_tse'_form")
-code_const approx' (Eval "Float'_Arith.approx'")
+code_reflect Float_Arith
+  datatypes float = Float
+  and floatarith = Add | Minus | Mult | Inverse | Cos | Arctan
+    | Abs | Max | Min | Pi | Sqrt | Exp | Ln | Power | Var | Num
+  and form = Bound | Assign | Less | LessEqual | AtLeastAtMost
+  functions approx_form approx_tse_form approx' approx_form_eval
 
 ML {*
   fun reorder_bounds_tac prems i =
--- a/src/HOL/Decision_Procs/Cooper.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Decision_Procs/Cooper.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -1909,10 +1909,9 @@
 
 ML {* @{code cooper_test} () *}
 
-(*
-code_reserved SML oo
-export_code pa in SML module_name GeneratedCooper file "~~/src/HOL/Tools/Qelim/raw_generated_cooper.ML"
-*)
+code_reflect Generated_Cooper
+  functions pa
+  file "~~/src/HOL/Tools/Qelim/generated_cooper.ML"
 
 oracle linzqe_oracle = {*
 let
--- a/src/HOL/Decision_Procs/MIR.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Decision_Procs/MIR.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -5791,8 +5791,9 @@
 ML {* @{code test4} () *}
 ML {* @{code test5} () *}
 
-(*export_code mircfrqe mirlfrqe
-  in SML module_name Mir file "raw_mir.ML"*)
+(*code_reflect Mir
+  functions mircfrqe mirlfrqe
+  file "mir.ML"*)
 
 oracle mirfr_oracle = {* fn (proofs, ct) =>
 let
--- a/src/HOL/FunDef.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/FunDef.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -314,8 +314,8 @@
 
 ML_val -- "setup inactive"
 {*
-  Context.theory_map (Function_Common.set_termination_prover (ScnpReconstruct.decomp_scnp 
-  [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])) 
+  Context.theory_map (Function_Common.set_termination_prover
+    (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))
 *}
 
 end
--- a/src/HOL/HOL.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/HOL.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -1962,6 +1962,10 @@
 
 subsubsection {* Evaluation and normalization by evaluation *}
 
+text {* Avoid some named infixes in evaluation environment *}
+
+code_reserved Eval oo ooo oooo upto downto orf andf mem mem_int mem_string
+
 setup {*
   Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of)
 *}
--- a/src/HOL/IsaMakefile	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/IsaMakefile	Thu Apr 29 11:42:34 2010 -0700
@@ -1299,8 +1299,8 @@
 HOL-Quotient_Examples: HOL $(LOG)/HOL-Quotient_Examples.gz
 
 $(LOG)/HOL-Quotient_Examples.gz: $(OUT)/HOL				\
-  Quotient_Examples/FSet.thy                                            \
-  Quotient_Examples/LarryInt.thy Quotient_Examples/LarryDatatype.thy
+  Quotient_Examples/FSet.thy Quotient_Examples/Quotient_Int.thy         \
+  Quotient_Examples/Quotient_Message.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Quotient_Examples
 
 
--- a/src/HOL/Lazy_Sequence.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Lazy_Sequence.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -123,41 +123,18 @@
 
 subsection {* Code setup *}
 
-ML {*
-signature LAZY_SEQUENCE =
-sig
-  datatype 'a lazy_sequence = Lazy_Sequence of unit -> ('a * 'a lazy_sequence) option
-  val yield : 'a lazy_sequence -> ('a * 'a lazy_sequence) option
-  val yieldn : int -> 'a lazy_sequence -> ('a list * 'a lazy_sequence)
-  val map : ('a -> 'b) -> 'a lazy_sequence -> 'b lazy_sequence
-end;
-
-structure Lazy_Sequence : LAZY_SEQUENCE =
-struct
-
-@{code_datatype lazy_sequence = Lazy_Sequence}
-
-val yield = @{code yield}
+fun anamorph :: "('a \<Rightarrow> ('b \<times> 'a) option) \<Rightarrow> code_numeral \<Rightarrow> 'a \<Rightarrow> 'b list \<times> 'a" where
+  "anamorph f k x = (if k = 0 then ([], x)
+    else case f x of None \<Rightarrow> ([], x) | Some (v, y) \<Rightarrow>
+      let (vs, z) = anamorph f (k - 1) y
+    in (v # vs, z))"
 
-fun anamorph f k x = (if k = 0 then ([], x)
-  else case f x
-   of NONE => ([], x)
-    | SOME (v, y) => let
-        val (vs, z) = anamorph f (k - 1) y
-      in (v :: vs, z) end);
-
-fun yieldn S = anamorph yield S;
+definition yieldn :: "code_numeral \<Rightarrow> 'a lazy_sequence \<Rightarrow> 'a list \<times> 'a lazy_sequence" where
+  "yieldn = anamorph yield"
 
-val map = @{code map}
-
-end;
-*}
-
-code_reserved Eval Lazy_Sequence
-
-code_type lazy_sequence (Eval "_/ Lazy'_Sequence.lazy'_sequence")
-
-code_const Lazy_Sequence (Eval "Lazy'_Sequence.Lazy'_Sequence")
+code_reflect Lazy_Sequence
+  datatypes lazy_sequence = Lazy_Sequence
+  functions map yield yieldn
 
 section {* With Hit Bound Value *}
 text {* assuming in negative context *}
--- a/src/HOL/Predicate.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Predicate.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -880,6 +880,10 @@
 
 code_abort not_unique
 
+code_reflect Predicate
+  datatypes pred = Seq and seq = Empty | Insert | Join
+  functions map
+
 ML {*
 signature PREDICATE =
 sig
@@ -893,15 +897,17 @@
 structure Predicate : PREDICATE =
 struct
 
-@{code_datatype pred = Seq};
-@{code_datatype seq = Empty | Insert | Join};
+datatype pred = datatype Predicate.pred
+datatype seq = datatype Predicate.seq
+
+fun map f = Predicate.map f;
 
-fun yield (@{code Seq} f) = next (f ())
-and next @{code Empty} = NONE
-  | next (@{code Insert} (x, P)) = SOME (x, P)
-  | next (@{code Join} (P, xq)) = (case yield P
+fun yield (Seq f) = next (f ())
+and next Empty = NONE
+  | next (Insert (x, P)) = SOME (x, P)
+  | next (Join (P, xq)) = (case yield P
      of NONE => next xq
-      | SOME (x, Q) => SOME (x, @{code Seq} (fn _ => @{code Join} (Q, xq))));
+      | SOME (x, Q) => SOME (x, Seq (fn _ => Join (Q, xq))));
 
 fun anamorph f k x = (if k = 0 then ([], x)
   else case f x
@@ -912,19 +918,9 @@
 
 fun yieldn P = anamorph yield P;
 
-fun map f = @{code map} f;
-
 end;
 *}
 
-code_reserved Eval Predicate
-
-code_type pred and seq
-  (Eval "_/ Predicate.pred" and "_/ Predicate.seq")
-
-code_const Seq and Empty and Insert and Join
-  (Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)")
-
 no_notation
   inf (infixl "\<sqinter>" 70) and
   sup (infixl "\<squnion>" 65) and
--- a/src/HOL/Quotient_Examples/FSet.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Quotient_Examples/FSet.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Quotient_Examples/Quotient.thy
+(*  Title:      HOL/Quotient_Examples/FSet.thy
     Author:     Cezary Kaliszyk, TU Munich
     Author:     Christian Urban, TU Munich
 
--- a/src/HOL/Quotient_Examples/LarryDatatype.thy	Thu Apr 29 11:41:04 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,394 +0,0 @@
-theory LarryDatatype
-imports Main Quotient_Syntax
-begin
-
-subsection{*Defining the Free Algebra*}
-
-datatype
-  freemsg = NONCE  nat
-        | MPAIR  freemsg freemsg
-        | CRYPT  nat freemsg  
-        | DECRYPT  nat freemsg
-
-inductive 
-  msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
-where 
-  CD:    "CRYPT K (DECRYPT K X) \<sim> X"
-| DC:    "DECRYPT K (CRYPT K X) \<sim> X"
-| NONCE: "NONCE N \<sim> NONCE N"
-| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
-| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
-| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
-| SYM:   "X \<sim> Y \<Longrightarrow> Y \<sim> X"
-| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
-
-lemmas msgrel.intros[intro]
-
-text{*Proving that it is an equivalence relation*}
-
-lemma msgrel_refl: "X \<sim> X"
-by (induct X, (blast intro: msgrel.intros)+)
-
-theorem equiv_msgrel: "equivp msgrel"
-proof (rule equivpI)
-  show "reflp msgrel" by (simp add: reflp_def msgrel_refl)
-  show "symp msgrel" by (simp add: symp_def, blast intro: msgrel.SYM)
-  show "transp msgrel" by (simp add: transp_def, blast intro: msgrel.TRANS)
-qed
-
-subsection{*Some Functions on the Free Algebra*}
-
-subsubsection{*The Set of Nonces*}
-
-fun
-  freenonces :: "freemsg \<Rightarrow> nat set"
-where
-  "freenonces (NONCE N) = {N}"
-| "freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"
-| "freenonces (CRYPT K X) = freenonces X"
-| "freenonces (DECRYPT K X) = freenonces X"
-
-theorem msgrel_imp_eq_freenonces: 
-  assumes a: "U \<sim> V"
-  shows "freenonces U = freenonces V"
-  using a by (induct) (auto) 
-
-subsubsection{*The Left Projection*}
-
-text{*A function to return the left part of the top pair in a message.  It will
-be lifted to the initial algrebra, to serve as an example of that process.*}
-fun
-  freeleft :: "freemsg \<Rightarrow> freemsg"
-where
-  "freeleft (NONCE N) = NONCE N"
-| "freeleft (MPAIR X Y) = X"
-| "freeleft (CRYPT K X) = freeleft X"
-| "freeleft (DECRYPT K X) = freeleft X"
-
-text{*This theorem lets us prove that the left function respects the
-equivalence relation.  It also helps us prove that MPair
-  (the abstract constructor) is injective*}
-lemma msgrel_imp_eqv_freeleft_aux:
-  shows "freeleft U \<sim> freeleft U"
-  by (induct rule: freeleft.induct) (auto)
-
-theorem msgrel_imp_eqv_freeleft:
-  assumes a: "U \<sim> V" 
-  shows "freeleft U \<sim> freeleft V"
-  using a
-  by (induct) (auto intro: msgrel_imp_eqv_freeleft_aux)
-
-subsubsection{*The Right Projection*}
-
-text{*A function to return the right part of the top pair in a message.*}
-fun
-  freeright :: "freemsg \<Rightarrow> freemsg"
-where
-  "freeright (NONCE N) = NONCE N"
-| "freeright (MPAIR X Y) = Y"
-| "freeright (CRYPT K X) = freeright X"
-| "freeright (DECRYPT K X) = freeright X"
-
-text{*This theorem lets us prove that the right function respects the
-equivalence relation.  It also helps us prove that MPair
-  (the abstract constructor) is injective*}
-lemma msgrel_imp_eqv_freeright_aux:
-  shows "freeright U \<sim> freeright U"
-  by (induct rule: freeright.induct) (auto)
-
-theorem msgrel_imp_eqv_freeright:
-  assumes a: "U \<sim> V" 
-  shows "freeright U \<sim> freeright V"
-  using a
-  by (induct) (auto intro: msgrel_imp_eqv_freeright_aux)
-
-subsubsection{*The Discriminator for Constructors*}
-
-text{*A function to distinguish nonces, mpairs and encryptions*}
-fun 
-  freediscrim :: "freemsg \<Rightarrow> int"
-where
-   "freediscrim (NONCE N) = 0"
- | "freediscrim (MPAIR X Y) = 1"
- | "freediscrim (CRYPT K X) = freediscrim X + 2"
- | "freediscrim (DECRYPT K X) = freediscrim X - 2"
-
-text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
-theorem msgrel_imp_eq_freediscrim:
-  assumes a: "U \<sim> V"
-  shows "freediscrim U = freediscrim V"
-  using a by (induct) (auto)
-
-subsection{*The Initial Algebra: A Quotiented Message Type*}
-
-quotient_type msg = freemsg / msgrel
-  by (rule equiv_msgrel)
-
-text{*The abstract message constructors*}
-
-quotient_definition
-  "Nonce :: nat \<Rightarrow> msg"
-is
-  "NONCE"
-
-quotient_definition
-  "MPair :: msg \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "MPAIR"
-
-quotient_definition
-  "Crypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "CRYPT"
-
-quotient_definition
-  "Decrypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "DECRYPT"
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim> ===> op \<sim>) CRYPT CRYPT"
-by (auto intro: CRYPT)
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim> ===> op \<sim>) DECRYPT DECRYPT"
-by (auto intro: DECRYPT)
-
-text{*Establishing these two equations is the point of the whole exercise*}
-theorem CD_eq [simp]: 
-  shows "Crypt K (Decrypt K X) = X"
-  by (lifting CD)
-
-theorem DC_eq [simp]: 
-  shows "Decrypt K (Crypt K X) = X"
-  by (lifting DC)
-
-subsection{*The Abstract Function to Return the Set of Nonces*}
-
-quotient_definition
-   "nonces:: msg \<Rightarrow> nat set"
-is
-  "freenonces"
-
-text{*Now prove the four equations for @{term nonces}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op =) freenonces freenonces"
-  by (simp add: msgrel_imp_eq_freenonces)
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim>) NONCE NONCE"
-  by (simp add: NONCE)
-
-lemma nonces_Nonce [simp]: 
-  shows "nonces (Nonce N) = {N}"
-  by (lifting freenonces.simps(1))
- 
-lemma [quot_respect]:
-  shows " (op \<sim> ===> op \<sim> ===> op \<sim>) MPAIR MPAIR"
-  by (simp add: MPAIR)
-
-lemma nonces_MPair [simp]: 
-  shows "nonces (MPair X Y) = nonces X \<union> nonces Y"
-  by (lifting freenonces.simps(2))
-
-lemma nonces_Crypt [simp]: 
-  shows "nonces (Crypt K X) = nonces X"
-  by (lifting freenonces.simps(3))
-
-lemma nonces_Decrypt [simp]: 
-  shows "nonces (Decrypt K X) = nonces X"
-  by (lifting freenonces.simps(4))
-
-subsection{*The Abstract Function to Return the Left Part*}
-
-quotient_definition
-  "left:: msg \<Rightarrow> msg"
-is
-  "freeleft"
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op \<sim>) freeleft freeleft"
-  by (simp add: msgrel_imp_eqv_freeleft)
-
-lemma left_Nonce [simp]: 
-  shows "left (Nonce N) = Nonce N"
-  by (lifting freeleft.simps(1))
-
-lemma left_MPair [simp]: 
-  shows "left (MPair X Y) = X"
-  by (lifting freeleft.simps(2))
-
-lemma left_Crypt [simp]: 
-  shows "left (Crypt K X) = left X"
-  by (lifting freeleft.simps(3))
-
-lemma left_Decrypt [simp]: 
-  shows "left (Decrypt K X) = left X"
-  by (lifting freeleft.simps(4))
-
-subsection{*The Abstract Function to Return the Right Part*}
-
-quotient_definition
-  "right:: msg \<Rightarrow> msg"
-is
-  "freeright"
-
-text{*Now prove the four equations for @{term right}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op \<sim>) freeright freeright"
-  by (simp add: msgrel_imp_eqv_freeright)
-
-lemma right_Nonce [simp]: 
-  shows "right (Nonce N) = Nonce N"
-  by (lifting freeright.simps(1))
-
-lemma right_MPair [simp]: 
-  shows "right (MPair X Y) = Y"
-  by (lifting freeright.simps(2))
-
-lemma right_Crypt [simp]: 
-  shows "right (Crypt K X) = right X"
-  by (lifting freeright.simps(3))
-
-lemma right_Decrypt [simp]: 
-  shows "right (Decrypt K X) = right X"
-  by (lifting freeright.simps(4))
-
-subsection{*Injectivity Properties of Some Constructors*}
-
-lemma NONCE_imp_eq: 
-  shows "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"
-  by (drule msgrel_imp_eq_freenonces, simp)
-
-text{*Can also be proved using the function @{term nonces}*}
-lemma Nonce_Nonce_eq [iff]: 
-  shows "(Nonce m = Nonce n) = (m = n)"
-proof
-  assume "Nonce m = Nonce n"
-  then show "m = n" by (lifting NONCE_imp_eq)
-next
-  assume "m = n" 
-  then show "Nonce m = Nonce n" by simp
-qed
-
-lemma MPAIR_imp_eqv_left: 
-  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"
-  by (drule msgrel_imp_eqv_freeleft) (simp)
-
-lemma MPair_imp_eq_left: 
-  assumes eq: "MPair X Y = MPair X' Y'" 
-  shows "X = X'"
-  using eq by (lifting MPAIR_imp_eqv_left)
-
-lemma MPAIR_imp_eqv_right: 
-  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"
-  by (drule msgrel_imp_eqv_freeright) (simp)
-
-lemma MPair_imp_eq_right: 
-  shows "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"
-  by (lifting  MPAIR_imp_eqv_right)
-
-theorem MPair_MPair_eq [iff]: 
-  shows "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')" 
-  by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)
-
-lemma NONCE_neqv_MPAIR: 
-  shows "\<not>(NONCE m \<sim> MPAIR X Y)"
-  by (auto dest: msgrel_imp_eq_freediscrim)
-
-theorem Nonce_neq_MPair [iff]: 
-  shows "Nonce N \<noteq> MPair X Y"
-  by (lifting NONCE_neqv_MPAIR)
-
-text{*Example suggested by a referee*}
-
-lemma CRYPT_NONCE_neq_NONCE:
-  shows "\<not>(CRYPT K (NONCE M) \<sim> NONCE N)"
-  by (auto dest: msgrel_imp_eq_freediscrim)
-
-theorem Crypt_Nonce_neq_Nonce: 
-  shows "Crypt K (Nonce M) \<noteq> Nonce N"
-  by (lifting CRYPT_NONCE_neq_NONCE)
-
-text{*...and many similar results*}
-lemma CRYPT2_NONCE_neq_NONCE: 
-  shows "\<not>(CRYPT K (CRYPT K' (NONCE M)) \<sim> NONCE N)"
-  by (auto dest: msgrel_imp_eq_freediscrim)  
-
-theorem Crypt2_Nonce_neq_Nonce: 
-  shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
-  by (lifting CRYPT2_NONCE_neq_NONCE) 
-
-theorem Crypt_Crypt_eq [iff]: 
-  shows "(Crypt K X = Crypt K X') = (X=X')" 
-proof
-  assume "Crypt K X = Crypt K X'"
-  hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp
-  thus "X = X'" by simp
-next
-  assume "X = X'"
-  thus "Crypt K X = Crypt K X'" by simp
-qed
-
-theorem Decrypt_Decrypt_eq [iff]: 
-  shows "(Decrypt K X = Decrypt K X') = (X=X')" 
-proof
-  assume "Decrypt K X = Decrypt K X'"
-  hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp
-  thus "X = X'" by simp
-next
-  assume "X = X'"
-  thus "Decrypt K X = Decrypt K X'" by simp
-qed
-
-lemma msg_induct_aux:
-  shows "\<lbrakk>\<And>N. P (Nonce N);
-          \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y);
-          \<And>K X. P X \<Longrightarrow> P (Crypt K X);
-          \<And>K X. P X \<Longrightarrow> P (Decrypt K X)\<rbrakk> \<Longrightarrow> P msg"
-  by (lifting freemsg.induct)
-
-lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:
-  assumes N: "\<And>N. P (Nonce N)"
-      and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"
-      and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"
-      and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"
-  shows "P msg"
-  using N M C D by (rule msg_induct_aux)
-
-subsection{*The Abstract Discriminator*}
-
-text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
-need this function in order to prove discrimination theorems.*}
-
-quotient_definition
-  "discrim:: msg \<Rightarrow> int"
-is
-  "freediscrim"
-
-text{*Now prove the four equations for @{term discrim}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op =) freediscrim freediscrim"
-  by (auto simp add: msgrel_imp_eq_freediscrim)
-
-lemma discrim_Nonce [simp]: 
-  shows "discrim (Nonce N) = 0"
-  by (lifting freediscrim.simps(1))
-
-lemma discrim_MPair [simp]: 
-  shows "discrim (MPair X Y) = 1"
-  by (lifting freediscrim.simps(2))
-
-lemma discrim_Crypt [simp]: 
-  shows "discrim (Crypt K X) = discrim X + 2"
-  by (lifting freediscrim.simps(3))
-
-lemma discrim_Decrypt [simp]: 
-  shows "discrim (Decrypt K X) = discrim X - 2"
-  by (lifting freediscrim.simps(4))
-
-end
-
--- a/src/HOL/Quotient_Examples/LarryInt.thy	Thu Apr 29 11:41:04 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,395 +0,0 @@
-
-header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
-
-theory LarryInt
-imports Main Quotient_Product
-begin
-
-fun
-  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" 
-where
-  "intrel (x, y) (u, v) = (x + v = u + y)"
-
-quotient_type int = "nat \<times> nat" / intrel
-  by (auto simp add: equivp_def expand_fun_eq)
-
-instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
-begin
-
-quotient_definition
-  Zero_int_def: "0::int" is "(0::nat, 0::nat)"
-
-quotient_definition
-  One_int_def: "1::int" is "(1::nat, 0::nat)"
-
-definition
-  "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
-
-quotient_definition
-  "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
-is
-  "add_raw"
-
-definition
-  "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
-
-quotient_definition
-  "uminus :: int \<Rightarrow> int" 
-is
-  "uminus_raw"
-
-fun
-  mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
-where
-  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
-
-quotient_definition
-  "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
-is
-  "mult_raw"
-
-definition
-  "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
-
-quotient_definition 
-  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
-is
-  "le_raw"
-
-definition
-  less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
-
-definition
-  diff_int_def:  "z - (w::int) \<equiv> z + (-w)"
-
-instance ..
-
-end
-
-subsection{*Construction of the Integers*}
-
-lemma zminus_zminus_raw:
-  "uminus_raw (uminus_raw z) = z"
-  by (cases z) (simp add: uminus_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel) uminus_raw uminus_raw"
-  by (simp add: uminus_raw_def)
-  
-lemma zminus_zminus:
-  fixes z::"int"
-  shows "- (- z) = z"
-  by(lifting zminus_zminus_raw)
-
-lemma zminus_0_raw:
-  shows "uminus_raw (0, 0) = (0, 0::nat)"
-  by (simp add: uminus_raw_def)
-
-lemma zminus_0: 
-  shows "- 0 = (0::int)"
-  by (lifting zminus_0_raw)
-
-subsection{*Integer Addition*}
-
-lemma zminus_zadd_distrib_raw:
-  shows "uminus_raw (add_raw z w) = add_raw (uminus_raw z) (uminus_raw w)"
-by (cases z, cases w)
-   (auto simp add: add_raw_def uminus_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> intrel) add_raw add_raw"
-by (simp add: add_raw_def)
-
-lemma zminus_zadd_distrib: 
-  fixes z w::"int"
-  shows "- (z + w) = (- z) + (- w)"
-  by(lifting zminus_zadd_distrib_raw)
-
-lemma zadd_commute_raw:
-  shows "add_raw z w = add_raw w z"
-by (cases z, cases w)
-   (simp add: add_raw_def)
-
-lemma zadd_commute:
-  fixes z w::"int"
-  shows "(z::int) + w = w + z"
-  by (lifting zadd_commute_raw)
-
-lemma zadd_assoc_raw:
-  shows "add_raw (add_raw z1 z2) z3 = add_raw z1 (add_raw z2 z3)"
-  by (cases z1, cases z2, cases z3) (simp add: add_raw_def)
-
-lemma zadd_assoc: 
-  fixes z1 z2 z3::"int"
-  shows "(z1 + z2) + z3 = z1 + (z2 + z3)"
-  by (lifting zadd_assoc_raw)
-
-lemma zadd_0_raw:
-  shows "add_raw (0, 0) z = z"
-  by (simp add: add_raw_def)
-
-
-text {*also for the instance declaration int :: plus_ac0*}
-lemma zadd_0: 
-  fixes z::"int"
-  shows "0 + z = z"
-  by (lifting zadd_0_raw)
-
-lemma zadd_zminus_inverse_raw:
-  shows "intrel (add_raw (uminus_raw z) z) (0, 0)"
-  by (cases z) (simp add: add_raw_def uminus_raw_def)
-
-lemma zadd_zminus_inverse2: 
-  fixes z::"int"
-  shows "(- z) + z = 0"
-  by (lifting zadd_zminus_inverse_raw)
-
-subsection{*Integer Multiplication*}
-
-lemma zmult_zminus_raw:
-  shows "mult_raw (uminus_raw z) w = uminus_raw (mult_raw z w)"
-apply(cases z, cases w)
-apply(simp add: uminus_raw_def)
-done
-
-lemma mult_raw_fst:
-  assumes a: "intrel x z"
-  shows "intrel (mult_raw x y) (mult_raw z y)"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma mult_raw_snd:
-  assumes a: "intrel x z"
-  shows "intrel (mult_raw y x) (mult_raw y z)"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> intrel) mult_raw mult_raw"
-apply(simp only: fun_rel_def)
-apply(rule allI | rule impI)+
-apply(rule equivp_transp[OF int_equivp])
-apply(rule mult_raw_fst)
-apply(assumption)
-apply(rule mult_raw_snd)
-apply(assumption)
-done
-
-lemma zmult_zminus: 
-  fixes z w::"int"
-  shows "(- z) * w = - (z * w)"
-  by (lifting zmult_zminus_raw)
-
-lemma zmult_commute_raw: 
-  shows "mult_raw z w = mult_raw w z"
-apply(cases z, cases w)
-apply(simp add: add_ac mult_ac)
-done
-
-lemma zmult_commute: 
-  fixes z w::"int"
-  shows "z * w = w * z"
-  by (lifting zmult_commute_raw)
-
-lemma zmult_assoc_raw:
-  shows "mult_raw (mult_raw z1 z2) z3 = mult_raw z1 (mult_raw z2 z3)"
-apply(cases z1, cases z2, cases z3)
-apply(simp add: add_mult_distrib2 mult_ac)
-done
-
-lemma zmult_assoc: 
-  fixes z1 z2 z3::"int"
-  shows "(z1 * z2) * z3 = z1 * (z2 * z3)"
-  by (lifting zmult_assoc_raw)
-
-lemma zadd_mult_distrib_raw:
-  shows "mult_raw (add_raw z1 z2) w = add_raw (mult_raw z1 w) (mult_raw z2 w)"
-apply(cases z1, cases z2, cases w)
-apply(simp add: add_mult_distrib2 mult_ac add_raw_def)
-done
-
-lemma zadd_zmult_distrib: 
-  fixes z1 z2 w::"int"
-  shows "(z1 + z2) * w = (z1 * w) + (z2 * w)"
-  by(lifting zadd_mult_distrib_raw)
-
-lemma zadd_zmult_distrib2: 
-  fixes w z1 z2::"int"
-  shows "w * (z1 + z2) = (w * z1) + (w * z2)"
-  by (simp add: zmult_commute [of w] zadd_zmult_distrib)
-
-lemma zdiff_zmult_distrib: 
-  fixes w z1 z2::"int"
-  shows "(z1 - z2) * w = (z1 * w) - (z2 * w)"
-  by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
-
-lemma zdiff_zmult_distrib2: 
-  fixes w z1 z2::"int"
-  shows "w * (z1 - z2) = (w * z1) - (w * z2)"
-  by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
-
-lemmas int_distrib =
-  zadd_zmult_distrib zadd_zmult_distrib2
-  zdiff_zmult_distrib zdiff_zmult_distrib2
-
-lemma zmult_1_raw:
-  shows "mult_raw (1, 0) z = z"
-  by (cases z) (auto)
-
-lemma zmult_1:
-  fixes z::"int"
-  shows "1 * z = z"
-  by (lifting zmult_1_raw)
-
-lemma zmult_1_right: 
-  fixes z::"int"
-  shows "z * 1 = z"
-  by (rule trans [OF zmult_commute zmult_1])
-
-lemma zero_not_one:
-  shows "\<not>(intrel (0, 0) (1::nat, 0::nat))"
-  by auto
-
-text{*The Integers Form A Ring*}
-instance int :: comm_ring_1
-proof
-  fix i j k :: int
-  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
-  show "i + j = j + i" by (simp add: zadd_commute)
-  show "0 + i = i" by (rule zadd_0)
-  show "- i + i = 0" by (rule zadd_zminus_inverse2)
-  show "i - j = i + (-j)" by (simp add: diff_int_def)
-  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
-  show "i * j = j * i" by (rule zmult_commute)
-  show "1 * i = i" by (rule zmult_1) 
-  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
-  show "0 \<noteq> (1::int)" by (lifting zero_not_one)
-qed
-
-
-subsection{*The @{text "\<le>"} Ordering*}
-
-lemma zle_refl_raw:
-  shows "le_raw w w"
-  by (cases w) (simp add: le_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> op =) le_raw le_raw"
-  by (auto) (simp_all add: le_raw_def)
-
-lemma zle_refl: 
-  fixes w::"int"
-  shows "w \<le> w"
-  by (lifting zle_refl_raw)
-
-
-lemma zle_trans_raw:
-  shows "\<lbrakk>le_raw i j; le_raw j k\<rbrakk> \<Longrightarrow> le_raw i k"
-apply(cases i, cases j, cases k)
-apply(auto simp add: le_raw_def)
-done
-
-lemma zle_trans: 
-  fixes i j k::"int"
-  shows "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> k"
-  by (lifting zle_trans_raw)
-
-lemma zle_anti_sym_raw:
-  shows "\<lbrakk>le_raw z w; le_raw w z\<rbrakk> \<Longrightarrow> intrel z w"
-apply(cases z, cases w)
-apply(auto iff: le_raw_def)
-done
-
-lemma zle_anti_sym: 
-  fixes z w::"int"
-  shows "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = w"
-  by (lifting zle_anti_sym_raw)
-
-
-(* Axiom 'order_less_le' of class 'order': *)
-lemma zless_le: 
-  fixes w z::"int"
-  shows "(w < z) = (w \<le> z & w \<noteq> z)"
-  by (simp add: less_int_def)
-
-instance int :: order
-apply (default)
-apply(auto simp add: zless_le zle_anti_sym)[1]
-apply(rule zle_refl)
-apply(erule zle_trans, assumption)
-apply(erule zle_anti_sym, assumption)
-done
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-
-lemma zle_linear_raw:
-  shows "le_raw z w \<or> le_raw w z"
-apply(cases w, cases z)
-apply(auto iff: le_raw_def)
-done
-
-lemma zle_linear: 
-  fixes z w::"int"
-  shows "z \<le> w \<or> w \<le> z"
-  by (lifting zle_linear_raw)
-
-instance int :: linorder
-apply(default)
-apply(rule zle_linear)
-done
-
-lemma zadd_left_mono_raw:
-  shows "le_raw i j \<Longrightarrow> le_raw (add_raw k i) (add_raw k j)"
-apply(cases k)
-apply(auto simp add: add_raw_def le_raw_def)
-done
-
-lemma zadd_left_mono: 
-  fixes i j::"int"
-  shows "i \<le> j \<Longrightarrow> k + i \<le> k + j"
-  by (lifting zadd_left_mono_raw)
-
-
-subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
-
-definition
-  "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
-
-quotient_definition
-  "nat2::int \<Rightarrow> nat"
-is
-  "nat_raw"
-
-abbreviation
-  "less_raw x y \<equiv> (le_raw x y \<and> \<not>(intrel x y))"
-
-lemma nat_le_eq_zle_raw:
-  shows "less_raw (0, 0) w \<or> le_raw (0, 0) z \<Longrightarrow> (nat_raw w \<le> nat_raw z) = (le_raw w z)"
-  apply (cases w)
-  apply (cases z)
-  apply (simp add: nat_raw_def le_raw_def)
-  by auto
-
-lemma [quot_respect]:
-  shows "(intrel ===> op =) nat_raw nat_raw"
-  by (auto iff: nat_raw_def)
-
-lemma nat_le_eq_zle: 
-  fixes w z::"int"
-  shows "0 < w \<or> 0 \<le> z \<Longrightarrow> (nat2 w \<le> nat2 z) = (w\<le>z)"
-  unfolding less_int_def
-  by (lifting nat_le_eq_zle_raw)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quotient_Examples/Quotient_Int.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -0,0 +1,380 @@
+(*  Title:      HOL/Quotient_Examples/Quotient_Int.thy
+    Author:     Cezary Kaliszyk
+    Author:     Christian Urban
+
+Integers based on Quotients, based on an older version by Larry Paulson.
+*)
+theory Quotient_Int
+imports Quotient_Product Nat
+begin
+
+fun
+  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
+where
+  "intrel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type int = "nat \<times> nat" / intrel
+  by (auto simp add: equivp_def expand_fun_eq)
+
+instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}"
+begin
+
+quotient_definition
+  "0 \<Colon> int" is "(0\<Colon>nat, 0\<Colon>nat)"
+
+quotient_definition
+  "1 \<Colon> int" is "(1\<Colon>nat, 0\<Colon>nat)"
+
+fun
+  plus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "plus_int_raw (x, y) (u, v) = (x + u, y + v)"
+
+quotient_definition
+  "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)" is "plus_int_raw"
+
+fun
+  uminus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "uminus_int_raw (x, y) = (y, x)"
+
+quotient_definition
+  "(uminus \<Colon> (int \<Rightarrow> int))" is "uminus_int_raw"
+
+definition
+  minus_int_def [code del]:  "z - w = z + (-w\<Colon>int)"
+
+fun
+  times_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "times_int_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+
+quotient_definition
+  "(op *) :: (int \<Rightarrow> int \<Rightarrow> int)" is "times_int_raw"
+
+fun
+  le_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
+where
+  "le_int_raw (x, y) (u, v) = (x+v \<le> u+y)"
+
+quotient_definition
+  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" is "le_int_raw"
+
+definition
+  less_int_def [code del]: "(z\<Colon>int) < w = (z \<le> w \<and> z \<noteq> w)"
+
+definition
+  zabs_def: "\<bar>i\<Colon>int\<bar> = (if i < 0 then - i else i)"
+
+definition
+  zsgn_def: "sgn (i\<Colon>int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
+
+instance ..
+
+end
+
+lemma [quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) plus_int_raw plus_int_raw"
+  and   "(op \<approx> ===> op \<approx>) uminus_int_raw uminus_int_raw"
+  and   "(op \<approx> ===> op \<approx> ===> op =) le_int_raw le_int_raw"
+  by auto
+
+lemma times_int_raw_fst:
+  assumes a: "x \<approx> z"
+  shows "times_int_raw x y \<approx> times_int_raw z y"
+  using a
+  apply(cases x, cases y, cases z)
+  apply(auto simp add: times_int_raw.simps intrel.simps)
+  apply(rename_tac u v w x y z)
+  apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+  apply(simp add: mult_ac)
+  apply(simp add: add_mult_distrib [symmetric])
+  done
+
+lemma times_int_raw_snd:
+  assumes a: "x \<approx> z"
+  shows "times_int_raw y x \<approx> times_int_raw y z"
+  using a
+  apply(cases x, cases y, cases z)
+  apply(auto simp add: times_int_raw.simps intrel.simps)
+  apply(rename_tac u v w x y z)
+  apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+  apply(simp add: mult_ac)
+  apply(simp add: add_mult_distrib [symmetric])
+  done
+
+lemma [quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) times_int_raw times_int_raw"
+  apply(simp only: fun_rel_def)
+  apply(rule allI | rule impI)+
+  apply(rule equivp_transp[OF int_equivp])
+  apply(rule times_int_raw_fst)
+  apply(assumption)
+  apply(rule times_int_raw_snd)
+  apply(assumption)
+  done
+
+lemma plus_assoc_raw:
+  shows "plus_int_raw (plus_int_raw i j) k \<approx> plus_int_raw i (plus_int_raw j k)"
+  by (cases i, cases j, cases k) (simp)
+
+lemma plus_sym_raw:
+  shows "plus_int_raw i j \<approx> plus_int_raw j i"
+  by (cases i, cases j) (simp)
+
+lemma plus_zero_raw:
+  shows "plus_int_raw (0, 0) i \<approx> i"
+  by (cases i) (simp)
+
+lemma plus_minus_zero_raw:
+  shows "plus_int_raw (uminus_int_raw i) i \<approx> (0, 0)"
+  by (cases i) (simp)
+
+lemma times_assoc_raw:
+  shows "times_int_raw (times_int_raw i j) k \<approx> times_int_raw i (times_int_raw j k)"
+  by (cases i, cases j, cases k)
+     (simp add: algebra_simps)
+
+lemma times_sym_raw:
+  shows "times_int_raw i j \<approx> times_int_raw j i"
+  by (cases i, cases j) (simp add: algebra_simps)
+
+lemma times_one_raw:
+  shows "times_int_raw  (1, 0) i \<approx> i"
+  by (cases i) (simp)
+
+lemma times_plus_comm_raw:
+  shows "times_int_raw (plus_int_raw i j) k \<approx> plus_int_raw (times_int_raw i k) (times_int_raw j k)"
+by (cases i, cases j, cases k)
+   (simp add: algebra_simps)
+
+lemma one_zero_distinct:
+  shows "\<not> (0, 0) \<approx> ((1::nat), (0::nat))"
+  by simp
+
+text{* The integers form a @{text comm_ring_1}*}
+
+instance int :: comm_ring_1
+proof
+  fix i j k :: int
+  show "(i + j) + k = i + (j + k)"
+    by (lifting plus_assoc_raw)
+  show "i + j = j + i"
+    by (lifting plus_sym_raw)
+  show "0 + i = (i::int)"
+    by (lifting plus_zero_raw)
+  show "- i + i = 0"
+    by (lifting plus_minus_zero_raw)
+  show "i - j = i + - j"
+    by (simp add: minus_int_def)
+  show "(i * j) * k = i * (j * k)"
+    by (lifting times_assoc_raw)
+  show "i * j = j * i"
+    by (lifting times_sym_raw)
+  show "1 * i = i"
+    by (lifting times_one_raw)
+  show "(i + j) * k = i * k + j * k"
+    by (lifting times_plus_comm_raw)
+  show "0 \<noteq> (1::int)"
+    by (lifting one_zero_distinct)
+qed
+
+lemma plus_int_raw_rsp_aux:
+  assumes a: "a \<approx> b" "c \<approx> d"
+  shows "plus_int_raw a c \<approx> plus_int_raw b d"
+  using a
+  by (cases a, cases b, cases c, cases d)
+     (simp)
+
+lemma add_abs_int:
+  "(abs_int (x,y)) + (abs_int (u,v)) =
+   (abs_int (x + u, y + v))"
+  apply(simp add: plus_int_def id_simps)
+  apply(fold plus_int_raw.simps)
+  apply(rule Quotient_rel_abs[OF Quotient_int])
+  apply(rule plus_int_raw_rsp_aux)
+  apply(simp_all add: rep_abs_rsp_left[OF Quotient_int])
+  done
+
+definition int_of_nat_raw:
+  "int_of_nat_raw m = (m :: nat, 0 :: nat)"
+
+quotient_definition
+  "int_of_nat :: nat \<Rightarrow> int" is "int_of_nat_raw"
+
+lemma[quot_respect]:
+  shows "(op = ===> op \<approx>) int_of_nat_raw int_of_nat_raw"
+  by (simp add: equivp_reflp[OF int_equivp])
+
+lemma int_of_nat:
+  shows "of_nat m = int_of_nat m"
+  by (induct m)
+     (simp_all add: zero_int_def one_int_def int_of_nat_def int_of_nat_raw add_abs_int)
+
+lemma le_antisym_raw:
+  shows "le_int_raw i j \<Longrightarrow> le_int_raw j i \<Longrightarrow> i \<approx> j"
+  by (cases i, cases j) (simp)
+
+lemma le_refl_raw:
+  shows "le_int_raw i i"
+  by (cases i) (simp)
+
+lemma le_trans_raw:
+  shows "le_int_raw i j \<Longrightarrow> le_int_raw j k \<Longrightarrow> le_int_raw i k"
+  by (cases i, cases j, cases k) (simp)
+
+lemma le_cases_raw:
+  shows "le_int_raw i j \<or> le_int_raw j i"
+  by (cases i, cases j)
+     (simp add: linorder_linear)
+
+instance int :: linorder
+proof
+  fix i j k :: int
+  show antisym: "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j"
+    by (lifting le_antisym_raw)
+  show "(i < j) = (i \<le> j \<and> \<not> j \<le> i)"
+    by (auto simp add: less_int_def dest: antisym)
+  show "i \<le> i"
+    by (lifting le_refl_raw)
+  show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k"
+    by (lifting le_trans_raw)
+  show "i \<le> j \<or> j \<le> i"
+    by (lifting le_cases_raw)
+qed
+
+instantiation int :: distrib_lattice
+begin
+
+definition
+  "(inf \<Colon> int \<Rightarrow> int \<Rightarrow> int) = min"
+
+definition
+  "(sup \<Colon> int \<Rightarrow> int \<Rightarrow> int) = max"
+
+instance
+  by default
+     (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
+
+end
+
+lemma le_plus_int_raw:
+  shows "le_int_raw i j \<Longrightarrow> le_int_raw (plus_int_raw k i) (plus_int_raw k j)"
+  by (cases i, cases j, cases k) (simp)
+
+instance int :: ordered_cancel_ab_semigroup_add
+proof
+  fix i j k :: int
+  show "i \<le> j \<Longrightarrow> k + i \<le> k + j"
+    by (lifting le_plus_int_raw)
+qed
+
+abbreviation
+  "less_int_raw i j \<equiv> le_int_raw i j \<and> \<not>(i \<approx> j)"
+
+lemma zmult_zless_mono2_lemma:
+  fixes i j::int
+  and   k::nat
+  shows "i < j \<Longrightarrow> 0 < k \<Longrightarrow> of_nat k * i < of_nat k * j"
+  apply(induct "k")
+  apply(simp)
+  apply(case_tac "k = 0")
+  apply(simp_all add: left_distrib add_strict_mono)
+  done
+
+lemma zero_le_imp_eq_int_raw:
+  fixes k::"(nat \<times> nat)"
+  shows "less_int_raw (0, 0) k \<Longrightarrow> (\<exists>n > 0. k \<approx> int_of_nat_raw n)"
+  apply(cases k)
+  apply(simp add:int_of_nat_raw)
+  apply(auto)
+  apply(rule_tac i="b" and j="a" in less_Suc_induct)
+  apply(auto)
+  done
+
+lemma zero_le_imp_eq_int:
+  fixes k::int
+  shows "0 < k \<Longrightarrow> \<exists>n > 0. k = of_nat n"
+  unfolding less_int_def int_of_nat
+  by (lifting zero_le_imp_eq_int_raw)
+
+lemma zmult_zless_mono2:
+  fixes i j k::int
+  assumes a: "i < j" "0 < k"
+  shows "k * i < k * j"
+  using a
+  by (drule_tac zero_le_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma)
+
+text{*The integers form an ordered integral domain*}
+
+instance int :: linordered_idom
+proof
+  fix i j k :: int
+  show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"
+    by (rule zmult_zless_mono2)
+  show "\<bar>i\<bar> = (if i < 0 then -i else i)"
+    by (simp only: zabs_def)
+  show "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
+    by (simp only: zsgn_def)
+qed
+
+lemmas int_distrib =
+  left_distrib [of "z1::int" "z2" "w", standard]
+  right_distrib [of "w::int" "z1" "z2", standard]
+  left_diff_distrib [of "z1::int" "z2" "w", standard]
+  right_diff_distrib [of "w::int" "z1" "z2", standard]
+  minus_add_distrib[of "z1::int" "z2", standard]
+
+lemma int_induct_raw:
+  assumes a: "P (0::nat, 0)"
+  and     b: "\<And>i. P i \<Longrightarrow> P (plus_int_raw i (1, 0))"
+  and     c: "\<And>i. P i \<Longrightarrow> P (plus_int_raw i (uminus_int_raw (1, 0)))"
+  shows      "P x"
+proof (cases x, clarify)
+  fix a b
+  show "P (a, b)"
+  proof (induct a arbitrary: b rule: Nat.induct)
+    case zero
+    show "P (0, b)" using assms by (induct b) auto
+  next
+    case (Suc n)
+    then show ?case using assms by auto
+  qed
+qed
+
+lemma int_induct:
+  fixes x :: int
+  assumes a: "P 0"
+  and     b: "\<And>i. P i \<Longrightarrow> P (i + 1)"
+  and     c: "\<And>i. P i \<Longrightarrow> P (i - 1)"
+  shows      "P x"
+  using a b c unfolding minus_int_def
+  by (lifting int_induct_raw)
+
+text {* Magnitide of an Integer, as a Natural Number: @{term nat} *}
+
+definition
+  "int_to_nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
+
+quotient_definition
+  "int_to_nat::int \<Rightarrow> nat"
+is
+  "int_to_nat_raw"
+
+lemma [quot_respect]:
+  shows "(intrel ===> op =) int_to_nat_raw int_to_nat_raw"
+  by (auto iff: int_to_nat_raw_def)
+
+lemma nat_le_eq_zle_raw:
+  assumes a: "less_int_raw (0, 0) w \<or> le_int_raw (0, 0) z"
+  shows "(int_to_nat_raw w \<le> int_to_nat_raw z) = (le_int_raw w z)"
+  using a
+  by (cases w, cases z) (auto simp add: int_to_nat_raw_def)
+
+lemma nat_le_eq_zle:
+  fixes w z::"int"
+  shows "0 < w \<or> 0 \<le> z \<Longrightarrow> (int_to_nat w \<le> int_to_nat z) = (w \<le> z)"
+  unfolding less_int_def
+  by (lifting nat_le_eq_zle_raw)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quotient_Examples/Quotient_Message.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -0,0 +1,399 @@
+(*  Title:      HOL/Quotient_Examples/Quotient_Message.thy
+    Author:     Christian Urban
+
+Message datatype, based on an older version by Larry Paulson.
+*)
+theory Quotient_Message
+imports Main Quotient_Syntax
+begin
+
+subsection{*Defining the Free Algebra*}
+
+datatype
+  freemsg = NONCE  nat
+        | MPAIR  freemsg freemsg
+        | CRYPT  nat freemsg
+        | DECRYPT  nat freemsg
+
+inductive
+  msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
+where
+  CD:    "CRYPT K (DECRYPT K X) \<sim> X"
+| DC:    "DECRYPT K (CRYPT K X) \<sim> X"
+| NONCE: "NONCE N \<sim> NONCE N"
+| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
+| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
+| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
+| SYM:   "X \<sim> Y \<Longrightarrow> Y \<sim> X"
+| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
+
+lemmas msgrel.intros[intro]
+
+text{*Proving that it is an equivalence relation*}
+
+lemma msgrel_refl: "X \<sim> X"
+by (induct X, (blast intro: msgrel.intros)+)
+
+theorem equiv_msgrel: "equivp msgrel"
+proof (rule equivpI)
+  show "reflp msgrel" by (simp add: reflp_def msgrel_refl)
+  show "symp msgrel" by (simp add: symp_def, blast intro: msgrel.SYM)
+  show "transp msgrel" by (simp add: transp_def, blast intro: msgrel.TRANS)
+qed
+
+subsection{*Some Functions on the Free Algebra*}
+
+subsubsection{*The Set of Nonces*}
+
+fun
+  freenonces :: "freemsg \<Rightarrow> nat set"
+where
+  "freenonces (NONCE N) = {N}"
+| "freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"
+| "freenonces (CRYPT K X) = freenonces X"
+| "freenonces (DECRYPT K X) = freenonces X"
+
+theorem msgrel_imp_eq_freenonces:
+  assumes a: "U \<sim> V"
+  shows "freenonces U = freenonces V"
+  using a by (induct) (auto)
+
+subsubsection{*The Left Projection*}
+
+text{*A function to return the left part of the top pair in a message.  It will
+be lifted to the initial algrebra, to serve as an example of that process.*}
+fun
+  freeleft :: "freemsg \<Rightarrow> freemsg"
+where
+  "freeleft (NONCE N) = NONCE N"
+| "freeleft (MPAIR X Y) = X"
+| "freeleft (CRYPT K X) = freeleft X"
+| "freeleft (DECRYPT K X) = freeleft X"
+
+text{*This theorem lets us prove that the left function respects the
+equivalence relation.  It also helps us prove that MPair
+  (the abstract constructor) is injective*}
+lemma msgrel_imp_eqv_freeleft_aux:
+  shows "freeleft U \<sim> freeleft U"
+  by (induct rule: freeleft.induct) (auto)
+
+theorem msgrel_imp_eqv_freeleft:
+  assumes a: "U \<sim> V"
+  shows "freeleft U \<sim> freeleft V"
+  using a
+  by (induct) (auto intro: msgrel_imp_eqv_freeleft_aux)
+
+subsubsection{*The Right Projection*}
+
+text{*A function to return the right part of the top pair in a message.*}
+fun
+  freeright :: "freemsg \<Rightarrow> freemsg"
+where
+  "freeright (NONCE N) = NONCE N"
+| "freeright (MPAIR X Y) = Y"
+| "freeright (CRYPT K X) = freeright X"
+| "freeright (DECRYPT K X) = freeright X"
+
+text{*This theorem lets us prove that the right function respects the
+equivalence relation.  It also helps us prove that MPair
+  (the abstract constructor) is injective*}
+lemma msgrel_imp_eqv_freeright_aux:
+  shows "freeright U \<sim> freeright U"
+  by (induct rule: freeright.induct) (auto)
+
+theorem msgrel_imp_eqv_freeright:
+  assumes a: "U \<sim> V"
+  shows "freeright U \<sim> freeright V"
+  using a
+  by (induct) (auto intro: msgrel_imp_eqv_freeright_aux)
+
+subsubsection{*The Discriminator for Constructors*}
+
+text{*A function to distinguish nonces, mpairs and encryptions*}
+fun
+  freediscrim :: "freemsg \<Rightarrow> int"
+where
+   "freediscrim (NONCE N) = 0"
+ | "freediscrim (MPAIR X Y) = 1"
+ | "freediscrim (CRYPT K X) = freediscrim X + 2"
+ | "freediscrim (DECRYPT K X) = freediscrim X - 2"
+
+text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
+theorem msgrel_imp_eq_freediscrim:
+  assumes a: "U \<sim> V"
+  shows "freediscrim U = freediscrim V"
+  using a by (induct) (auto)
+
+subsection{*The Initial Algebra: A Quotiented Message Type*}
+
+quotient_type msg = freemsg / msgrel
+  by (rule equiv_msgrel)
+
+text{*The abstract message constructors*}
+
+quotient_definition
+  "Nonce :: nat \<Rightarrow> msg"
+is
+  "NONCE"
+
+quotient_definition
+  "MPair :: msg \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "MPAIR"
+
+quotient_definition
+  "Crypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "CRYPT"
+
+quotient_definition
+  "Decrypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "DECRYPT"
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim> ===> op \<sim>) CRYPT CRYPT"
+by (auto intro: CRYPT)
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim> ===> op \<sim>) DECRYPT DECRYPT"
+by (auto intro: DECRYPT)
+
+text{*Establishing these two equations is the point of the whole exercise*}
+theorem CD_eq [simp]:
+  shows "Crypt K (Decrypt K X) = X"
+  by (lifting CD)
+
+theorem DC_eq [simp]:
+  shows "Decrypt K (Crypt K X) = X"
+  by (lifting DC)
+
+subsection{*The Abstract Function to Return the Set of Nonces*}
+
+quotient_definition
+   "nonces:: msg \<Rightarrow> nat set"
+is
+  "freenonces"
+
+text{*Now prove the four equations for @{term nonces}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op =) freenonces freenonces"
+  by (simp add: msgrel_imp_eq_freenonces)
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim>) NONCE NONCE"
+  by (simp add: NONCE)
+
+lemma nonces_Nonce [simp]:
+  shows "nonces (Nonce N) = {N}"
+  by (lifting freenonces.simps(1))
+
+lemma [quot_respect]:
+  shows " (op \<sim> ===> op \<sim> ===> op \<sim>) MPAIR MPAIR"
+  by (simp add: MPAIR)
+
+lemma nonces_MPair [simp]:
+  shows "nonces (MPair X Y) = nonces X \<union> nonces Y"
+  by (lifting freenonces.simps(2))
+
+lemma nonces_Crypt [simp]:
+  shows "nonces (Crypt K X) = nonces X"
+  by (lifting freenonces.simps(3))
+
+lemma nonces_Decrypt [simp]:
+  shows "nonces (Decrypt K X) = nonces X"
+  by (lifting freenonces.simps(4))
+
+subsection{*The Abstract Function to Return the Left Part*}
+
+quotient_definition
+  "left:: msg \<Rightarrow> msg"
+is
+  "freeleft"
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op \<sim>) freeleft freeleft"
+  by (simp add: msgrel_imp_eqv_freeleft)
+
+lemma left_Nonce [simp]:
+  shows "left (Nonce N) = Nonce N"
+  by (lifting freeleft.simps(1))
+
+lemma left_MPair [simp]:
+  shows "left (MPair X Y) = X"
+  by (lifting freeleft.simps(2))
+
+lemma left_Crypt [simp]:
+  shows "left (Crypt K X) = left X"
+  by (lifting freeleft.simps(3))
+
+lemma left_Decrypt [simp]:
+  shows "left (Decrypt K X) = left X"
+  by (lifting freeleft.simps(4))
+
+subsection{*The Abstract Function to Return the Right Part*}
+
+quotient_definition
+  "right:: msg \<Rightarrow> msg"
+is
+  "freeright"
+
+text{*Now prove the four equations for @{term right}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op \<sim>) freeright freeright"
+  by (simp add: msgrel_imp_eqv_freeright)
+
+lemma right_Nonce [simp]:
+  shows "right (Nonce N) = Nonce N"
+  by (lifting freeright.simps(1))
+
+lemma right_MPair [simp]:
+  shows "right (MPair X Y) = Y"
+  by (lifting freeright.simps(2))
+
+lemma right_Crypt [simp]:
+  shows "right (Crypt K X) = right X"
+  by (lifting freeright.simps(3))
+
+lemma right_Decrypt [simp]:
+  shows "right (Decrypt K X) = right X"
+  by (lifting freeright.simps(4))
+
+subsection{*Injectivity Properties of Some Constructors*}
+
+lemma NONCE_imp_eq:
+  shows "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"
+  by (drule msgrel_imp_eq_freenonces, simp)
+
+text{*Can also be proved using the function @{term nonces}*}
+lemma Nonce_Nonce_eq [iff]:
+  shows "(Nonce m = Nonce n) = (m = n)"
+proof
+  assume "Nonce m = Nonce n"
+  then show "m = n" by (lifting NONCE_imp_eq)
+next
+  assume "m = n"
+  then show "Nonce m = Nonce n" by simp
+qed
+
+lemma MPAIR_imp_eqv_left:
+  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"
+  by (drule msgrel_imp_eqv_freeleft) (simp)
+
+lemma MPair_imp_eq_left:
+  assumes eq: "MPair X Y = MPair X' Y'"
+  shows "X = X'"
+  using eq by (lifting MPAIR_imp_eqv_left)
+
+lemma MPAIR_imp_eqv_right:
+  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"
+  by (drule msgrel_imp_eqv_freeright) (simp)
+
+lemma MPair_imp_eq_right:
+  shows "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"
+  by (lifting  MPAIR_imp_eqv_right)
+
+theorem MPair_MPair_eq [iff]:
+  shows "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')"
+  by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)
+
+lemma NONCE_neqv_MPAIR:
+  shows "\<not>(NONCE m \<sim> MPAIR X Y)"
+  by (auto dest: msgrel_imp_eq_freediscrim)
+
+theorem Nonce_neq_MPair [iff]:
+  shows "Nonce N \<noteq> MPair X Y"
+  by (lifting NONCE_neqv_MPAIR)
+
+text{*Example suggested by a referee*}
+
+lemma CRYPT_NONCE_neq_NONCE:
+  shows "\<not>(CRYPT K (NONCE M) \<sim> NONCE N)"
+  by (auto dest: msgrel_imp_eq_freediscrim)
+
+theorem Crypt_Nonce_neq_Nonce:
+  shows "Crypt K (Nonce M) \<noteq> Nonce N"
+  by (lifting CRYPT_NONCE_neq_NONCE)
+
+text{*...and many similar results*}
+lemma CRYPT2_NONCE_neq_NONCE:
+  shows "\<not>(CRYPT K (CRYPT K' (NONCE M)) \<sim> NONCE N)"
+  by (auto dest: msgrel_imp_eq_freediscrim)
+
+theorem Crypt2_Nonce_neq_Nonce:
+  shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
+  by (lifting CRYPT2_NONCE_neq_NONCE)
+
+theorem Crypt_Crypt_eq [iff]:
+  shows "(Crypt K X = Crypt K X') = (X=X')"
+proof
+  assume "Crypt K X = Crypt K X'"
+  hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp
+  thus "X = X'" by simp
+next
+  assume "X = X'"
+  thus "Crypt K X = Crypt K X'" by simp
+qed
+
+theorem Decrypt_Decrypt_eq [iff]:
+  shows "(Decrypt K X = Decrypt K X') = (X=X')"
+proof
+  assume "Decrypt K X = Decrypt K X'"
+  hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp
+  thus "X = X'" by simp
+next
+  assume "X = X'"
+  thus "Decrypt K X = Decrypt K X'" by simp
+qed
+
+lemma msg_induct_aux:
+  shows "\<lbrakk>\<And>N. P (Nonce N);
+          \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y);
+          \<And>K X. P X \<Longrightarrow> P (Crypt K X);
+          \<And>K X. P X \<Longrightarrow> P (Decrypt K X)\<rbrakk> \<Longrightarrow> P msg"
+  by (lifting freemsg.induct)
+
+lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:
+  assumes N: "\<And>N. P (Nonce N)"
+      and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"
+      and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"
+      and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"
+  shows "P msg"
+  using N M C D by (rule msg_induct_aux)
+
+subsection{*The Abstract Discriminator*}
+
+text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
+need this function in order to prove discrimination theorems.*}
+
+quotient_definition
+  "discrim:: msg \<Rightarrow> int"
+is
+  "freediscrim"
+
+text{*Now prove the four equations for @{term discrim}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op =) freediscrim freediscrim"
+  by (auto simp add: msgrel_imp_eq_freediscrim)
+
+lemma discrim_Nonce [simp]:
+  shows "discrim (Nonce N) = 0"
+  by (lifting freediscrim.simps(1))
+
+lemma discrim_MPair [simp]:
+  shows "discrim (MPair X Y) = 1"
+  by (lifting freediscrim.simps(2))
+
+lemma discrim_Crypt [simp]:
+  shows "discrim (Crypt K X) = discrim X + 2"
+  by (lifting freediscrim.simps(3))
+
+lemma discrim_Decrypt [simp]:
+  shows "discrim (Decrypt K X) = discrim X - 2"
+  by (lifting freediscrim.simps(4))
+
+end
+
--- a/src/HOL/Quotient_Examples/ROOT.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Quotient_Examples/ROOT.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -4,5 +4,5 @@
 Testing the quotient package.
 *)
 
-use_thys ["FSet", "LarryInt", "LarryDatatype"];
+use_thys ["FSet", "Quotient_Int", "Quotient_Message"];
 
--- a/src/HOL/Random.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Random.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -138,10 +138,15 @@
 
 subsection {* @{text ML} interface *}
 
+code_reflect Random_Engine
+  functions range select select_weight
+
 ML {*
 structure Random_Engine =
 struct
 
+open Random_Engine;
+
 type seed = int * int;
 
 local
--- a/src/HOL/Statespace/state_space.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Statespace/state_space.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -198,7 +198,7 @@
   if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name
   then
     let val n' = lookupI (op =) (Variable.fixes_of ctxt) name
-    in SOME (Free (n',ProofContext.infer_type ctxt n')) end
+    in SOME (Free (n',ProofContext.infer_type ctxt (n', dummyT))) end
   else NONE
 
 
@@ -430,7 +430,7 @@
           let
             fun upd (n,v) =
               let
-                val nT = ProofContext.infer_type (Local_Theory.target_of lthy) n
+                val nT = ProofContext.infer_type (Local_Theory.target_of lthy) (n, dummyT)
               in Context.proof_map
                   (update_declinfo (Morphism.term phi (Free (n,nT)),v))
               end;
--- a/src/HOL/TLA/Memory/MemoryParameters.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/TLA/Memory/MemoryParameters.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -12,7 +12,7 @@
 begin
 
 (* the memory operations *)
-datatype memOp = read Locs | write Locs Vals
+datatype memOp = read Locs | "write" Locs Vals
 
 consts
   (* memory locations and contents *)
--- a/src/HOL/Tools/Function/fun.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Function/fun.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -7,12 +7,12 @@
 
 signature FUNCTION_FUN =
 sig
-  val add_fun : Function_Common.function_config ->
-    (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
-    bool -> local_theory -> Proof.context
-  val add_fun_cmd : Function_Common.function_config ->
-    (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
-    bool -> local_theory -> Proof.context
+  val add_fun : (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> Function_Common.function_config ->
+    local_theory -> Proof.context
+  val add_fun_cmd : (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> Function_Common.function_config ->
+    local_theory -> Proof.context
 
   val setup : theory -> theory
 end
@@ -56,15 +56,6 @@
     ()
   end
 
-val by_pat_completeness_auto =
-  Proof.global_future_terminal_proof
-    (Method.Basic Pat_Completeness.pat_completeness,
-     SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
-
-fun termination_by method int =
-  Function.termination_proof NONE
-  #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
-
 fun mk_catchall fixes arity_of =
   let
     fun mk_eqn ((fname, fT), _) =
@@ -148,24 +139,32 @@
 val fun_config = FunctionConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*), 
   domintros=false, partials=false, tailrec=false }
 
-fun gen_fun add config fixes statements int lthy =
-  lthy
-  |> add fixes statements config
-  |> by_pat_completeness_auto int
-  |> Local_Theory.restore
-  |> termination_by (Function_Common.get_termination_prover lthy) int
+fun gen_add_fun add fixes statements config lthy =
+  let
+    fun pat_completeness_auto ctxt =
+      Pat_Completeness.pat_completeness_tac ctxt 1
+      THEN auto_tac (clasimpset_of ctxt)
+    fun prove_termination lthy =
+      Function.prove_termination NONE
+        (Function_Common.get_termination_prover lthy lthy) lthy
+  in
+    lthy
+    |> add fixes statements config pat_completeness_auto |> snd
+    |> Local_Theory.restore
+    |> prove_termination
+  end
 
-val add_fun = gen_fun Function.add_function
-val add_fun_cmd = gen_fun Function.add_function_cmd
+val add_fun = gen_add_fun Function.add_function
+val add_fun_cmd = gen_add_fun Function.add_function_cmd
 
 
 
 local structure P = OuterParse and K = OuterKeyword in
 
 val _ =
-  OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
+  OuterSyntax.local_theory "fun" "define general recursive functions (short version)" K.thy_decl
   (function_parser fun_config
-     >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
+     >> (fn ((config, fixes), statements) => add_fun_cmd fixes statements config));
 
 end
 
--- a/src/HOL/Tools/Function/function.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Function/function.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -11,14 +11,25 @@
 
   val add_function: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> Function_Common.function_config ->
-    local_theory -> Proof.state
+    (Proof.context -> tactic) -> local_theory -> info * local_theory
 
   val add_function_cmd: (binding * string option * mixfix) list ->
     (Attrib.binding * string) list -> Function_Common.function_config ->
+    (Proof.context -> tactic) -> local_theory -> info * local_theory
+
+  val function: (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> Function_Common.function_config ->
     local_theory -> Proof.state
 
-  val termination_proof : term option -> local_theory -> Proof.state
-  val termination_proof_cmd : string option -> local_theory -> Proof.state
+  val function_cmd: (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> Function_Common.function_config ->
+    local_theory -> Proof.state
+
+  val prove_termination: term option -> tactic -> local_theory -> local_theory
+  val prove_termination_cmd: string option -> tactic -> local_theory -> local_theory
+
+  val termination : term option -> local_theory -> Proof.state
+  val termination_cmd : string option -> local_theory -> Proof.state
 
   val setup : theory -> theory
   val get_congs : Proof.context -> thm list
@@ -65,7 +76,7 @@
     (saved_simps, fold2 add_for_f fnames simps_by_f lthy)
   end
 
-fun gen_add_function is_external prep default_constraint fixspec eqns config lthy =
+fun prepare_function is_external prep default_constraint fixspec eqns config lthy =
   let
     val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
     val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
@@ -76,7 +87,7 @@
     val defname = mk_defname fixes
     val FunctionConfig {partials, ...} = config
 
-    val ((goalstate, cont), lthy) =
+    val ((goal_state, cont), lthy') =
       Function_Mutual.prepare_function_mutual config defname fixes eqs lthy
 
     fun afterqed [[proof]] lthy =
@@ -115,20 +126,45 @@
           if not is_external then ()
           else Specification.print_consts lthy (K false) (map fst fixes)
       in
-        lthy
-        |> Local_Theory.declaration false (add_function_data o morph_function_data info)
+        (info, 
+         lthy |> Local_Theory.declaration false (add_function_data o morph_function_data info))
       end
   in
-    lthy
-    |> Proof.theorem NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
-    |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
+    ((goal_state, afterqed), lthy')
+  end
+
+fun gen_add_function is_external prep default_constraint fixspec eqns config tac lthy =
+  let
+    val ((goal_state, afterqed), lthy') =
+      prepare_function is_external prep default_constraint fixspec eqns config lthy
+    val pattern_thm =
+      case SINGLE (tac lthy') goal_state of
+        NONE => error "pattern completeness and compatibility proof failed"
+      | SOME st => Goal.finish lthy' st
+  in
+    lthy'
+    |> afterqed [[pattern_thm]]
   end
 
 val add_function =
   gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
 val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
-                                                
-fun gen_termination_proof prep_term raw_term_opt lthy =
+
+fun gen_function is_external prep default_constraint fixspec eqns config lthy =
+  let
+    val ((goal_state, afterqed), lthy') =
+      prepare_function is_external prep default_constraint fixspec eqns config lthy
+  in
+    lthy'
+    |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
+    |> Proof.refine (Method.primitive_text (K goal_state)) |> Seq.hd
+  end
+
+val function =
+  gen_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
+val function_cmd = gen_function true Specification.read_spec "_::type"
+
+fun prepare_termination_proof prep_term raw_term_opt lthy =
   let
     val term_opt = Option.map (prep_term lthy) raw_term_opt
     val info = the (case term_opt of
@@ -169,6 +205,26 @@
             end)
         end
   in
+    (goal, afterqed, termination)
+  end
+
+fun gen_prove_termination prep_term raw_term_opt tac lthy =
+  let
+    val (goal, afterqed, termination) =
+      prepare_termination_proof prep_term raw_term_opt lthy
+
+    val totality = Goal.prove lthy [] [] goal (K tac)
+  in
+    afterqed [[totality]] lthy
+end
+
+val prove_termination = gen_prove_termination Syntax.check_term
+val prove_termination_cmd = gen_prove_termination Syntax.read_term
+
+fun gen_termination prep_term raw_term_opt lthy =
+  let
+    val (goal, afterqed, termination) = prepare_termination_proof prep_term raw_term_opt lthy
+  in
     lthy
     |> ProofContext.note_thmss ""
        [((Binding.empty, [Context_Rules.rule_del]), [([allI], [])])] |> snd
@@ -180,8 +236,8 @@
     |> Proof.theorem NONE afterqed [[(goal, [])]]
   end
 
-val termination_proof = gen_termination_proof Syntax.check_term
-val termination_proof_cmd = gen_termination_proof Syntax.read_term
+val termination = gen_termination Syntax.check_term
+val termination_cmd = gen_termination Syntax.read_term
 
 
 (* Datatype hook to declare datatype congs as "function_congs" *)
@@ -221,11 +277,11 @@
 val _ =
   OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
   (function_parser default_config
-     >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config))
+     >> (fn ((config, fixes), statements) => function_cmd fixes statements config))
 
 val _ =
   OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
-  (Scan.option P.term >> termination_proof_cmd)
+  (Scan.option P.term >> termination_cmd)
 
 end
 
--- a/src/HOL/Tools/Function/function_common.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Function/function_common.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -172,7 +172,7 @@
 
 structure TerminationProver = Generic_Data
 (
-  type T = Proof.context -> Proof.method
+  type T = Proof.context -> tactic
   val empty = (fn _ => error "Termination prover not configured")
   val extend = I
   fun merge (a, b) = b  (* FIXME ? *)
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -225,6 +225,6 @@
   Method.setup @{binding lexicographic_order}
     (Method.sections clasimp_modifiers >> (K lexicographic_order))
     "termination prover for lexicographic orderings"
-  #> Context.theory_map (Function_Common.set_termination_prover lexicographic_order)
+  #> Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false))
 
 end;
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -10,7 +10,7 @@
 
   val sizechange_tac : Proof.context -> tactic -> tactic
 
-  val decomp_scnp : ScnpSolve.label list -> Proof.context -> Proof.method
+  val decomp_scnp_tac : ScnpSolve.label list -> Proof.context -> tactic
 
   val setup : theory -> theory
 
@@ -396,13 +396,12 @@
 fun sizechange_tac ctxt autom_tac =
   gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt (K (K all_tac))
 
-fun decomp_scnp orders ctxt =
+fun decomp_scnp_tac orders ctxt =
   let
     val extra_simps = Function_Common.Termination_Simps.get ctxt
     val autom_tac = auto_tac (clasimpset_of ctxt addsimps2 extra_simps)
   in
-    SIMPLE_METHOD
-      (gen_sizechange_tac orders autom_tac ctxt (print_error ctxt))
+     gen_sizechange_tac orders autom_tac ctxt (print_error ctxt)
   end
 
 
@@ -416,7 +415,8 @@
   || Scan.succeed [MAX, MS, MIN]
 
 val setup = Method.setup @{binding size_change}
-  (Scan.lift orders --| Method.sections clasimp_modifiers >> decomp_scnp)
+  (Scan.lift orders --| Method.sections clasimp_modifiers >>
+    (fn orders => SIMPLE_METHOD o decomp_scnp_tac orders))
   "termination prover with graph decomposition and the NP subset of size change termination"
 
 end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -3232,14 +3232,14 @@
                 (Code_Eval.eval NONE
                   ("Predicate_Compile_Core.new_random_dseq_stats_eval_ref", new_random_dseq_stats_eval_ref)
                   (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
-                    |> Lazy_Sequence.map (apfst proc))
+                    |> Lazy_Sequence.mapa (apfst proc))
                     thy t' [] nrandom size seed depth))))
             else rpair NONE
               (fst (Lazy_Sequence.yieldn k
                 (Code_Eval.eval NONE
                   ("Predicate_Compile_Core.new_random_dseq_eval_ref", new_random_dseq_eval_ref)
                   (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
-                    |> Lazy_Sequence.map proc)
+                    |> Lazy_Sequence.mapa proc)
                     thy t' [] nrandom size seed depth)))
           end
       | _ =>
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -267,7 +267,7 @@
               Code_Eval.eval (SOME target)
                 ("Predicate_Compile_Quickcheck.new_test_ref", new_test_ref)
                 (fn proc => fn g => fn nrandom => fn size => fn s => fn depth =>
-                  g nrandom size s depth |> (Lazy_Sequence.map o map) proc)
+                  g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc)
                   thy4 qc_term []
           in
             fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield 
--- a/src/HOL/Tools/Qelim/cooper.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Qelim/cooper.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -536,7 +536,7 @@
 structure Coopereif =
 struct
 
-open GeneratedCooper;
+open Generated_Cooper;
 
 fun cooper s = raise Cooper.COOPER ("Cooper oracle failed", ERROR s);
 fun i_of_term vs t = case t
--- a/src/HOL/Tools/Qelim/generated_cooper.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Tools/Qelim/generated_cooper.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -1,49 +1,263 @@
-(*  Title:      HOL/Tools/Qelim/generated_cooper.ML
+(* Generated from Cooper.thy; DO NOT EDIT! *)
 
-This file is generated from HOL/Decision_Procs/Cooper.thy.  DO NOT EDIT.
-*)
-
-structure GeneratedCooper = 
-struct
+structure Generated_Cooper : sig
+  type 'a eq
+  val eq : 'a eq -> 'a -> 'a -> bool
+  val eqa : 'a eq -> 'a -> 'a -> bool
+  val leta : 'a -> ('a -> 'b) -> 'b
+  val suc : IntInf.int -> IntInf.int
+  datatype num = C of IntInf.int | Bound of IntInf.int |
+    Cn of IntInf.int * IntInf.int * num | Neg of num | Add of num * num |
+    Sub of num * num | Mul of IntInf.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 IntInf.int * num | NDvd of IntInf.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 IntInf.int |
+    NClosed of IntInf.int
+  val map : ('a -> 'b) -> 'a list -> 'b list
+  val append : 'a list -> 'a list -> 'a list
+  val disjuncts : fm -> fm list
+  val fm_case :
+    'a -> 'a -> (num -> 'a) ->
+                  (num -> 'a) ->
+                    (num -> 'a) ->
+                      (num -> 'a) ->
+                        (num -> 'a) ->
+                          (num -> 'a) ->
+                            (IntInf.int -> num -> 'a) ->
+                              (IntInf.int -> num -> 'a) ->
+                                (fm -> 'a) ->
+                                  (fm -> fm -> 'a) ->
+                                    (fm -> fm -> 'a) ->
+                                      (fm -> fm -> 'a) ->
+(fm -> fm -> 'a) ->
+  (fm -> 'a) ->
+    (fm -> 'a) -> (IntInf.int -> 'a) -> (IntInf.int -> 'a) -> fm -> 'a
+  val eq_num : num -> num -> bool
+  val eq_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 dj : (fm -> fm) -> fm -> fm
+  val disj : fm -> fm -> fm
+  val minus_nat : IntInf.int -> IntInf.int -> IntInf.int
+  val decrnum : num -> num
+  val decr : fm -> fm
+  val concat_map : ('a -> 'b list) -> 'a list -> 'b list
+  val numsubst0 : num -> num -> num
+  val subst0 : num -> fm -> fm
+  val minusinf : fm -> fm
+  val eq_int : IntInf.int eq
+  val zero_int : IntInf.int
+  type 'a zero
+  val zero : 'a zero -> 'a
+  val zero_inta : IntInf.int zero
+  type 'a times
+  val times : 'a times -> 'a -> 'a -> '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
+  val times_int : IntInf.int times
+  val no_zero_divisors_int : IntInf.int no_zero_divisors
+  type 'a one
+  val one : 'a one -> 'a
+  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 semigroup_mult
+  val times_semigroup_mult : 'a semigroup_mult -> 'a times
+  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 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 monoid_add
+  val semigroup_add_monoid_add : 'a monoid_add -> 'a semigroup_add
+  val zero_monoid_add : 'a monoid_add -> 'a zero
+  type 'a comm_monoid_add
+  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 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 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 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 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 dvd
+  val times_dvd : 'a dvd -> 'a times
+  type 'a ab_semigroup_mult
+  val semigroup_mult_ab_semigroup_mult :
+    'a ab_semigroup_mult -> 'a semigroup_mult
+  type 'a comm_semiring
+  val ab_semigroup_mult_comm_semiring : 'a comm_semiring -> 'a ab_semigroup_mult
+  val semiring_comm_semiring : 'a comm_semiring -> 'a semiring
+  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_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 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 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
+  val comm_semiring_1_comm_semiring_1_cancel :
+    '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 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 one_int : IntInf.int
+  val one_inta : IntInf.int one
+  val zero_neq_one_int : IntInf.int zero_neq_one
+  val semigroup_mult_int : IntInf.int semigroup_mult
+  val plus_int : IntInf.int plus
+  val semigroup_add_int : IntInf.int semigroup_add
+  val ab_semigroup_add_int : IntInf.int ab_semigroup_add
+  val semiring_int : IntInf.int semiring
+  val mult_zero_int : IntInf.int mult_zero
+  val monoid_add_int : IntInf.int monoid_add
+  val comm_monoid_add_int : IntInf.int comm_monoid_add
+  val semiring_0_int : IntInf.int semiring_0
+  val power_int : IntInf.int power
+  val monoid_mult_int : IntInf.int monoid_mult
+  val semiring_1_int : IntInf.int semiring_1
+  val cancel_semigroup_add_int : IntInf.int cancel_semigroup_add
+  val cancel_ab_semigroup_add_int : IntInf.int cancel_ab_semigroup_add
+  val cancel_comm_monoid_add_int : IntInf.int cancel_comm_monoid_add
+  val semiring_0_cancel_int : IntInf.int semiring_0_cancel
+  val semiring_1_cancel_int : IntInf.int semiring_1_cancel
+  val dvd_int : IntInf.int dvd
+  val ab_semigroup_mult_int : IntInf.int ab_semigroup_mult
+  val comm_semiring_int : IntInf.int comm_semiring
+  val comm_semiring_0_int : IntInf.int comm_semiring_0
+  val comm_monoid_mult_int : IntInf.int comm_monoid_mult
+  val comm_semiring_1_int : IntInf.int comm_semiring_1
+  val comm_semiring_0_cancel_int : IntInf.int comm_semiring_0_cancel
+  val comm_semiring_1_cancel_int : IntInf.int comm_semiring_1_cancel
+  val abs_int : IntInf.int -> IntInf.int
+  val split : ('a -> 'b -> 'c) -> 'a * 'b -> 'c
+  val sgn_int : IntInf.int -> IntInf.int
+  val apsnd : ('a -> 'b) -> 'c * 'a -> 'c * 'b
+  val divmod_int : IntInf.int -> IntInf.int -> IntInf.int * IntInf.int
+  val snd : 'a * 'b -> 'b
+  val mod_int : IntInf.int -> IntInf.int -> IntInf.int
+  val fst : 'a * 'b -> 'a
+  val div_int : IntInf.int -> IntInf.int -> IntInf.int
+  val div_inta : IntInf.int diva
+  val semiring_div_int : IntInf.int semiring_div
+  val dvd : 'a semiring_div * 'a eq -> 'a -> 'a -> bool
+  val num_case :
+    (IntInf.int -> 'a) ->
+      (IntInf.int -> 'a) ->
+        (IntInf.int -> IntInf.int -> num -> 'a) ->
+          (num -> 'a) ->
+            (num -> num -> 'a) ->
+              (num -> num -> 'a) -> (IntInf.int -> num -> 'a) -> num -> 'a
+  val nummul : IntInf.int -> num -> num
+  val numneg : num -> num
+  val numadd : num * num -> num
+  val numsub : num -> num -> num
+  val simpnum : num -> num
+  val nota : fm -> fm
+  val iffa : fm -> fm -> fm
+  val impa : fm -> fm -> fm
+  val conj : fm -> fm -> fm
+  val simpfm : fm -> fm
+  val iupt : IntInf.int -> IntInf.int -> IntInf.int list
+  val mirror : fm -> fm
+  val size_list : 'a list -> IntInf.int
+  val alpha : fm -> num list
+  val beta : fm -> num list
+  val eq_numa : num eq
+  val member : 'a eq -> 'a -> 'a list -> bool
+  val remdups : 'a eq -> 'a list -> 'a list
+  val gcd_int : IntInf.int -> IntInf.int -> IntInf.int
+  val lcm_int : IntInf.int -> IntInf.int -> IntInf.int
+  val delta : fm -> IntInf.int
+  val a_beta : fm -> IntInf.int -> fm
+  val zeta : fm -> IntInf.int
+  val zsplit0 : num -> IntInf.int * num
+  val zlfm : fm -> fm
+  val unita : fm -> fm * (num list * IntInf.int)
+  val cooper : fm -> fm
+  val prep : fm -> fm
+  val qelim : fm -> (fm -> fm) -> fm
+  val pa : fm -> fm
+end = struct
 
 type 'a eq = {eq : 'a -> 'a -> bool};
-fun eq (A_:'a eq) = #eq A_;
-
-val eq_nat = {eq = (fn a => fn b => ((a : IntInf.int) = b))} : IntInf.int eq;
-
-fun eqop A_ a b = eq A_ a b;
-
-fun divmod n m = (if eqop eq_nat m 0 then (0, n) else IntInf.divMod (n, m));
-
-fun snd (a, b) = b;
+val eq = #eq : 'a eq -> 'a -> 'a -> bool;
 
-fun mod_nat m n = snd (divmod m n);
-
-fun gcd m n = (if eqop eq_nat n 0 then m else gcd n (mod_nat m n));
-
-fun fst (a, b) = a;
-
-fun div_nat m n = fst (divmod m n);
-
-fun lcm m n = div_nat (IntInf.* (m, n)) (gcd m n);
+fun eqa A_ a b = eq A_ a b;
 
 fun leta s f = f s;
 
-fun suc n = IntInf.+ (n, 1);
-
-datatype num = Mul of IntInf.int * num | Sub of num * num | Add of num * num |
-  Neg of num | Cn of IntInf.int * IntInf.int * num | Bound of IntInf.int |
-  C of IntInf.int;
+fun suc n = IntInf.+ (n, (1 : IntInf.int));
 
-datatype fm = NClosed of IntInf.int | Closed of IntInf.int | A of fm | E of fm |
-  Iff of fm * fm | Imp of fm * fm | Or of fm * fm | And of fm * fm | Not of fm |
-  NDvd of IntInf.int * num | Dvd of IntInf.int * num | NEq of num | Eq of num |
-  Ge of num | Gt of num | Le of num | Lt of num | F | T;
+datatype num = C of IntInf.int | Bound of IntInf.int |
+  Cn of IntInf.int * IntInf.int * num | Neg of num | Add of num * num |
+  Sub of num * num | Mul of IntInf.int * num;
 
-fun abs_int i = (if IntInf.< (i, (0 : IntInf.int)) then IntInf.~ i else i);
-
-fun zlcm i j =
-  (lcm (IntInf.max (0, (abs_int i))) (IntInf.max (0, (abs_int j))));
+datatype fm = T | F | Lt of num | Le of num | Gt of num | Ge of num | Eq of num
+  | NEq of num | Dvd of IntInf.int * num | NDvd of IntInf.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 IntInf.int | NClosed of IntInf.int;
 
 fun map f [] = []
   | map f (x :: xs) = f x :: map f xs;
@@ -110,449 +324,441 @@
   | fm_case f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 T
     = f1;
 
-fun eq_num (Mul (c, d)) (Sub (a, b)) = false
-  | eq_num (Mul (c, d)) (Add (a, b)) = false
-  | eq_num (Sub (c, d)) (Add (a, b)) = false
-  | eq_num (Mul (b, c)) (Neg a) = false
-  | eq_num (Sub (b, c)) (Neg a) = false
-  | eq_num (Add (b, c)) (Neg a) = false
-  | eq_num (Mul (d, e)) (Cn (a, b, c)) = false
-  | eq_num (Sub (d, e)) (Cn (a, b, c)) = false
-  | eq_num (Add (d, e)) (Cn (a, b, c)) = false
-  | eq_num (Neg d) (Cn (a, b, c)) = false
-  | eq_num (Mul (b, c)) (Bound a) = false
-  | eq_num (Sub (b, c)) (Bound a) = false
-  | eq_num (Add (b, c)) (Bound a) = false
-  | eq_num (Neg b) (Bound a) = false
-  | eq_num (Cn (b, c, d)) (Bound a) = false
-  | eq_num (Mul (b, c)) (C a) = false
-  | eq_num (Sub (b, c)) (C a) = false
-  | eq_num (Add (b, c)) (C a) = false
-  | eq_num (Neg b) (C a) = false
-  | eq_num (Cn (b, c, d)) (C a) = false
-  | eq_num (Bound b) (C a) = false
-  | eq_num (Sub (a, b)) (Mul (c, d)) = false
-  | eq_num (Add (a, b)) (Mul (c, d)) = false
-  | eq_num (Add (a, b)) (Sub (c, d)) = false
-  | eq_num (Neg a) (Mul (b, c)) = false
-  | eq_num (Neg a) (Sub (b, c)) = false
-  | eq_num (Neg a) (Add (b, c)) = false
-  | eq_num (Cn (a, b, c)) (Mul (d, e)) = false
-  | eq_num (Cn (a, b, c)) (Sub (d, e)) = false
-  | eq_num (Cn (a, b, c)) (Add (d, e)) = false
-  | eq_num (Cn (a, b, c)) (Neg d) = false
-  | eq_num (Bound a) (Mul (b, c)) = false
-  | eq_num (Bound a) (Sub (b, c)) = false
-  | eq_num (Bound a) (Add (b, c)) = false
-  | eq_num (Bound a) (Neg b) = false
-  | eq_num (Bound a) (Cn (b, c, d)) = false
-  | eq_num (C a) (Mul (b, c)) = false
-  | eq_num (C a) (Sub (b, c)) = false
-  | eq_num (C a) (Add (b, c)) = false
-  | eq_num (C a) (Neg b) = false
-  | eq_num (C a) (Cn (b, c, d)) = false
-  | eq_num (C a) (Bound b) = false
-  | eq_num (Mul (inta, num)) (Mul (int', num')) =
-    ((inta : IntInf.int) = int') andalso eq_num num num'
-  | eq_num (Sub (num1, num2)) (Sub (num1', num2')) =
-    eq_num num1 num1' andalso eq_num num2 num2'
-  | eq_num (Add (num1, num2)) (Add (num1', num2')) =
-    eq_num num1 num1' andalso eq_num num2 num2'
-  | eq_num (Neg num) (Neg num') = eq_num num num'
-  | eq_num (Cn (nat, inta, num)) (Cn (nat', int', num')) =
-    ((nat : IntInf.int) = nat') andalso
-      (((inta : IntInf.int) = int') andalso eq_num num num')
-  | eq_num (Bound nat) (Bound nat') = ((nat : IntInf.int) = nat')
-  | eq_num (C inta) (C int') = ((inta : IntInf.int) = int');
+fun eq_num (C intaa) (C inta) = ((intaa : IntInf.int) = inta)
+  | eq_num (Bound nata) (Bound nat) = ((nata : IntInf.int) = nat)
+  | eq_num (Cn (nata, intaa, numa)) (Cn (nat, inta, num)) =
+    ((nata : IntInf.int) = nat) andalso
+      (((intaa : IntInf.int) = inta) andalso eq_num numa num)
+  | eq_num (Neg numa) (Neg num) = eq_num numa num
+  | eq_num (Add (num1a, num2a)) (Add (num1, num2)) =
+    eq_num num1a num1 andalso eq_num num2a num2
+  | eq_num (Sub (num1a, num2a)) (Sub (num1, num2)) =
+    eq_num num1a num1 andalso eq_num num2a num2
+  | eq_num (Mul (intaa, numa)) (Mul (inta, num)) =
+    ((intaa : IntInf.int) = inta) andalso eq_num numa num
+  | eq_num (C inta) (Bound nat) = false
+  | eq_num (Bound nat) (C inta) = false
+  | eq_num (C intaa) (Cn (nat, inta, num)) = false
+  | eq_num (Cn (nat, intaa, num)) (C inta) = false
+  | eq_num (C inta) (Neg num) = false
+  | eq_num (Neg num) (C inta) = false
+  | eq_num (C inta) (Add (num1, num2)) = false
+  | eq_num (Add (num1, num2)) (C inta) = false
+  | eq_num (C inta) (Sub (num1, num2)) = false
+  | eq_num (Sub (num1, num2)) (C inta) = false
+  | eq_num (C intaa) (Mul (inta, num)) = false
+  | eq_num (Mul (intaa, num)) (C inta) = false
+  | eq_num (Bound nata) (Cn (nat, inta, num)) = false
+  | eq_num (Cn (nata, inta, num)) (Bound nat) = false
+  | eq_num (Bound nat) (Neg num) = false
+  | eq_num (Neg num) (Bound nat) = false
+  | eq_num (Bound nat) (Add (num1, num2)) = false
+  | eq_num (Add (num1, num2)) (Bound nat) = false
+  | eq_num (Bound nat) (Sub (num1, num2)) = false
+  | eq_num (Sub (num1, num2)) (Bound nat) = false
+  | eq_num (Bound nat) (Mul (inta, num)) = false
+  | eq_num (Mul (inta, num)) (Bound nat) = false
+  | eq_num (Cn (nat, inta, numa)) (Neg num) = false
+  | eq_num (Neg numa) (Cn (nat, inta, num)) = false
+  | eq_num (Cn (nat, inta, num)) (Add (num1, num2)) = false
+  | eq_num (Add (num1, num2)) (Cn (nat, inta, num)) = false
+  | eq_num (Cn (nat, inta, num)) (Sub (num1, num2)) = false
+  | eq_num (Sub (num1, num2)) (Cn (nat, inta, num)) = false
+  | eq_num (Cn (nat, intaa, numa)) (Mul (inta, num)) = false
+  | eq_num (Mul (intaa, numa)) (Cn (nat, inta, num)) = false
+  | eq_num (Neg num) (Add (num1, num2)) = false
+  | eq_num (Add (num1, num2)) (Neg num) = false
+  | eq_num (Neg num) (Sub (num1, num2)) = false
+  | eq_num (Sub (num1, num2)) (Neg num) = false
+  | eq_num (Neg numa) (Mul (inta, num)) = false
+  | eq_num (Mul (inta, numa)) (Neg num) = false
+  | eq_num (Add (num1a, num2a)) (Sub (num1, num2)) = false
+  | eq_num (Sub (num1a, num2a)) (Add (num1, num2)) = false
+  | eq_num (Add (num1, num2)) (Mul (inta, num)) = false
+  | eq_num (Mul (inta, num)) (Add (num1, num2)) = false
+  | eq_num (Sub (num1, num2)) (Mul (inta, num)) = false
+  | eq_num (Mul (inta, num)) (Sub (num1, num2)) = false;
 
-fun eq_fm (NClosed b) (Closed a) = false
-  | eq_fm (NClosed b) (A a) = false
-  | eq_fm (Closed b) (A a) = false
-  | eq_fm (NClosed b) (E a) = false
-  | eq_fm (Closed b) (E a) = false
-  | eq_fm (A b) (E a) = false
-  | eq_fm (NClosed c) (Iff (a, b)) = false
-  | eq_fm (Closed c) (Iff (a, b)) = false
-  | eq_fm (A c) (Iff (a, b)) = false
-  | eq_fm (E c) (Iff (a, b)) = false
-  | eq_fm (NClosed c) (Imp (a, b)) = false
-  | eq_fm (Closed c) (Imp (a, b)) = false
-  | eq_fm (A c) (Imp (a, b)) = false
-  | eq_fm (E c) (Imp (a, b)) = false
-  | eq_fm (Iff (c, d)) (Imp (a, b)) = false
-  | eq_fm (NClosed c) (Or (a, b)) = false
-  | eq_fm (Closed c) (Or (a, b)) = false
-  | eq_fm (A c) (Or (a, b)) = false
-  | eq_fm (E c) (Or (a, b)) = false
-  | eq_fm (Iff (c, d)) (Or (a, b)) = false
-  | eq_fm (Imp (c, d)) (Or (a, b)) = false
-  | eq_fm (NClosed c) (And (a, b)) = false
-  | eq_fm (Closed c) (And (a, b)) = false
-  | eq_fm (A c) (And (a, b)) = false
-  | eq_fm (E c) (And (a, b)) = false
-  | eq_fm (Iff (c, d)) (And (a, b)) = false
-  | eq_fm (Imp (c, d)) (And (a, b)) = false
-  | eq_fm (Or (c, d)) (And (a, b)) = false
-  | eq_fm (NClosed b) (Not a) = false
-  | eq_fm (Closed b) (Not a) = false
-  | eq_fm (A b) (Not a) = false
-  | eq_fm (E b) (Not a) = false
-  | eq_fm (Iff (b, c)) (Not a) = false
-  | eq_fm (Imp (b, c)) (Not a) = false
-  | eq_fm (Or (b, c)) (Not a) = false
-  | eq_fm (And (b, c)) (Not a) = false
-  | eq_fm (NClosed c) (NDvd (a, b)) = false
-  | eq_fm (Closed c) (NDvd (a, b)) = false
-  | eq_fm (A c) (NDvd (a, b)) = false
-  | eq_fm (E c) (NDvd (a, b)) = false
-  | eq_fm (Iff (c, d)) (NDvd (a, b)) = false
-  | eq_fm (Imp (c, d)) (NDvd (a, b)) = false
-  | eq_fm (Or (c, d)) (NDvd (a, b)) = false
-  | eq_fm (And (c, d)) (NDvd (a, b)) = false
-  | eq_fm (Not c) (NDvd (a, b)) = false
-  | eq_fm (NClosed c) (Dvd (a, b)) = false
-  | eq_fm (Closed c) (Dvd (a, b)) = false
-  | eq_fm (A c) (Dvd (a, b)) = false
-  | eq_fm (E c) (Dvd (a, b)) = false
-  | eq_fm (Iff (c, d)) (Dvd (a, b)) = false
-  | eq_fm (Imp (c, d)) (Dvd (a, b)) = false
-  | eq_fm (Or (c, d)) (Dvd (a, b)) = false
-  | eq_fm (And (c, d)) (Dvd (a, b)) = false
-  | eq_fm (Not c) (Dvd (a, b)) = false
-  | eq_fm (NDvd (c, d)) (Dvd (a, b)) = false
-  | eq_fm (NClosed b) (NEq a) = false
-  | eq_fm (Closed b) (NEq a) = false
-  | eq_fm (A b) (NEq a) = false
-  | eq_fm (E b) (NEq a) = false
-  | eq_fm (Iff (b, c)) (NEq a) = false
-  | eq_fm (Imp (b, c)) (NEq a) = false
-  | eq_fm (Or (b, c)) (NEq a) = false
-  | eq_fm (And (b, c)) (NEq a) = false
-  | eq_fm (Not b) (NEq a) = false
-  | eq_fm (NDvd (b, c)) (NEq a) = false
-  | eq_fm (Dvd (b, c)) (NEq a) = false
-  | eq_fm (NClosed b) (Eq a) = false
-  | eq_fm (Closed b) (Eq a) = false
-  | eq_fm (A b) (Eq a) = false
-  | eq_fm (E b) (Eq a) = false
-  | eq_fm (Iff (b, c)) (Eq a) = false
-  | eq_fm (Imp (b, c)) (Eq a) = false
-  | eq_fm (Or (b, c)) (Eq a) = false
-  | eq_fm (And (b, c)) (Eq a) = false
-  | eq_fm (Not b) (Eq a) = false
-  | eq_fm (NDvd (b, c)) (Eq a) = false
-  | eq_fm (Dvd (b, c)) (Eq a) = false
-  | eq_fm (NEq b) (Eq a) = false
-  | eq_fm (NClosed b) (Ge a) = false
-  | eq_fm (Closed b) (Ge a) = false
-  | eq_fm (A b) (Ge a) = false
-  | eq_fm (E b) (Ge a) = false
-  | eq_fm (Iff (b, c)) (Ge a) = false
-  | eq_fm (Imp (b, c)) (Ge a) = false
-  | eq_fm (Or (b, c)) (Ge a) = false
-  | eq_fm (And (b, c)) (Ge a) = false
-  | eq_fm (Not b) (Ge a) = false
-  | eq_fm (NDvd (b, c)) (Ge a) = false
-  | eq_fm (Dvd (b, c)) (Ge a) = false
-  | eq_fm (NEq b) (Ge a) = false
-  | eq_fm (Eq b) (Ge a) = false
-  | eq_fm (NClosed b) (Gt a) = false
-  | eq_fm (Closed b) (Gt a) = false
-  | eq_fm (A b) (Gt a) = false
-  | eq_fm (E b) (Gt a) = false
-  | eq_fm (Iff (b, c)) (Gt a) = false
-  | eq_fm (Imp (b, c)) (Gt a) = false
-  | eq_fm (Or (b, c)) (Gt a) = false
-  | eq_fm (And (b, c)) (Gt a) = false
-  | eq_fm (Not b) (Gt a) = false
-  | eq_fm (NDvd (b, c)) (Gt a) = false
-  | eq_fm (Dvd (b, c)) (Gt a) = false
-  | eq_fm (NEq b) (Gt a) = false
-  | eq_fm (Eq b) (Gt a) = false
-  | eq_fm (Ge b) (Gt a) = false
-  | eq_fm (NClosed b) (Le a) = false
-  | eq_fm (Closed b) (Le a) = false
-  | eq_fm (A b) (Le a) = false
-  | eq_fm (E b) (Le a) = false
-  | eq_fm (Iff (b, c)) (Le a) = false
-  | eq_fm (Imp (b, c)) (Le a) = false
-  | eq_fm (Or (b, c)) (Le a) = false
-  | eq_fm (And (b, c)) (Le a) = false
-  | eq_fm (Not b) (Le a) = false
-  | eq_fm (NDvd (b, c)) (Le a) = false
-  | eq_fm (Dvd (b, c)) (Le a) = false
-  | eq_fm (NEq b) (Le a) = false
-  | eq_fm (Eq b) (Le a) = false
-  | eq_fm (Ge b) (Le a) = false
-  | eq_fm (Gt b) (Le a) = false
-  | eq_fm (NClosed b) (Lt a) = false
-  | eq_fm (Closed b) (Lt a) = false
-  | eq_fm (A b) (Lt a) = false
-  | eq_fm (E b) (Lt a) = false
-  | eq_fm (Iff (b, c)) (Lt a) = false
-  | eq_fm (Imp (b, c)) (Lt a) = false
-  | eq_fm (Or (b, c)) (Lt a) = false
-  | eq_fm (And (b, c)) (Lt a) = false
-  | eq_fm (Not b) (Lt a) = false
-  | eq_fm (NDvd (b, c)) (Lt a) = false
-  | eq_fm (Dvd (b, c)) (Lt a) = false
-  | eq_fm (NEq b) (Lt a) = false
-  | eq_fm (Eq b) (Lt a) = false
-  | eq_fm (Ge b) (Lt a) = false
-  | eq_fm (Gt b) (Lt a) = false
-  | eq_fm (Le b) (Lt a) = false
-  | eq_fm (NClosed a) F = false
-  | eq_fm (Closed a) F = false
-  | eq_fm (A a) F = false
-  | eq_fm (E a) F = false
-  | eq_fm (Iff (a, b)) F = false
-  | eq_fm (Imp (a, b)) F = false
-  | eq_fm (Or (a, b)) F = false
-  | eq_fm (And (a, b)) F = false
-  | eq_fm (Not a) F = false
-  | eq_fm (NDvd (a, b)) F = false
-  | eq_fm (Dvd (a, b)) F = false
-  | eq_fm (NEq a) F = false
-  | eq_fm (Eq a) F = false
-  | eq_fm (Ge a) F = false
-  | eq_fm (Gt a) F = false
-  | eq_fm (Le a) F = false
-  | eq_fm (Lt a) F = false
-  | eq_fm (NClosed a) T = false
-  | eq_fm (Closed a) T = false
-  | eq_fm (A a) T = false
-  | eq_fm (E a) T = false
-  | eq_fm (Iff (a, b)) T = false
-  | eq_fm (Imp (a, b)) T = false
-  | eq_fm (Or (a, b)) T = false
-  | eq_fm (And (a, b)) T = false
-  | eq_fm (Not a) T = false
-  | eq_fm (NDvd (a, b)) T = false
-  | eq_fm (Dvd (a, b)) T = false
-  | eq_fm (NEq a) T = false
-  | eq_fm (Eq a) T = false
-  | eq_fm (Ge a) T = false
-  | eq_fm (Gt a) T = false
-  | eq_fm (Le a) T = false
-  | eq_fm (Lt a) T = false
+fun eq_fm T T = true
+  | eq_fm F F = true
+  | eq_fm (Lt numa) (Lt num) = eq_num numa num
+  | eq_fm (Le numa) (Le num) = eq_num numa num
+  | eq_fm (Gt numa) (Gt num) = eq_num numa num
+  | eq_fm (Ge numa) (Ge num) = eq_num numa num
+  | eq_fm (Eq numa) (Eq num) = eq_num numa num
+  | eq_fm (NEq numa) (NEq num) = eq_num numa num
+  | eq_fm (Dvd (intaa, numa)) (Dvd (inta, num)) =
+    ((intaa : IntInf.int) = inta) andalso eq_num numa num
+  | eq_fm (NDvd (intaa, numa)) (NDvd (inta, num)) =
+    ((intaa : IntInf.int) = inta) andalso eq_num numa num
+  | eq_fm (Not fma) (Not fm) = eq_fm fma fm
+  | eq_fm (And (fm1a, fm2a)) (And (fm1, fm2)) =
+    eq_fm fm1a fm1 andalso eq_fm fm2a fm2
+  | eq_fm (Or (fm1a, fm2a)) (Or (fm1, fm2)) =
+    eq_fm fm1a fm1 andalso eq_fm fm2a fm2
+  | eq_fm (Imp (fm1a, fm2a)) (Imp (fm1, fm2)) =
+    eq_fm fm1a fm1 andalso eq_fm fm2a fm2
+  | eq_fm (Iff (fm1a, fm2a)) (Iff (fm1, fm2)) =
+    eq_fm fm1a fm1 andalso eq_fm fm2a fm2
+  | eq_fm (E fma) (E fm) = eq_fm fma fm
+  | eq_fm (A fma) (A fm) = eq_fm fma fm
+  | eq_fm (Closed nata) (Closed nat) = ((nata : IntInf.int) = nat)
+  | eq_fm (NClosed nata) (NClosed nat) = ((nata : IntInf.int) = nat)
+  | eq_fm T F = false
   | eq_fm F T = false
-  | eq_fm (Closed a) (NClosed b) = false
-  | eq_fm (A a) (NClosed b) = false
-  | eq_fm (A a) (Closed b) = false
-  | eq_fm (E a) (NClosed b) = false
-  | eq_fm (E a) (Closed b) = false
-  | eq_fm (E a) (A b) = false
-  | eq_fm (Iff (a, b)) (NClosed c) = false
-  | eq_fm (Iff (a, b)) (Closed c) = false
-  | eq_fm (Iff (a, b)) (A c) = false
-  | eq_fm (Iff (a, b)) (E c) = false
-  | eq_fm (Imp (a, b)) (NClosed c) = false
-  | eq_fm (Imp (a, b)) (Closed c) = false
-  | eq_fm (Imp (a, b)) (A c) = false
-  | eq_fm (Imp (a, b)) (E c) = false
-  | eq_fm (Imp (a, b)) (Iff (c, d)) = false
-  | eq_fm (Or (a, b)) (NClosed c) = false
-  | eq_fm (Or (a, b)) (Closed c) = false
-  | eq_fm (Or (a, b)) (A c) = false
-  | eq_fm (Or (a, b)) (E c) = false
-  | eq_fm (Or (a, b)) (Iff (c, d)) = false
-  | eq_fm (Or (a, b)) (Imp (c, d)) = false
-  | eq_fm (And (a, b)) (NClosed c) = false
-  | eq_fm (And (a, b)) (Closed c) = false
-  | eq_fm (And (a, b)) (A c) = false
-  | eq_fm (And (a, b)) (E c) = false
-  | eq_fm (And (a, b)) (Iff (c, d)) = false
-  | eq_fm (And (a, b)) (Imp (c, d)) = false
-  | eq_fm (And (a, b)) (Or (c, d)) = false
-  | eq_fm (Not a) (NClosed b) = false
-  | eq_fm (Not a) (Closed b) = false
-  | eq_fm (Not a) (A b) = false
-  | eq_fm (Not a) (E b) = false
-  | eq_fm (Not a) (Iff (b, c)) = false
-  | eq_fm (Not a) (Imp (b, c)) = false
-  | eq_fm (Not a) (Or (b, c)) = false
-  | eq_fm (Not a) (And (b, c)) = false
-  | eq_fm (NDvd (a, b)) (NClosed c) = false
-  | eq_fm (NDvd (a, b)) (Closed c) = false
-  | eq_fm (NDvd (a, b)) (A c) = false
-  | eq_fm (NDvd (a, b)) (E c) = false
-  | eq_fm (NDvd (a, b)) (Iff (c, d)) = false
-  | eq_fm (NDvd (a, b)) (Imp (c, d)) = false
-  | eq_fm (NDvd (a, b)) (Or (c, d)) = false
-  | eq_fm (NDvd (a, b)) (And (c, d)) = false
-  | eq_fm (NDvd (a, b)) (Not c) = false
-  | eq_fm (Dvd (a, b)) (NClosed c) = false
-  | eq_fm (Dvd (a, b)) (Closed c) = false
-  | eq_fm (Dvd (a, b)) (A c) = false
-  | eq_fm (Dvd (a, b)) (E c) = false
-  | eq_fm (Dvd (a, b)) (Iff (c, d)) = false
-  | eq_fm (Dvd (a, b)) (Imp (c, d)) = false
-  | eq_fm (Dvd (a, b)) (Or (c, d)) = false
-  | eq_fm (Dvd (a, b)) (And (c, d)) = false
-  | eq_fm (Dvd (a, b)) (Not c) = false
-  | eq_fm (Dvd (a, b)) (NDvd (c, d)) = false
-  | eq_fm (NEq a) (NClosed b) = false
-  | eq_fm (NEq a) (Closed b) = false
-  | eq_fm (NEq a) (A b) = false
-  | eq_fm (NEq a) (E b) = false
-  | eq_fm (NEq a) (Iff (b, c)) = false
-  | eq_fm (NEq a) (Imp (b, c)) = false
-  | eq_fm (NEq a) (Or (b, c)) = false
-  | eq_fm (NEq a) (And (b, c)) = false
-  | eq_fm (NEq a) (Not b) = false
-  | eq_fm (NEq a) (NDvd (b, c)) = false
-  | eq_fm (NEq a) (Dvd (b, c)) = false
-  | eq_fm (Eq a) (NClosed b) = false
-  | eq_fm (Eq a) (Closed b) = false
-  | eq_fm (Eq a) (A b) = false
-  | eq_fm (Eq a) (E b) = false
-  | eq_fm (Eq a) (Iff (b, c)) = false
-  | eq_fm (Eq a) (Imp (b, c)) = false
-  | eq_fm (Eq a) (Or (b, c)) = false
-  | eq_fm (Eq a) (And (b, c)) = false
-  | eq_fm (Eq a) (Not b) = false
-  | eq_fm (Eq a) (NDvd (b, c)) = false
-  | eq_fm (Eq a) (Dvd (b, c)) = false
-  | eq_fm (Eq a) (NEq b) = false
-  | eq_fm (Ge a) (NClosed b) = false
-  | eq_fm (Ge a) (Closed b) = false
-  | eq_fm (Ge a) (A b) = false
-  | eq_fm (Ge a) (E b) = false
-  | eq_fm (Ge a) (Iff (b, c)) = false
-  | eq_fm (Ge a) (Imp (b, c)) = false
-  | eq_fm (Ge a) (Or (b, c)) = false
-  | eq_fm (Ge a) (And (b, c)) = false
-  | eq_fm (Ge a) (Not b) = false
-  | eq_fm (Ge a) (NDvd (b, c)) = false
-  | eq_fm (Ge a) (Dvd (b, c)) = false
-  | eq_fm (Ge a) (NEq b) = false
-  | eq_fm (Ge a) (Eq b) = false
-  | eq_fm (Gt a) (NClosed b) = false
-  | eq_fm (Gt a) (Closed b) = false
-  | eq_fm (Gt a) (A b) = false
-  | eq_fm (Gt a) (E b) = false
-  | eq_fm (Gt a) (Iff (b, c)) = false
-  | eq_fm (Gt a) (Imp (b, c)) = false
-  | eq_fm (Gt a) (Or (b, c)) = false
-  | eq_fm (Gt a) (And (b, c)) = false
-  | eq_fm (Gt a) (Not b) = false
-  | eq_fm (Gt a) (NDvd (b, c)) = false
-  | eq_fm (Gt a) (Dvd (b, c)) = false
-  | eq_fm (Gt a) (NEq b) = false
-  | eq_fm (Gt a) (Eq b) = false
-  | eq_fm (Gt a) (Ge b) = false
-  | eq_fm (Le a) (NClosed b) = false
-  | eq_fm (Le a) (Closed b) = false
-  | eq_fm (Le a) (A b) = false
-  | eq_fm (Le a) (E b) = false
-  | eq_fm (Le a) (Iff (b, c)) = false
-  | eq_fm (Le a) (Imp (b, c)) = false
-  | eq_fm (Le a) (Or (b, c)) = false
-  | eq_fm (Le a) (And (b, c)) = false
-  | eq_fm (Le a) (Not b) = false
-  | eq_fm (Le a) (NDvd (b, c)) = false
-  | eq_fm (Le a) (Dvd (b, c)) = false
-  | eq_fm (Le a) (NEq b) = false
-  | eq_fm (Le a) (Eq b) = false
-  | eq_fm (Le a) (Ge b) = false
-  | eq_fm (Le a) (Gt b) = false
-  | eq_fm (Lt a) (NClosed b) = false
-  | eq_fm (Lt a) (Closed b) = false
-  | eq_fm (Lt a) (A b) = false
-  | eq_fm (Lt a) (E b) = false
-  | eq_fm (Lt a) (Iff (b, c)) = false
-  | eq_fm (Lt a) (Imp (b, c)) = false
-  | eq_fm (Lt a) (Or (b, c)) = false
-  | eq_fm (Lt a) (And (b, c)) = false
-  | eq_fm (Lt a) (Not b) = false
-  | eq_fm (Lt a) (NDvd (b, c)) = false
-  | eq_fm (Lt a) (Dvd (b, c)) = false
-  | eq_fm (Lt a) (NEq b) = false
-  | eq_fm (Lt a) (Eq b) = false
-  | eq_fm (Lt a) (Ge b) = false
-  | eq_fm (Lt a) (Gt b) = false
-  | eq_fm (Lt a) (Le b) = false
-  | eq_fm F (NClosed a) = false
-  | eq_fm F (Closed a) = false
-  | eq_fm F (A a) = false
-  | eq_fm F (E a) = false
-  | eq_fm F (Iff (a, b)) = false
-  | eq_fm F (Imp (a, b)) = false
-  | eq_fm F (Or (a, b)) = false
-  | eq_fm F (And (a, b)) = false
-  | eq_fm F (Not a) = false
-  | eq_fm F (NDvd (a, b)) = false
-  | eq_fm F (Dvd (a, b)) = false
-  | eq_fm F (NEq a) = false
-  | eq_fm F (Eq a) = false
-  | eq_fm F (Ge a) = false
-  | eq_fm F (Gt a) = false
-  | eq_fm F (Le a) = false
-  | eq_fm F (Lt a) = false
-  | eq_fm T (NClosed a) = false
-  | eq_fm T (Closed a) = false
-  | eq_fm T (A a) = false
-  | eq_fm T (E a) = false
-  | eq_fm T (Iff (a, b)) = false
-  | eq_fm T (Imp (a, b)) = false
-  | eq_fm T (Or (a, b)) = false
-  | eq_fm T (And (a, b)) = false
-  | eq_fm T (Not a) = false
-  | eq_fm T (NDvd (a, b)) = false
-  | eq_fm T (Dvd (a, b)) = false
-  | eq_fm T (NEq a) = false
-  | eq_fm T (Eq a) = false
-  | eq_fm T (Ge a) = false
-  | eq_fm T (Gt a) = false
-  | eq_fm T (Le a) = false
-  | eq_fm T (Lt a) = false
-  | eq_fm T F = false
-  | eq_fm (NClosed nat) (NClosed nat') = ((nat : IntInf.int) = nat')
-  | eq_fm (Closed nat) (Closed nat') = ((nat : IntInf.int) = nat')
-  | eq_fm (A fm) (A fm') = eq_fm fm fm'
-  | eq_fm (E fm) (E fm') = eq_fm fm fm'
-  | eq_fm (Iff (fm1, fm2)) (Iff (fm1', fm2')) =
-    eq_fm fm1 fm1' andalso eq_fm fm2 fm2'
-  | eq_fm (Imp (fm1, fm2)) (Imp (fm1', fm2')) =
-    eq_fm fm1 fm1' andalso eq_fm fm2 fm2'
-  | eq_fm (Or (fm1, fm2)) (Or (fm1', fm2')) =
-    eq_fm fm1 fm1' andalso eq_fm fm2 fm2'
-  | eq_fm (And (fm1, fm2)) (And (fm1', fm2')) =
-    eq_fm fm1 fm1' andalso eq_fm fm2 fm2'
-  | eq_fm (Not fm) (Not fm') = eq_fm fm fm'
-  | eq_fm (NDvd (inta, num)) (NDvd (int', num')) =
-    ((inta : IntInf.int) = int') andalso eq_num num num'
-  | eq_fm (Dvd (inta, num)) (Dvd (int', num')) =
-    ((inta : IntInf.int) = int') andalso eq_num num num'
-  | eq_fm (NEq num) (NEq num') = eq_num num num'
-  | eq_fm (Eq num) (Eq num') = eq_num num num'
-  | eq_fm (Ge num) (Ge num') = eq_num num num'
-  | eq_fm (Gt num) (Gt num') = eq_num num num'
-  | eq_fm (Le num) (Le num') = eq_num num num'
-  | eq_fm (Lt num) (Lt num') = eq_num num num'
-  | eq_fm F F = true
-  | eq_fm T T = true;
-
-val eq_fma = {eq = eq_fm} : fm eq;
+  | eq_fm T (Lt num) = false
+  | eq_fm (Lt num) T = false
+  | eq_fm T (Le num) = false
+  | eq_fm (Le num) T = false
+  | eq_fm T (Gt num) = false
+  | eq_fm (Gt num) T = false
+  | eq_fm T (Ge num) = false
+  | eq_fm (Ge num) T = false
+  | eq_fm T (Eq num) = false
+  | eq_fm (Eq num) T = false
+  | eq_fm T (NEq num) = false
+  | eq_fm (NEq num) T = false
+  | eq_fm T (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) T = false
+  | eq_fm T (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) T = false
+  | eq_fm T (Not fm) = false
+  | eq_fm (Not fm) T = false
+  | eq_fm T (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) T = false
+  | eq_fm T (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) T = false
+  | eq_fm T (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) T = false
+  | eq_fm T (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) T = false
+  | eq_fm T (E fm) = false
+  | eq_fm (E fm) T = false
+  | eq_fm T (A fm) = false
+  | eq_fm (A fm) T = false
+  | eq_fm T (Closed nat) = false
+  | eq_fm (Closed nat) T = false
+  | eq_fm T (NClosed nat) = false
+  | eq_fm (NClosed nat) T = false
+  | eq_fm F (Lt num) = false
+  | eq_fm (Lt num) F = false
+  | eq_fm F (Le num) = false
+  | eq_fm (Le num) F = false
+  | eq_fm F (Gt num) = false
+  | eq_fm (Gt num) F = false
+  | eq_fm F (Ge num) = false
+  | eq_fm (Ge num) F = false
+  | eq_fm F (Eq num) = false
+  | eq_fm (Eq num) F = false
+  | eq_fm F (NEq num) = false
+  | eq_fm (NEq num) F = false
+  | eq_fm F (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) F = false
+  | eq_fm F (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) F = false
+  | eq_fm F (Not fm) = false
+  | eq_fm (Not fm) F = false
+  | eq_fm F (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) F = false
+  | eq_fm F (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) F = false
+  | eq_fm F (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) F = false
+  | eq_fm F (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) F = false
+  | eq_fm F (E fm) = false
+  | eq_fm (E fm) F = false
+  | eq_fm F (A fm) = false
+  | eq_fm (A fm) F = false
+  | eq_fm F (Closed nat) = false
+  | eq_fm (Closed nat) F = false
+  | eq_fm F (NClosed nat) = false
+  | eq_fm (NClosed nat) F = false
+  | eq_fm (Lt numa) (Le num) = false
+  | eq_fm (Le numa) (Lt num) = false
+  | eq_fm (Lt numa) (Gt num) = false
+  | eq_fm (Gt numa) (Lt num) = false
+  | eq_fm (Lt numa) (Ge num) = false
+  | eq_fm (Ge numa) (Lt num) = false
+  | eq_fm (Lt numa) (Eq num) = false
+  | eq_fm (Eq numa) (Lt num) = false
+  | eq_fm (Lt numa) (NEq num) = false
+  | eq_fm (NEq numa) (Lt num) = false
+  | eq_fm (Lt numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (Lt num) = false
+  | eq_fm (Lt numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (Lt num) = false
+  | eq_fm (Lt num) (Not fm) = false
+  | eq_fm (Not fm) (Lt num) = false
+  | eq_fm (Lt num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Lt num) = false
+  | eq_fm (Lt num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Lt num) = false
+  | eq_fm (Lt num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Lt num) = false
+  | eq_fm (Lt num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Lt num) = false
+  | eq_fm (Lt num) (E fm) = false
+  | eq_fm (E fm) (Lt num) = false
+  | eq_fm (Lt num) (A fm) = false
+  | eq_fm (A fm) (Lt num) = false
+  | eq_fm (Lt num) (Closed nat) = false
+  | eq_fm (Closed nat) (Lt num) = false
+  | eq_fm (Lt num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Lt num) = false
+  | eq_fm (Le numa) (Gt num) = false
+  | eq_fm (Gt numa) (Le num) = false
+  | eq_fm (Le numa) (Ge num) = false
+  | eq_fm (Ge numa) (Le num) = false
+  | eq_fm (Le numa) (Eq num) = false
+  | eq_fm (Eq numa) (Le num) = false
+  | eq_fm (Le numa) (NEq num) = false
+  | eq_fm (NEq numa) (Le num) = false
+  | eq_fm (Le numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (Le num) = false
+  | eq_fm (Le numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (Le num) = false
+  | eq_fm (Le num) (Not fm) = false
+  | eq_fm (Not fm) (Le num) = false
+  | eq_fm (Le num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Le num) = false
+  | eq_fm (Le num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Le num) = false
+  | eq_fm (Le num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Le num) = false
+  | eq_fm (Le num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Le num) = false
+  | eq_fm (Le num) (E fm) = false
+  | eq_fm (E fm) (Le num) = false
+  | eq_fm (Le num) (A fm) = false
+  | eq_fm (A fm) (Le num) = false
+  | eq_fm (Le num) (Closed nat) = false
+  | eq_fm (Closed nat) (Le num) = false
+  | eq_fm (Le num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Le num) = false
+  | eq_fm (Gt numa) (Ge num) = false
+  | eq_fm (Ge numa) (Gt num) = false
+  | eq_fm (Gt numa) (Eq num) = false
+  | eq_fm (Eq numa) (Gt num) = false
+  | eq_fm (Gt numa) (NEq num) = false
+  | eq_fm (NEq numa) (Gt num) = false
+  | eq_fm (Gt numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (Gt num) = false
+  | eq_fm (Gt numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (Gt num) = false
+  | eq_fm (Gt num) (Not fm) = false
+  | eq_fm (Not fm) (Gt num) = false
+  | eq_fm (Gt num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Gt num) = false
+  | eq_fm (Gt num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Gt num) = false
+  | eq_fm (Gt num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Gt num) = false
+  | eq_fm (Gt num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Gt num) = false
+  | eq_fm (Gt num) (E fm) = false
+  | eq_fm (E fm) (Gt num) = false
+  | eq_fm (Gt num) (A fm) = false
+  | eq_fm (A fm) (Gt num) = false
+  | eq_fm (Gt num) (Closed nat) = false
+  | eq_fm (Closed nat) (Gt num) = false
+  | eq_fm (Gt num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Gt num) = false
+  | eq_fm (Ge numa) (Eq num) = false
+  | eq_fm (Eq numa) (Ge num) = false
+  | eq_fm (Ge numa) (NEq num) = false
+  | eq_fm (NEq numa) (Ge num) = false
+  | eq_fm (Ge numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (Ge num) = false
+  | eq_fm (Ge numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (Ge num) = false
+  | eq_fm (Ge num) (Not fm) = false
+  | eq_fm (Not fm) (Ge num) = false
+  | eq_fm (Ge num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Ge num) = false
+  | eq_fm (Ge num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Ge num) = false
+  | eq_fm (Ge num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Ge num) = false
+  | eq_fm (Ge num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Ge num) = false
+  | eq_fm (Ge num) (E fm) = false
+  | eq_fm (E fm) (Ge num) = false
+  | eq_fm (Ge num) (A fm) = false
+  | eq_fm (A fm) (Ge num) = false
+  | eq_fm (Ge num) (Closed nat) = false
+  | eq_fm (Closed nat) (Ge num) = false
+  | eq_fm (Ge num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Ge num) = false
+  | eq_fm (Eq numa) (NEq num) = false
+  | eq_fm (NEq numa) (Eq num) = false
+  | eq_fm (Eq numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (Eq num) = false
+  | eq_fm (Eq numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (Eq num) = false
+  | eq_fm (Eq num) (Not fm) = false
+  | eq_fm (Not fm) (Eq num) = false
+  | eq_fm (Eq num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Eq num) = false
+  | eq_fm (Eq num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Eq num) = false
+  | eq_fm (Eq num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Eq num) = false
+  | eq_fm (Eq num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Eq num) = false
+  | eq_fm (Eq num) (E fm) = false
+  | eq_fm (E fm) (Eq num) = false
+  | eq_fm (Eq num) (A fm) = false
+  | eq_fm (A fm) (Eq num) = false
+  | eq_fm (Eq num) (Closed nat) = false
+  | eq_fm (Closed nat) (Eq num) = false
+  | eq_fm (Eq num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Eq num) = false
+  | eq_fm (NEq numa) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, numa)) (NEq num) = false
+  | eq_fm (NEq numa) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, numa)) (NEq num) = false
+  | eq_fm (NEq num) (Not fm) = false
+  | eq_fm (Not fm) (NEq num) = false
+  | eq_fm (NEq num) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (NEq num) = false
+  | eq_fm (NEq num) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (NEq num) = false
+  | eq_fm (NEq num) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (NEq num) = false
+  | eq_fm (NEq num) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (NEq num) = false
+  | eq_fm (NEq num) (E fm) = false
+  | eq_fm (E fm) (NEq num) = false
+  | eq_fm (NEq num) (A fm) = false
+  | eq_fm (A fm) (NEq num) = false
+  | eq_fm (NEq num) (Closed nat) = false
+  | eq_fm (Closed nat) (NEq num) = false
+  | eq_fm (NEq num) (NClosed nat) = false
+  | eq_fm (NClosed nat) (NEq num) = false
+  | eq_fm (Dvd (intaa, numa)) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (intaa, numa)) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (Not fm) = false
+  | eq_fm (Not fm) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (E fm) = false
+  | eq_fm (E fm) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (A fm) = false
+  | eq_fm (A fm) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (Closed nat) = false
+  | eq_fm (Closed nat) (Dvd (inta, num)) = false
+  | eq_fm (Dvd (inta, num)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Dvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (Not fm) = false
+  | eq_fm (Not fm) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (E fm) = false
+  | eq_fm (E fm) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (A fm) = false
+  | eq_fm (A fm) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (Closed nat) = false
+  | eq_fm (Closed nat) (NDvd (inta, num)) = false
+  | eq_fm (NDvd (inta, num)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (NDvd (inta, num)) = false
+  | eq_fm (Not fm) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Not fm) = false
+  | eq_fm (Not fm) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Not fm) = false
+  | eq_fm (Not fm) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Not fm) = false
+  | eq_fm (Not fm) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Not fm) = false
+  | eq_fm (Not fma) (E fm) = false
+  | eq_fm (E fma) (Not fm) = false
+  | eq_fm (Not fma) (A fm) = false
+  | eq_fm (A fma) (Not fm) = false
+  | eq_fm (Not fm) (Closed nat) = false
+  | eq_fm (Closed nat) (Not fm) = false
+  | eq_fm (Not fm) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Not fm) = false
+  | eq_fm (And (fm1a, fm2a)) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1a, fm2a)) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1a, fm2a)) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1a, fm2a)) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1a, fm2a)) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1a, fm2a)) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (E fm) = false
+  | eq_fm (E fm) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (A fm) = false
+  | eq_fm (A fm) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (Closed nat) = false
+  | eq_fm (Closed nat) (And (fm1, fm2)) = false
+  | eq_fm (And (fm1, fm2)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (And (fm1, fm2)) = false
+  | eq_fm (Or (fm1a, fm2a)) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1a, fm2a)) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1a, fm2a)) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1a, fm2a)) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (E fm) = false
+  | eq_fm (E fm) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (A fm) = false
+  | eq_fm (A fm) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (Closed nat) = false
+  | eq_fm (Closed nat) (Or (fm1, fm2)) = false
+  | eq_fm (Or (fm1, fm2)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Or (fm1, fm2)) = false
+  | eq_fm (Imp (fm1a, fm2a)) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1a, fm2a)) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (E fm) = false
+  | eq_fm (E fm) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (A fm) = false
+  | eq_fm (A fm) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (Closed nat) = false
+  | eq_fm (Closed nat) (Imp (fm1, fm2)) = false
+  | eq_fm (Imp (fm1, fm2)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Imp (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (E fm) = false
+  | eq_fm (E fm) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (A fm) = false
+  | eq_fm (A fm) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (Closed nat) = false
+  | eq_fm (Closed nat) (Iff (fm1, fm2)) = false
+  | eq_fm (Iff (fm1, fm2)) (NClosed nat) = false
+  | eq_fm (NClosed nat) (Iff (fm1, fm2)) = false
+  | eq_fm (E fma) (A fm) = false
+  | eq_fm (A fma) (E fm) = false
+  | eq_fm (E fm) (Closed nat) = false
+  | eq_fm (Closed nat) (E fm) = false
+  | eq_fm (E fm) (NClosed nat) = false
+  | eq_fm (NClosed nat) (E fm) = false
+  | eq_fm (A fm) (Closed nat) = false
+  | eq_fm (Closed nat) (A fm) = false
+  | eq_fm (A fm) (NClosed nat) = false
+  | eq_fm (NClosed nat) (A fm) = false
+  | eq_fm (Closed nata) (NClosed nat) = false
+  | eq_fm (NClosed nata) (Closed nat) = false;
 
 fun djf f p q =
-  (if eqop eq_fma q T then T
-    else (if eqop eq_fma q F then f p
-           else let
-                  val a = f p;
-                in
-                  (case a of T => T | F => q | Lt num => Or (f p, q)
-                     | Le num => Or (f p, q) | Gt num => Or (f p, q)
-                     | Ge num => Or (f p, q) | Eq num => Or (f p, q)
-                     | NEq num => Or (f p, q) | Dvd (inta, num) => Or (f p, q)
-                     | NDvd (inta, num) => Or (f p, q) | Not fm => Or (f p, q)
-                     | And (fm1, fm2) => Or (f p, q)
-                     | Or (fm1, fm2) => Or (f p, q)
-                     | Imp (fm1, fm2) => Or (f p, q)
-                     | Iff (fm1, fm2) => Or (f p, q) | E fm => Or (f p, q)
-                     | A fm => Or (f p, q) | Closed nat => Or (f p, q)
-                     | NClosed nat => Or (f p, q))
-                end));
+  (if eq_fm q T then T
+    else (if eq_fm q F then f p
+           else (case f p of T => T | F => q | Lt _ => Or (f p, q)
+                  | Le _ => Or (f p, q) | Gt _ => Or (f p, q)
+                  | Ge _ => Or (f p, q) | Eq _ => Or (f p, q)
+                  | NEq _ => Or (f p, q) | Dvd (_, _) => Or (f p, q)
+                  | NDvd (_, _) => Or (f p, q) | Not _ => Or (f p, q)
+                  | And (_, _) => Or (f p, q) | Or (_, _) => Or (f p, q)
+                  | Imp (_, _) => Or (f p, q) | Iff (_, _) => Or (f p, q)
+                  | 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);
@@ -562,18 +768,17 @@
 fun dj f p = evaldjf f (disjuncts p);
 
 fun disj p q =
-  (if eqop eq_fma p T orelse eqop eq_fma q T then T
-    else (if eqop eq_fma p F then q
-           else (if eqop eq_fma q F then p else Or (p, q))));
+  (if eq_fm p T orelse eq_fm q T then T
+    else (if eq_fm p F then q else (if eq_fm q F then p else Or (p, q))));
 
 fun minus_nat n m = IntInf.max (0, (IntInf.- (n, m)));
 
-fun decrnum (Bound n) = Bound (minus_nat n 1)
+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, i, decrnum a)
+  | decrnum (Cn (n, i, a)) = Cn (minus_nat n (1 : IntInf.int), i, decrnum a)
   | decrnum (C u) = C u;
 
 fun decr (Lt a) = Lt (decrnum a)
@@ -596,20 +801,20 @@
   | decr (Closed aq) = Closed aq
   | decr (NClosed ar) = NClosed ar;
 
-fun concat [] = []
-  | concat (x :: xs) = append x (concat xs);
-
-fun split f (a, b) = f a b;
+fun concat_map f [] = []
+  | concat_map f (x :: xs) = append (f x) (concat_map f xs);
 
 fun numsubst0 t (C c) = C c
-  | numsubst0 t (Bound n) = (if eqop eq_nat n 0 then t else Bound n)
+  | numsubst0 t (Bound n) =
+    (if ((n : IntInf.int) = (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 eqop eq_nat v 0 then Add (Mul (i, t), numsubst0 t a)
-      else Cn (suc (minus_nat v 1), i, numsubst0 t a));
+    (if ((v : IntInf.int) = (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
@@ -679,49 +884,417 @@
   | minusinf (Closed ap) = Closed ap
   | minusinf (NClosed aq) = NClosed aq
   | minusinf (Lt (Cn (cm, c, e))) =
-    (if eqop eq_nat cm 0 then T else Lt (Cn (suc (minus_nat cm 1), c, e)))
+    (if ((cm : IntInf.int) = (0 : IntInf.int)) then T
+      else Lt (Cn (suc (minus_nat cm (1 : IntInf.int)), c, e)))
   | minusinf (Le (Cn (dm, c, e))) =
-    (if eqop eq_nat dm 0 then T else Le (Cn (suc (minus_nat dm 1), c, e)))
+    (if ((dm : IntInf.int) = (0 : IntInf.int)) then T
+      else Le (Cn (suc (minus_nat dm (1 : IntInf.int)), c, e)))
   | minusinf (Gt (Cn (em, c, e))) =
-    (if eqop eq_nat em 0 then F else Gt (Cn (suc (minus_nat em 1), c, e)))
+    (if ((em : IntInf.int) = (0 : IntInf.int)) then F
+      else Gt (Cn (suc (minus_nat em (1 : IntInf.int)), c, e)))
   | minusinf (Ge (Cn (fm, c, e))) =
-    (if eqop eq_nat fm 0 then F else Ge (Cn (suc (minus_nat fm 1), c, e)))
+    (if ((fm : IntInf.int) = (0 : IntInf.int)) then F
+      else Ge (Cn (suc (minus_nat fm (1 : IntInf.int)), c, e)))
   | minusinf (Eq (Cn (gm, c, e))) =
-    (if eqop eq_nat gm 0 then F else Eq (Cn (suc (minus_nat gm 1), c, e)))
+    (if ((gm : IntInf.int) = (0 : IntInf.int)) then F
+      else Eq (Cn (suc (minus_nat gm (1 : IntInf.int)), c, e)))
   | minusinf (NEq (Cn (hm, c, e))) =
-    (if eqop eq_nat hm 0 then T else NEq (Cn (suc (minus_nat hm 1), c, e)));
+    (if ((hm : IntInf.int) = (0 : IntInf.int)) then T
+      else NEq (Cn (suc (minus_nat hm (1 : IntInf.int)), c, e)));
 
 val eq_int = {eq = (fn a => fn b => ((a : IntInf.int) = b))} : IntInf.int eq;
 
+val zero_int : IntInf.int = (0 : IntInf.int);
+
+type 'a zero = {zero : 'a};
+val zero = #zero : 'a zero -> 'a;
+
+val zero_inta = {zero = zero_int} : IntInf.int zero;
+
+type 'a times = {times : 'a -> 'a -> 'a};
+val times = #times : 'a times -> 'a -> 'a -> '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;
+
+val times_int = {times = (fn a => fn b => IntInf.* (a, b))} : IntInf.int times;
+
+val no_zero_divisors_int =
+  {times_no_zero_divisors = times_int, zero_no_zero_divisors = zero_inta} :
+  IntInf.int no_zero_divisors;
+
+type 'a one = {one : 'a};
+val one = #one : 'a one -> 'a;
+
+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 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 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 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 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 dvd = {times_dvd : 'a times};
+val times_dvd = #times_dvd : 'a dvd -> 'a times;
+
+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 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 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;
+
+val one_int : IntInf.int = (1 : IntInf.int);
+
+val one_inta = {one = one_int} : IntInf.int one;
+
+val zero_neq_one_int =
+  {one_zero_neq_one = one_inta, zero_zero_neq_one = zero_inta} :
+  IntInf.int zero_neq_one;
+
+val semigroup_mult_int = {times_semigroup_mult = times_int} :
+  IntInf.int semigroup_mult;
+
+val plus_int = {plus = (fn a => fn b => IntInf.+ (a, b))} : IntInf.int plus;
+
+val semigroup_add_int = {plus_semigroup_add = plus_int} :
+  IntInf.int semigroup_add;
+
+val ab_semigroup_add_int = {semigroup_add_ab_semigroup_add = semigroup_add_int}
+  : IntInf.int ab_semigroup_add;
+
+val semiring_int =
+  {ab_semigroup_add_semiring = ab_semigroup_add_int,
+    semigroup_mult_semiring = semigroup_mult_int}
+  : IntInf.int semiring;
+
+val mult_zero_int = {times_mult_zero = times_int, zero_mult_zero = zero_inta} :
+  IntInf.int mult_zero;
+
+val monoid_add_int =
+  {semigroup_add_monoid_add = semigroup_add_int, zero_monoid_add = zero_inta} :
+  IntInf.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}
+  : IntInf.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}
+  : IntInf.int semiring_0;
+
+val power_int = {one_power = one_inta, times_power = times_int} :
+  IntInf.int power;
+
+val monoid_mult_int =
+  {semigroup_mult_monoid_mult = semigroup_mult_int,
+    power_monoid_mult = power_int}
+  : IntInf.int monoid_mult;
+
+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}
+  : IntInf.int semiring_1;
+
+val cancel_semigroup_add_int =
+  {semigroup_add_cancel_semigroup_add = semigroup_add_int} :
+  IntInf.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}
+  : IntInf.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}
+  : IntInf.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}
+  : IntInf.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}
+  : IntInf.int semiring_1_cancel;
+
+val dvd_int = {times_dvd = times_int} : IntInf.int dvd;
+
+val ab_semigroup_mult_int =
+  {semigroup_mult_ab_semigroup_mult = semigroup_mult_int} :
+  IntInf.int ab_semigroup_mult;
+
+val comm_semiring_int =
+  {ab_semigroup_mult_comm_semiring = ab_semigroup_mult_int,
+    semiring_comm_semiring = semiring_int}
+  : IntInf.int comm_semiring;
+
+val comm_semiring_0_int =
+  {comm_semiring_comm_semiring_0 = comm_semiring_int,
+    semiring_0_comm_semiring_0 = semiring_0_int}
+  : IntInf.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}
+  : IntInf.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}
+  : IntInf.int comm_semiring_1;
+
+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}
+  : IntInf.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}
+  : IntInf.int comm_semiring_1_cancel;
+
+fun abs_int i = (if IntInf.< (i, (0 : IntInf.int)) then IntInf.~ i else i);
+
+fun split f (a, b) = f a b;
+
 fun sgn_int i =
-  (if eqop eq_int i (0 : IntInf.int) then (0 : IntInf.int)
+  (if ((i : IntInf.int) = (0 : IntInf.int)) then (0 : IntInf.int)
     else (if IntInf.< ((0 : IntInf.int), i) then (1 : IntInf.int)
            else IntInf.~ (1 : IntInf.int)));
 
 fun apsnd f (x, y) = (x, f y);
 
-fun divmoda k l =
-  (if eqop eq_int k (0 : IntInf.int) then ((0 : IntInf.int), (0 : IntInf.int))
-    else (if eqop eq_int l (0 : IntInf.int) then ((0 : IntInf.int), k)
+fun divmod_int k l =
+  (if ((k : IntInf.int) = (0 : IntInf.int))
+    then ((0 : IntInf.int), (0 : IntInf.int))
+    else (if ((l : IntInf.int) = (0 : IntInf.int)) then ((0 : IntInf.int), k)
            else apsnd (fn a => IntInf.* (sgn_int l, a))
-                  (if eqop eq_int (sgn_int k) (sgn_int l)
-                    then (fn k => fn l => IntInf.divMod (IntInf.abs k,
-                           IntInf.abs l))
-                           k l
+                  (if (((sgn_int k) : IntInf.int) = (sgn_int l))
+                    then IntInf.divMod (IntInf.abs k, IntInf.abs l)
                     else let
-                           val a =
-                             (fn k => fn l => IntInf.divMod (IntInf.abs k,
-                               IntInf.abs l))
-                               k l;
-                           val (r, s) = a;
+                           val (r, s) =
+                             IntInf.divMod (IntInf.abs k, IntInf.abs l);
                          in
-                           (if eqop eq_int s (0 : IntInf.int)
+                           (if ((s : IntInf.int) = (0 : IntInf.int))
                              then (IntInf.~ r, (0 : IntInf.int))
                              else (IntInf.- (IntInf.~ r, (1 : IntInf.int)),
                                     IntInf.- (abs_int l, s)))
                          end)));
 
-fun mod_int a b = snd (divmoda a b);
+fun snd (a, b) = b;
+
+fun mod_int a b = snd (divmod_int a b);
+
+fun fst (a, b) = a;
+
+fun div_int a b = fst (divmod_int a b);
+
+val div_inta = {dvd_div = dvd_int, diva = div_int, moda = mod_int} :
+  IntInf.int diva;
+
+val semiring_div_int =
+  {div_semiring_div = div_inta,
+    comm_semiring_1_cancel_semiring_div = comm_semiring_1_cancel_int,
+    no_zero_divisors_semiring_div = no_zero_divisors_int}
+  : IntInf.int semiring_div;
+
+fun dvd (A1_, A2_) a b =
+  eqa A2_ (moda (div_semiring_div A1_) b a)
+    (zero ((zero_no_zero_divisors o no_zero_divisors_semiring_div) A1_));
 
 fun num_case f1 f2 f3 f4 f5 f6 f7 (Mul (inta, num)) = f7 inta num
   | num_case f1 f2 f3 f4 f5 f6 f7 (Sub (num1, num2)) = f6 num1 num2
@@ -742,11 +1315,11 @@
 fun numneg t = nummul (IntInf.~ (1 : IntInf.int)) t;
 
 fun numadd (Cn (n1, c1, r1), Cn (n2, c2, r2)) =
-  (if eqop eq_nat n1 n2
+  (if ((n1 : IntInf.int) = n2)
     then let
            val c = IntInf.+ (c1, c2);
          in
-           (if eqop eq_int c (0 : IntInf.int) then numadd (r1, r2)
+           (if ((c : IntInf.int) = (0 : IntInf.int)) then numadd (r1, r2)
              else Cn (n1, c, numadd (r1, r2)))
          end
     else (if IntInf.<= (n1, n2)
@@ -807,10 +1380,8 @@
   | numadd (Mul (at, au), Sub (hp, hq)) = Add (Mul (at, au), Sub (hp, hq))
   | numadd (Mul (at, au), Mul (hr, hs)) = Add (Mul (at, au), Mul (hr, hs));
 
-val eq_numa = {eq = eq_num} : num eq;
-
 fun numsub s t =
-  (if eqop eq_numa s t then C (0 : IntInf.int) else numadd (s, numneg t));
+  (if eq_num 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))
@@ -818,7 +1389,7 @@
   | simpnum (Add (t, s)) = numadd (simpnum t, simpnum s)
   | simpnum (Sub (t, s)) = numsub (simpnum t) (simpnum s)
   | simpnum (Mul (i, t)) =
-    (if eqop eq_int i (0 : IntInf.int) then C (0 : IntInf.int)
+    (if ((i : IntInf.int) = (0 : IntInf.int)) then C (0 : IntInf.int)
       else nummul i (simpnum t))
   | simpnum (Cn (v, va, vb)) = Cn (v, va, vb);
 
@@ -843,23 +1414,20 @@
   | nota (NClosed v) = Not (NClosed v);
 
 fun iffa p q =
-  (if eqop eq_fma p q then T
-    else (if eqop eq_fma p (nota q) orelse eqop eq_fma (nota p) q then F
-           else (if eqop eq_fma p F then nota q
-                  else (if eqop eq_fma q F then nota p
-                         else (if eqop eq_fma p T then q
-                                else (if eqop eq_fma q T then p
-                                       else Iff (p, q)))))));
+  (if eq_fm p q then T
+    else (if eq_fm p (nota q) orelse eq_fm (nota p) q then F
+           else (if eq_fm p F then nota q
+                  else (if eq_fm q F then nota p
+                         else (if eq_fm p T then q
+                                else (if eq_fm q T then p else Iff (p, q)))))));
 
 fun impa p q =
-  (if eqop eq_fma p F orelse eqop eq_fma q T then T
-    else (if eqop eq_fma p T then q
-           else (if eqop eq_fma q F then nota p else Imp (p, q))));
+  (if eq_fm p F orelse eq_fm q T then T
+    else (if eq_fm p T then q else (if eq_fm q F then nota p else Imp (p, q))));
 
 fun conj p q =
-  (if eqop eq_fma p F orelse eqop eq_fma q F then F
-    else (if eqop eq_fma p T then q
-           else (if eqop eq_fma q T then p else And (p, q))));
+  (if eq_fm p F orelse eq_fm q F then F
+    else (if eq_fm p T then q else (if eq_fm q T then p else And (p, q))));
 
 fun simpfm (And (p, q)) = conj (simpfm p) (simpfm q)
   | simpfm (Or (p, q)) = disj (simpfm p) (simpfm q)
@@ -868,91 +1436,80 @@
   | simpfm (Not p) = nota (simpfm p)
   | simpfm (Lt a) =
     let
-      val a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if IntInf.< (v, (0 : IntInf.int)) then T else F)
-         | Bound nat => Lt a' | Cn (nat, inta, num) => Lt a' | Neg num => Lt a'
-         | Add (num1, num2) => Lt a' | Sub (num1, num2) => Lt a'
-         | Mul (inta, num) => Lt a')
+      (case aa of C v => (if IntInf.< (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 a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if IntInf.<= (v, (0 : IntInf.int)) then T else F)
-         | Bound nat => Le a' | Cn (nat, inta, num) => Le a' | Neg num => Le a'
-         | Add (num1, num2) => Le a' | Sub (num1, num2) => Le a'
-         | Mul (inta, num) => Le a')
+      (case aa of C v => (if IntInf.<= (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 a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if IntInf.< ((0 : IntInf.int), v) then T else F)
-         | Bound nat => Gt a' | Cn (nat, inta, num) => Gt a' | Neg num => Gt a'
-         | Add (num1, num2) => Gt a' | Sub (num1, num2) => Gt a'
-         | Mul (inta, num) => Gt a')
+      (case aa of C v => (if IntInf.< ((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 a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if IntInf.<= ((0 : IntInf.int), v) then T else F)
-         | Bound nat => Ge a' | Cn (nat, inta, num) => Ge a' | Neg num => Ge a'
-         | Add (num1, num2) => Ge a' | Sub (num1, num2) => Ge a'
-         | Mul (inta, num) => Ge a')
+      (case aa of C v => (if IntInf.<= ((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 a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if eqop eq_int v (0 : IntInf.int) then T else F)
-         | Bound nat => Eq a' | Cn (nat, inta, num) => Eq a' | Neg num => Eq a'
-         | Add (num1, num2) => Eq a' | Sub (num1, num2) => Eq a'
-         | Mul (inta, num) => Eq a')
+      (case aa
+        of C v => (if ((v : IntInf.int) = (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 a' = simpnum a;
+      val aa = simpnum a;
     in
-      (case a' of C v => (if not (eqop eq_int v (0 : IntInf.int)) then T else F)
-         | Bound nat => NEq a' | Cn (nat, inta, num) => NEq a'
-         | Neg num => NEq a' | Add (num1, num2) => NEq a'
-         | Sub (num1, num2) => NEq a' | Mul (inta, num) => NEq a')
+      (case aa
+        of C v => (if not ((v : IntInf.int) = (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 eqop eq_int i (0 : IntInf.int) then simpfm (Eq a)
-      else (if eqop eq_int (abs_int i) (1 : IntInf.int) then T
+    (if ((i : IntInf.int) = (0 : IntInf.int)) then simpfm (Eq a)
+      else (if (((abs_int i) : IntInf.int) = (1 : IntInf.int)) then T
              else let
-                    val a' = simpnum a;
+                    val aa = simpnum a;
                   in
-                    (case a'
-                       of C v =>
-                         (if eqop eq_int (mod_int v i) (0 : IntInf.int) then T
-                           else F)
-                       | Bound nat => Dvd (i, a')
-                       | Cn (nat, inta, num) => Dvd (i, a')
-                       | Neg num => Dvd (i, a')
-                       | Add (num1, num2) => Dvd (i, a')
-                       | Sub (num1, num2) => Dvd (i, a')
-                       | Mul (inta, num) => Dvd (i, a'))
+                    (case aa
+                      of C v =>
+                        (if dvd (semiring_div_int, eq_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 eqop eq_int i (0 : IntInf.int) then simpfm (NEq a)
-      else (if eqop eq_int (abs_int i) (1 : IntInf.int) then F
+    (if ((i : IntInf.int) = (0 : IntInf.int)) then simpfm (NEq a)
+      else (if (((abs_int i) : IntInf.int) = (1 : IntInf.int)) then F
              else let
-                    val a' = simpnum a;
+                    val aa = simpnum a;
                   in
-                    (case a'
-                       of C v =>
-                         (if not (eqop eq_int (mod_int v i) (0 : IntInf.int))
-                           then T else F)
-                       | Bound nat => NDvd (i, a')
-                       | Cn (nat, inta, num) => NDvd (i, a')
-                       | Neg num => NDvd (i, a')
-                       | Add (num1, num2) => NDvd (i, a')
-                       | Sub (num1, num2) => NDvd (i, a')
-                       | Mul (inta, num) => NDvd (i, a'))
+                    (case aa
+                      of C v =>
+                        (if not (dvd (semiring_div_int, eq_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
@@ -1025,32 +1582,40 @@
   | mirror (Closed ap) = Closed ap
   | mirror (NClosed aq) = NClosed aq
   | mirror (Lt (Cn (cm, c, e))) =
-    (if eqop eq_nat cm 0 then Gt (Cn (0, c, Neg e))
-      else Lt (Cn (suc (minus_nat cm 1), c, e)))
+    (if ((cm : IntInf.int) = (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 eqop eq_nat dm 0 then Ge (Cn (0, c, Neg e))
-      else Le (Cn (suc (minus_nat dm 1), c, e)))
+    (if ((dm : IntInf.int) = (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 eqop eq_nat em 0 then Lt (Cn (0, c, Neg e))
-      else Gt (Cn (suc (minus_nat em 1), c, e)))
+    (if ((em : IntInf.int) = (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 eqop eq_nat fm 0 then Le (Cn (0, c, Neg e))
-      else Ge (Cn (suc (minus_nat fm 1), c, e)))
+    (if ((fm : IntInf.int) = (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 eqop eq_nat gm 0 then Eq (Cn (0, c, Neg e))
-      else Eq (Cn (suc (minus_nat gm 1), c, e)))
+    (if ((gm : IntInf.int) = (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 eqop eq_nat hm 0 then NEq (Cn (0, c, Neg e))
-      else NEq (Cn (suc (minus_nat hm 1), c, e)))
+    (if ((hm : IntInf.int) = (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 eqop eq_nat im 0 then Dvd (i, Cn (0, c, Neg e))
-      else Dvd (i, Cn (suc (minus_nat im 1), c, e)))
+    (if ((im : IntInf.int) = (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 eqop eq_nat jm 0 then NDvd (i, Cn (0, c, Neg e))
-      else NDvd (i, Cn (suc (minus_nat jm 1), c, e)));
+    (if ((jm : IntInf.int) = (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
-  | size_list (a :: lista) = IntInf.+ (size_list lista, suc 0);
+fun size_list [] = (0 : IntInf.int)
+  | size_list (a :: lista) = IntInf.+ (size_list lista, suc (0 : IntInf.int));
 
 fun alpha (And (p, q)) = append (alpha p) (alpha q)
   | alpha (Or (p, q)) = append (alpha p) (alpha q)
@@ -1101,14 +1666,20 @@
   | alpha (A ao) = []
   | alpha (Closed ap) = []
   | alpha (NClosed aq) = []
-  | alpha (Lt (Cn (cm, c, e))) = (if eqop eq_nat cm 0 then [e] else [])
+  | alpha (Lt (Cn (cm, c, e))) =
+    (if ((cm : IntInf.int) = (0 : IntInf.int)) then [e] else [])
   | alpha (Le (Cn (dm, c, e))) =
-    (if eqop eq_nat dm 0 then [Add (C (~1 : IntInf.int), e)] else [])
-  | alpha (Gt (Cn (em, c, e))) = (if eqop eq_nat em 0 then [] else [])
-  | alpha (Ge (Cn (fm, c, e))) = (if eqop eq_nat fm 0 then [] else [])
+    (if ((dm : IntInf.int) = (0 : IntInf.int))
+      then [Add (C (~1 : IntInf.int), e)] else [])
+  | alpha (Gt (Cn (em, c, e))) =
+    (if ((em : IntInf.int) = (0 : IntInf.int)) then [] else [])
+  | alpha (Ge (Cn (fm, c, e))) =
+    (if ((fm : IntInf.int) = (0 : IntInf.int)) then [] else [])
   | alpha (Eq (Cn (gm, c, e))) =
-    (if eqop eq_nat gm 0 then [Add (C (~1 : IntInf.int), e)] else [])
-  | alpha (NEq (Cn (hm, c, e))) = (if eqop eq_nat hm 0 then [e] else []);
+    (if ((gm : IntInf.int) = (0 : IntInf.int))
+      then [Add (C (~1 : IntInf.int), e)] else [])
+  | alpha (NEq (Cn (hm, c, e))) =
+    (if ((hm : IntInf.int) = (0 : IntInf.int)) then [e] else []);
 
 fun beta (And (p, q)) = append (beta p) (beta q)
   | beta (Or (p, q)) = append (beta p) (beta q)
@@ -1159,24 +1730,39 @@
   | beta (A ao) = []
   | beta (Closed ap) = []
   | beta (NClosed aq) = []
-  | beta (Lt (Cn (cm, c, e))) = (if eqop eq_nat cm 0 then [] else [])
-  | beta (Le (Cn (dm, c, e))) = (if eqop eq_nat dm 0 then [] else [])
-  | beta (Gt (Cn (em, c, e))) = (if eqop eq_nat em 0 then [Neg e] else [])
+  | beta (Lt (Cn (cm, c, e))) =
+    (if ((cm : IntInf.int) = (0 : IntInf.int)) then [] else [])
+  | beta (Le (Cn (dm, c, e))) =
+    (if ((dm : IntInf.int) = (0 : IntInf.int)) then [] else [])
+  | beta (Gt (Cn (em, c, e))) =
+    (if ((em : IntInf.int) = (0 : IntInf.int)) then [Neg e] else [])
   | beta (Ge (Cn (fm, c, e))) =
-    (if eqop eq_nat fm 0 then [Sub (C (~1 : IntInf.int), e)] else [])
+    (if ((fm : IntInf.int) = (0 : IntInf.int))
+      then [Sub (C (~1 : IntInf.int), e)] else [])
   | beta (Eq (Cn (gm, c, e))) =
-    (if eqop eq_nat gm 0 then [Sub (C (~1 : IntInf.int), e)] else [])
-  | beta (NEq (Cn (hm, c, e))) = (if eqop eq_nat hm 0 then [Neg e] else []);
+    (if ((gm : IntInf.int) = (0 : IntInf.int))
+      then [Sub (C (~1 : IntInf.int), e)] else [])
+  | beta (NEq (Cn (hm, c, e))) =
+    (if ((hm : IntInf.int) = (0 : IntInf.int)) then [Neg e] else []);
+
+val eq_numa = {eq = eq_num} : num eq;
 
 fun member A_ x [] = false
-  | member A_ x (y :: ys) = eqop A_ x y orelse member A_ x ys;
+  | member A_ x (y :: ys) = eqa A_ x y orelse member A_ x ys;
 
 fun remdups A_ [] = []
   | remdups A_ (x :: xs) =
     (if member A_ x xs then remdups A_ xs else x :: remdups A_ xs);
 
-fun delta (And (p, q)) = zlcm (delta p) (delta q)
-  | delta (Or (p, q)) = zlcm (delta p) (delta q)
+fun gcd_int k l =
+  abs_int
+    (if ((l : IntInf.int) = (0 : IntInf.int)) then k
+      else gcd_int l (mod_int (abs_int k) (abs_int l)));
+
+fun lcm_int a b = div_int (IntInf.* (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 = (1 : IntInf.int)
   | delta F = (1 : IntInf.int)
   | delta (Lt u) = (1 : IntInf.int)
@@ -1205,110 +1791,117 @@
   | delta (Closed ap) = (1 : IntInf.int)
   | delta (NClosed aq) = (1 : IntInf.int)
   | delta (Dvd (i, Cn (cm, c, e))) =
-    (if eqop eq_nat cm 0 then i else (1 : IntInf.int))
+    (if ((cm : IntInf.int) = (0 : IntInf.int)) then i else (1 : IntInf.int))
   | delta (NDvd (i, Cn (dm, c, e))) =
-    (if eqop eq_nat dm 0 then i else (1 : IntInf.int));
-
-fun div_int a b = fst (divmoda a b);
+    (if ((dm : IntInf.int) = (0 : IntInf.int)) then i else (1 : IntInf.int));
 
 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 k => T)
-  | a_beta F = (fn k => F)
-  | a_beta (Lt (C bo)) = (fn k => Lt (C bo))
-  | a_beta (Lt (Bound bp)) = (fn k => Lt (Bound bp))
-  | a_beta (Lt (Neg bt)) = (fn k => Lt (Neg bt))
-  | a_beta (Lt (Add (bu, bv))) = (fn k => Lt (Add (bu, bv)))
-  | a_beta (Lt (Sub (bw, bx))) = (fn k => Lt (Sub (bw, bx)))
-  | a_beta (Lt (Mul (by, bz))) = (fn k => Lt (Mul (by, bz)))
-  | a_beta (Le (C co)) = (fn k => Le (C co))
-  | a_beta (Le (Bound cp)) = (fn k => Le (Bound cp))
-  | a_beta (Le (Neg ct)) = (fn k => Le (Neg ct))
-  | a_beta (Le (Add (cu, cv))) = (fn k => Le (Add (cu, cv)))
-  | a_beta (Le (Sub (cw, cx))) = (fn k => Le (Sub (cw, cx)))
-  | a_beta (Le (Mul (cy, cz))) = (fn k => Le (Mul (cy, cz)))
-  | a_beta (Gt (C doa)) = (fn k => Gt (C doa))
-  | a_beta (Gt (Bound dp)) = (fn k => Gt (Bound dp))
-  | a_beta (Gt (Neg dt)) = (fn k => Gt (Neg dt))
-  | a_beta (Gt (Add (du, dv))) = (fn k => Gt (Add (du, dv)))
-  | a_beta (Gt (Sub (dw, dx))) = (fn k => Gt (Sub (dw, dx)))
-  | a_beta (Gt (Mul (dy, dz))) = (fn k => Gt (Mul (dy, dz)))
-  | a_beta (Ge (C eo)) = (fn k => Ge (C eo))
-  | a_beta (Ge (Bound ep)) = (fn k => Ge (Bound ep))
-  | a_beta (Ge (Neg et)) = (fn k => Ge (Neg et))
-  | a_beta (Ge (Add (eu, ev))) = (fn k => Ge (Add (eu, ev)))
-  | a_beta (Ge (Sub (ew, ex))) = (fn k => Ge (Sub (ew, ex)))
-  | a_beta (Ge (Mul (ey, ez))) = (fn k => Ge (Mul (ey, ez)))
-  | a_beta (Eq (C fo)) = (fn k => Eq (C fo))
-  | a_beta (Eq (Bound fp)) = (fn k => Eq (Bound fp))
-  | a_beta (Eq (Neg ft)) = (fn k => Eq (Neg ft))
-  | a_beta (Eq (Add (fu, fv))) = (fn k => Eq (Add (fu, fv)))
-  | a_beta (Eq (Sub (fw, fx))) = (fn k => Eq (Sub (fw, fx)))
-  | a_beta (Eq (Mul (fy, fz))) = (fn k => Eq (Mul (fy, fz)))
-  | a_beta (NEq (C go)) = (fn k => NEq (C go))
-  | a_beta (NEq (Bound gp)) = (fn k => NEq (Bound gp))
-  | a_beta (NEq (Neg gt)) = (fn k => NEq (Neg gt))
-  | a_beta (NEq (Add (gu, gv))) = (fn k => NEq (Add (gu, gv)))
-  | a_beta (NEq (Sub (gw, gx))) = (fn k => NEq (Sub (gw, gx)))
-  | a_beta (NEq (Mul (gy, gz))) = (fn k => NEq (Mul (gy, gz)))
-  | a_beta (Dvd (aa, C ho)) = (fn k => Dvd (aa, C ho))
-  | a_beta (Dvd (aa, Bound hp)) = (fn k => Dvd (aa, Bound hp))
-  | a_beta (Dvd (aa, Neg ht)) = (fn k => Dvd (aa, Neg ht))
-  | a_beta (Dvd (aa, Add (hu, hv))) = (fn k => Dvd (aa, Add (hu, hv)))
-  | a_beta (Dvd (aa, Sub (hw, hx))) = (fn k => Dvd (aa, Sub (hw, hx)))
-  | a_beta (Dvd (aa, Mul (hy, hz))) = (fn k => Dvd (aa, Mul (hy, hz)))
-  | a_beta (NDvd (ac, C io)) = (fn k => NDvd (ac, C io))
-  | a_beta (NDvd (ac, Bound ip)) = (fn k => NDvd (ac, Bound ip))
-  | a_beta (NDvd (ac, Neg it)) = (fn k => NDvd (ac, Neg it))
-  | a_beta (NDvd (ac, Add (iu, iv))) = (fn k => NDvd (ac, Add (iu, iv)))
-  | a_beta (NDvd (ac, Sub (iw, ix))) = (fn k => NDvd (ac, Sub (iw, ix)))
-  | a_beta (NDvd (ac, Mul (iy, iz))) = (fn k => NDvd (ac, Mul (iy, iz)))
-  | a_beta (Not ae) = (fn k => Not ae)
-  | a_beta (Imp (aj, ak)) = (fn k => Imp (aj, ak))
-  | a_beta (Iff (al, am)) = (fn k => Iff (al, am))
-  | a_beta (E an) = (fn k => E an)
-  | a_beta (A ao) = (fn k => A ao)
-  | a_beta (Closed ap) = (fn k => Closed ap)
-  | a_beta (NClosed aq) = (fn k => NClosed aq)
+  | 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 eqop eq_nat cm 0
-      then (fn k => Lt (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Lt (Cn (suc (minus_nat cm 1), c, e))))
+    (if ((cm : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             Lt (Cn ((0 : IntInf.int), (1 : IntInf.int), Mul (div_int k c, e))))
+      else (fn _ => Lt (Cn (suc (minus_nat cm (1 : IntInf.int)), c, e))))
   | a_beta (Le (Cn (dm, c, e))) =
-    (if eqop eq_nat dm 0
-      then (fn k => Le (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Le (Cn (suc (minus_nat dm 1), c, e))))
+    (if ((dm : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             Le (Cn ((0 : IntInf.int), (1 : IntInf.int), Mul (div_int k c, e))))
+      else (fn _ => Le (Cn (suc (minus_nat dm (1 : IntInf.int)), c, e))))
   | a_beta (Gt (Cn (em, c, e))) =
-    (if eqop eq_nat em 0
-      then (fn k => Gt (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Gt (Cn (suc (minus_nat em 1), c, e))))
+    (if ((em : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             Gt (Cn ((0 : IntInf.int), (1 : IntInf.int), Mul (div_int k c, e))))
+      else (fn _ => Gt (Cn (suc (minus_nat em (1 : IntInf.int)), c, e))))
   | a_beta (Ge (Cn (fm, c, e))) =
-    (if eqop eq_nat fm 0
-      then (fn k => Ge (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Ge (Cn (suc (minus_nat fm 1), c, e))))
+    (if ((fm : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             Ge (Cn ((0 : IntInf.int), (1 : IntInf.int), Mul (div_int k c, e))))
+      else (fn _ => Ge (Cn (suc (minus_nat fm (1 : IntInf.int)), c, e))))
   | a_beta (Eq (Cn (gm, c, e))) =
-    (if eqop eq_nat gm 0
-      then (fn k => Eq (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Eq (Cn (suc (minus_nat gm 1), c, e))))
+    (if ((gm : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             Eq (Cn ((0 : IntInf.int), (1 : IntInf.int), Mul (div_int k c, e))))
+      else (fn _ => Eq (Cn (suc (minus_nat gm (1 : IntInf.int)), c, e))))
   | a_beta (NEq (Cn (hm, c, e))) =
-    (if eqop eq_nat hm 0
-      then (fn k => NEq (Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => NEq (Cn (suc (minus_nat hm 1), c, e))))
+    (if ((hm : IntInf.int) = (0 : IntInf.int))
+      then (fn k =>
+             NEq (Cn ((0 : IntInf.int), (1 : IntInf.int),
+                       Mul (div_int 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 eqop eq_nat im 0
+    (if ((im : IntInf.int) = (0 : IntInf.int))
       then (fn k =>
              Dvd (IntInf.* (div_int k c, i),
-                   Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => Dvd (i, Cn (suc (minus_nat im 1), c, e))))
+                   Cn ((0 : IntInf.int), (1 : IntInf.int),
+                        Mul (div_int 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 eqop eq_nat jm 0
+    (if ((jm : IntInf.int) = (0 : IntInf.int))
       then (fn k =>
              NDvd (IntInf.* (div_int k c, i),
-                    Cn (0, (1 : IntInf.int), Mul (div_int k c, e))))
-      else (fn k => NDvd (i, Cn (suc (minus_nat jm 1), c, e))));
+                    Cn ((0 : IntInf.int), (1 : IntInf.int),
+                         Mul (div_int k c, e))))
+      else (fn _ => NDvd (i, Cn (suc (minus_nat jm (1 : IntInf.int)), c, e))));
 
-fun zeta (And (p, q)) = zlcm (zeta p) (zeta q)
-  | zeta (Or (p, q)) = zlcm (zeta p) (zeta q)
+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)
@@ -1367,64 +1960,59 @@
   | zeta (Closed ap) = (1 : IntInf.int)
   | zeta (NClosed aq) = (1 : IntInf.int)
   | zeta (Lt (Cn (cm, c, e))) =
-    (if eqop eq_nat cm 0 then c else (1 : IntInf.int))
+    (if ((cm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (Le (Cn (dm, c, e))) =
-    (if eqop eq_nat dm 0 then c else (1 : IntInf.int))
+    (if ((dm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (Gt (Cn (em, c, e))) =
-    (if eqop eq_nat em 0 then c else (1 : IntInf.int))
+    (if ((em : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (Ge (Cn (fm, c, e))) =
-    (if eqop eq_nat fm 0 then c else (1 : IntInf.int))
+    (if ((fm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (Eq (Cn (gm, c, e))) =
-    (if eqop eq_nat gm 0 then c else (1 : IntInf.int))
+    (if ((gm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (NEq (Cn (hm, c, e))) =
-    (if eqop eq_nat hm 0 then c else (1 : IntInf.int))
+    (if ((hm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (Dvd (i, Cn (im, c, e))) =
-    (if eqop eq_nat im 0 then c else (1 : IntInf.int))
+    (if ((im : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int))
   | zeta (NDvd (i, Cn (jm, c, e))) =
-    (if eqop eq_nat jm 0 then c else (1 : IntInf.int));
+    (if ((jm : IntInf.int) = (0 : IntInf.int)) then c else (1 : IntInf.int));
 
 fun zsplit0 (C c) = ((0 : IntInf.int), C c)
   | zsplit0 (Bound n) =
-    (if eqop eq_nat n 0 then ((1 : IntInf.int), C (0 : IntInf.int))
+    (if ((n : IntInf.int) = (0 : IntInf.int))
+      then ((1 : IntInf.int), C (0 : IntInf.int))
       else ((0 : IntInf.int), Bound n))
   | zsplit0 (Cn (n, i, a)) =
     let
-      val aa = zsplit0 a;
-      val (i', a') = aa;
+      val (ia, aa) = zsplit0 a;
     in
-      (if eqop eq_nat n 0 then (IntInf.+ (i, i'), a') else (i', Cn (n, i, a')))
+      (if ((n : IntInf.int) = (0 : IntInf.int)) then (IntInf.+ (i, ia), aa)
+        else (ia, Cn (n, i, aa)))
     end
   | zsplit0 (Neg a) =
     let
-      val aa = zsplit0 a;
-      val (i', a') = aa;
+      val (i, aa) = zsplit0 a;
     in
-      (IntInf.~ i', Neg a')
+      (IntInf.~ i, Neg aa)
     end
   | zsplit0 (Add (a, b)) =
     let
-      val aa = zsplit0 a;
-      val (ia, a') = aa;
-      val ab = zsplit0 b;
-      val (ib, b') = ab;
+      val (ia, aa) = zsplit0 a;
+      val (ib, ba) = zsplit0 b;
     in
-      (IntInf.+ (ia, ib), Add (a', b'))
+      (IntInf.+ (ia, ib), Add (aa, ba))
     end
   | zsplit0 (Sub (a, b)) =
     let
-      val aa = zsplit0 a;
-      val (ia, a') = aa;
-      val ab = zsplit0 b;
-      val (ib, b') = ab;
+      val (ia, aa) = zsplit0 a;
+      val (ib, ba) = zsplit0 b;
     in
-      (IntInf.- (ia, ib), Sub (a', b'))
+      (IntInf.- (ia, ib), Sub (aa, ba))
     end
   | zsplit0 (Mul (i, a)) =
     let
-      val aa = zsplit0 a;
-      val (i', a') = aa;
+      val (ia, aa) = zsplit0 a;
     in
-      (IntInf.* (i, i'), Mul (i, a'))
+      (IntInf.* (i, ia), Mul (i, aa))
     end;
 
 fun zlfm (And (p, q)) = And (zlfm p, zlfm q)
@@ -1434,79 +2022,79 @@
     Or (And (zlfm p, zlfm q), And (zlfm (Not p), zlfm (Not q)))
   | zlfm (Lt a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then Lt r
-        else (if IntInf.< ((0 : IntInf.int), c) then Lt (Cn (0, c, r))
-               else Gt (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then Lt r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then Lt (Cn ((0 : IntInf.int), c, r))
+               else Gt (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (Le a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then Le r
-        else (if IntInf.< ((0 : IntInf.int), c) then Le (Cn (0, c, r))
-               else Ge (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then Le r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then Le (Cn ((0 : IntInf.int), c, r))
+               else Ge (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (Gt a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then Gt r
-        else (if IntInf.< ((0 : IntInf.int), c) then Gt (Cn (0, c, r))
-               else Lt (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then Gt r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then Gt (Cn ((0 : IntInf.int), c, r))
+               else Lt (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (Ge a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then Ge r
-        else (if IntInf.< ((0 : IntInf.int), c) then Ge (Cn (0, c, r))
-               else Le (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then Ge r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then Ge (Cn ((0 : IntInf.int), c, r))
+               else Le (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (Eq a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then Eq r
-        else (if IntInf.< ((0 : IntInf.int), c) then Eq (Cn (0, c, r))
-               else Eq (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then Eq r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then Eq (Cn ((0 : IntInf.int), c, r))
+               else Eq (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (NEq a) =
     let
-      val aa = zsplit0 a;
-      val (c, r) = aa;
+      val (c, r) = zsplit0 a;
     in
-      (if eqop eq_int c (0 : IntInf.int) then NEq r
-        else (if IntInf.< ((0 : IntInf.int), c) then NEq (Cn (0, c, r))
-               else NEq (Cn (0, IntInf.~ c, Neg r))))
+      (if ((c : IntInf.int) = (0 : IntInf.int)) then NEq r
+        else (if IntInf.< ((0 : IntInf.int), c)
+               then NEq (Cn ((0 : IntInf.int), c, r))
+               else NEq (Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
     end
   | zlfm (Dvd (i, a)) =
-    (if eqop eq_int i (0 : IntInf.int) then zlfm (Eq a)
+    (if ((i : IntInf.int) = (0 : IntInf.int)) then zlfm (Eq a)
       else let
-             val aa = zsplit0 a;
-             val (c, r) = aa;
+             val (c, r) = zsplit0 a;
            in
-             (if eqop eq_int c (0 : IntInf.int) then Dvd (abs_int i, r)
+             (if ((c : IntInf.int) = (0 : IntInf.int)) then Dvd (abs_int i, r)
                else (if IntInf.< ((0 : IntInf.int), c)
-                      then Dvd (abs_int i, Cn (0, c, r))
-                      else Dvd (abs_int i, Cn (0, IntInf.~ c, Neg r))))
+                      then Dvd (abs_int i, Cn ((0 : IntInf.int), c, r))
+                      else Dvd (abs_int i,
+                                 Cn ((0 : IntInf.int), IntInf.~ c, Neg r))))
            end)
   | zlfm (NDvd (i, a)) =
-    (if eqop eq_int i (0 : IntInf.int) then zlfm (NEq a)
+    (if ((i : IntInf.int) = (0 : IntInf.int)) then zlfm (NEq a)
       else let
-             val aa = zsplit0 a;
-             val (c, r) = aa;
+             val (c, r) = zsplit0 a;
            in
-             (if eqop eq_int c (0 : IntInf.int) then NDvd (abs_int i, r)
+             (if ((c : IntInf.int) = (0 : IntInf.int)) then NDvd (abs_int i, r)
                else (if IntInf.< ((0 : IntInf.int), c)
-                      then NDvd (abs_int i, Cn (0, c, r))
-                      else NDvd (abs_int i, Cn (0, IntInf.~ c, Neg r))))
+                      then NDvd (abs_int i, Cn ((0 : IntInf.int), c, r))
+                      else NDvd (abs_int i,
+                                  Cn ((0 : IntInf.int), IntInf.~ 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))
@@ -1537,10 +2125,11 @@
 
 fun unita p =
   let
-    val p' = zlfm p;
-    val l = zeta p';
+    val pa = zlfm p;
+    val l = zeta pa;
     val q =
-      And (Dvd (l, Cn (0, (1 : IntInf.int), C (0 : IntInf.int))), a_beta p' l);
+      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 eq_numa (map simpnum (beta q));
     val a = remdups eq_numa (map simpnum (alpha q));
@@ -1551,18 +2140,16 @@
 
 fun cooper p =
   let
-    val a = unita p;
-    val (q, aa) = a;
-    val (b, d) = aa;
+    val (q, (b, d)) = unita p;
     val js = iupt (1 : IntInf.int) d;
     val mq = simpfm (minusinf q);
     val md = evaldjf (fn j => simpfm (subst0 (C j) mq)) js;
   in
-    (if eqop eq_fma md T then T
+    (if eq_fm md T then T
       else let
              val qd =
-               evaldjf (fn ab as (ba, j) => simpfm (subst0 (Add (ba, C j)) q))
-                 (concat (map (fn ba => map (fn ab => (ba, ab)) js) b));
+               evaldjf (fn (ba, j) => simpfm (subst0 (Add (ba, C j)) q))
+                 (concat_map (fn ba => map (fn a => (ba, a)) js) b);
            in
              decr (disj md qd)
            end)
@@ -1669,37 +2256,19 @@
   | 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 y => simpfm T)
-  | qelim F = (fn y => simpfm F)
-  | qelim (Lt u) = (fn y => simpfm (Lt u))
-  | qelim (Le v) = (fn y => simpfm (Le v))
-  | qelim (Gt w) = (fn y => simpfm (Gt w))
-  | qelim (Ge x) = (fn y => simpfm (Ge x))
-  | qelim (Eq y) = (fn ya => simpfm (Eq y))
-  | qelim (NEq z) = (fn y => simpfm (NEq z))
-  | qelim (Dvd (aa, ab)) = (fn y => simpfm (Dvd (aa, ab)))
-  | qelim (NDvd (ac, ad)) = (fn y => simpfm (NDvd (ac, ad)))
-  | qelim (Closed ap) = (fn y => simpfm (Closed ap))
-  | qelim (NClosed aq) = (fn y => simpfm (NClosed aq));
+  | qelim T = (fn _ => simpfm T)
+  | qelim F = (fn _ => simpfm F)
+  | qelim (Lt u) = (fn _ => simpfm (Lt u))
+  | qelim (Le v) = (fn _ => simpfm (Le v))
+  | qelim (Gt w) = (fn _ => simpfm (Gt w))
+  | qelim (Ge x) = (fn _ => simpfm (Ge x))
+  | qelim (Eq y) = (fn _ => simpfm (Eq y))
+  | qelim (NEq z) = (fn _ => simpfm (NEq z))
+  | qelim (Dvd (aa, ab)) = (fn _ => simpfm (Dvd (aa, ab)))
+  | qelim (NDvd (ac, ad)) = (fn _ => simpfm (NDvd (ac, ad)))
+  | qelim (Closed ap) = (fn _ => simpfm (Closed ap))
+  | qelim (NClosed aq) = (fn _ => simpfm (NClosed aq));
 
 fun pa p = qelim (prep p) cooper;
 
-fun neg z = IntInf.< (z, (0 : IntInf.int));
-
-fun nat_aux i n =
-  (if IntInf.<= (i, (0 : IntInf.int)) then n
-    else nat_aux (IntInf.- (i, (1 : IntInf.int))) (suc n));
-
-fun adjust b =
-  (fn a as (q, r) =>
-    (if IntInf.<= ((0 : IntInf.int), IntInf.- (r, b))
-      then (IntInf.+ (IntInf.* ((2 : IntInf.int), q), (1 : IntInf.int)),
-             IntInf.- (r, b))
-      else (IntInf.* ((2 : IntInf.int), q), r)));
-
-fun posDivAlg a b =
-  (if IntInf.< (a, b) orelse IntInf.<= (b, (0 : IntInf.int))
-    then ((0 : IntInf.int), a)
-    else adjust b (posDivAlg a (IntInf.* ((2 : IntInf.int), b))));
-
-end; (*struct GeneratedCooper*)
+end; (*struct Generated_Cooper*)
--- a/src/HOL/Unix/Unix.thy	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/HOL/Unix/Unix.thy	Thu Apr 29 11:42:34 2010 -0700
@@ -358,7 +358,7 @@
   read:
     "access root path uid {Readable} = Some (Val (att, text)) \<Longrightarrow>
       root \<midarrow>(Read uid text path)\<rightarrow> root" |
-  write:
+  "write":
     "access root path uid {Writable} = Some (Val (att, text')) \<Longrightarrow>
       root \<midarrow>(Write uid text path)\<rightarrow> update path (Some (Val (att, text))) root" |
 
@@ -436,7 +436,7 @@
   case read
   with root' show ?thesis by cases auto
 next
-  case write
+  case "write"
   with root' show ?thesis by cases auto
 next
   case chmod
--- a/src/Pure/Isar/args.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Isar/args.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -205,7 +205,7 @@
   >> (fn Type (c, _) => c | TFree (a, _) => a | _ => "");
 
 fun const strict =
-  Scan.peek (fn ctxt => named_term (ProofContext.read_const (Context.proof_of ctxt) strict))
+  Scan.peek (fn ctxt => named_term (ProofContext.read_const (Context.proof_of ctxt) strict dummyT))
   >> (fn Const (c, _) => c | Free (x, _) => x | _ => "");
 
 fun const_proper strict =
--- a/src/Pure/Isar/isar_syn.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Isar/isar_syn.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -225,22 +225,22 @@
     >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
 
 val _ =
-  OuterSyntax.local_theory "type_notation" "add notation for type constructors" K.thy_decl
+  OuterSyntax.local_theory "type_notation" "add concrete syntax for type constructors" K.thy_decl
     (opt_mode -- P.and_list1 (P.xname -- P.mixfix)
     >> (fn (mode, args) => Specification.type_notation_cmd true mode args));
 
 val _ =
-  OuterSyntax.local_theory "no_type_notation" "delete notation for type constructors" K.thy_decl
+  OuterSyntax.local_theory "no_type_notation" "delete concrete syntax for type constructors" K.thy_decl
     (opt_mode -- P.and_list1 (P.xname -- P.mixfix)
     >> (fn (mode, args) => Specification.type_notation_cmd false mode args));
 
 val _ =
-  OuterSyntax.local_theory "notation" "add notation for constants / fixed variables" K.thy_decl
+  OuterSyntax.local_theory "notation" "add concrete syntax for constants / fixed variables" K.thy_decl
     (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
     >> (fn (mode, args) => Specification.notation_cmd true mode args));
 
 val _ =
-  OuterSyntax.local_theory "no_notation" "delete notation for constants / fixed variables" K.thy_decl
+  OuterSyntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables" K.thy_decl
     (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
     >> (fn (mode, args) => Specification.notation_cmd false mode args));
 
@@ -615,6 +615,12 @@
     (P.and_list1 (P.and_list1 P.term -- (P.$$$ "=" |-- P.term))
       >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));
 
+val _ =
+  OuterSyntax.command "write" "add concrete syntax for constants / fixed variables"
+    (K.tag_proof K.prf_decl)
+    (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
+    >> (fn (mode, args) => Toplevel.print o Toplevel.proof (Proof.write_cmd mode args)));
+
 val case_spec =
   (P.$$$ "(" |-- P.!!! (P.xname -- Scan.repeat1 (P.maybe P.name) --| P.$$$ ")") ||
     P.xname >> rpair []) -- SpecParse.opt_attribs >> P.triple1;
--- a/src/Pure/Isar/proof.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Isar/proof.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -43,6 +43,8 @@
   val simple_goal: state -> {context: context, goal: thm}
   val let_bind: (term list * term) list -> state -> state
   val let_bind_cmd: (string list * string) list -> state -> state
+  val write: Syntax.mode -> (term * mixfix) list -> state -> state
+  val write_cmd: Syntax.mode -> (string * mixfix) list -> state -> state
   val fix: (binding * typ option * mixfix) list -> state -> state
   val fix_cmd: (binding * string option * mixfix) list -> state -> state
   val assm: Assumption.export ->
@@ -539,6 +541,26 @@
 end;
 
 
+(* concrete syntax *)
+
+local
+
+fun gen_write prep_arg mode args =
+  assert_forward
+  #> map_context (fn ctxt => ctxt |> ProofContext.notation true mode (map (prep_arg ctxt) args))
+  #> put_facts NONE;
+
+in
+
+val write = gen_write (K I);
+
+val write_cmd =
+  gen_write (fn ctxt => fn (c, mx) =>
+    (ProofContext.read_const ctxt false (Syntax.mixfixT mx) c, mx));
+
+end;
+
+
 (* fix *)
 
 local
--- a/src/Pure/Isar/proof_context.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Isar/proof_context.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -55,13 +55,13 @@
   val cert_typ_abbrev: Proof.context -> typ -> typ
   val get_skolem: Proof.context -> string -> string
   val revert_skolem: Proof.context -> string -> string
-  val infer_type: Proof.context -> string -> typ
+  val infer_type: Proof.context -> string * typ -> typ
   val inferred_param: string -> Proof.context -> typ * Proof.context
   val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
   val read_type_name: Proof.context -> bool -> string -> typ
   val read_type_name_proper: Proof.context -> bool -> string -> typ
   val read_const_proper: Proof.context -> bool -> string -> term
-  val read_const: Proof.context -> bool -> string -> term
+  val read_const: Proof.context -> bool -> typ -> string -> term
   val allow_dummies: Proof.context -> Proof.context
   val check_tvar: Proof.context -> indexname * sort -> indexname * sort
   val check_tfree: Proof.context -> string * sort -> string * sort
@@ -438,11 +438,10 @@
 (* inferred types of parameters *)
 
 fun infer_type ctxt x =
-  Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt))
-    (Free (x, dummyT)));
+  Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt)) (Free x));
 
 fun inferred_param x ctxt =
-  let val T = infer_type ctxt x
+  let val T = infer_type ctxt (x, dummyT)
   in (T, ctxt |> Variable.declare_term (Free (x, T))) end;
 
 fun inferred_fixes ctxt =
@@ -505,13 +504,13 @@
 
 fun read_const_proper ctxt strict = prep_const_proper ctxt strict o token_content;
 
-fun read_const ctxt strict text =
+fun read_const ctxt strict ty text =
   let val (c, pos) = token_content text in
     (case (lookup_skolem ctxt c, Variable.is_const ctxt c) of
       (SOME x, false) =>
         (Position.report (Markup.name x
             (if can Name.dest_skolem x then Markup.skolem else Markup.free)) pos;
-          Free (x, infer_type ctxt x))
+          Free (x, infer_type ctxt (x, ty)))
     | _ => prep_const_proper ctxt strict (c, pos))
   end;
 
@@ -738,7 +737,7 @@
   let
     val (syms, pos) = Syntax.parse_token Markup.sort text;
     val S = Syntax.standard_parse_sort ctxt (syn_of ctxt) (syms, pos)
-      handle ERROR msg => cat_error msg  ("Failed to parse sort" ^ Position.str_of pos)
+      handle ERROR msg => cat_error msg ("Failed to parse sort" ^ Position.str_of pos)
   in Type.minimize_sort (tsig_of ctxt) S end;
 
 fun parse_typ ctxt text =
@@ -1175,16 +1174,6 @@
 
 (* fixes *)
 
-local
-
-fun prep_mixfix (x, T, mx) =
-  if mx <> NoSyn andalso mx <> Structure andalso
-      (can Name.dest_internal x orelse can Name.dest_skolem x) then
-    error ("Illegal mixfix syntax for internal/skolem constant " ^ quote x)
-  else (Local_Syntax.Fixed, (x, T, mx));
-
-in
-
 fun add_fixes raw_vars ctxt =
   let
     val (vars, _) = cert_vars raw_vars ctxt;
@@ -1192,13 +1181,11 @@
     val ctxt'' =
       ctxt'
       |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)
-      |-> (map_syntax o Local_Syntax.add_syntax (theory_of ctxt) o map prep_mixfix);
+      |-> (map_syntax o Local_Syntax.add_syntax (theory_of ctxt) o map (pair Local_Syntax.Fixed));
     val _ = (vars ~~ xs') |> List.app (fn ((b, _, _), x') =>
       Context_Position.report_visible ctxt (Markup.fixed_decl x') (Binding.pos_of b));
   in (xs', ctxt'') end;
 
-end;
-
 
 (* fixes vs. frees *)
 
--- a/src/Pure/Isar/specification.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Isar/specification.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -284,7 +284,7 @@
 val type_notation_cmd = gen_type_notation (fn ctxt => ProofContext.read_type_name ctxt false);
 
 val notation = gen_notation (K I);
-val notation_cmd = gen_notation (fn ctxt => ProofContext.read_const ctxt false);
+val notation_cmd = gen_notation (fn ctxt => ProofContext.read_const ctxt false dummyT);
 
 end;
 
--- a/src/Pure/Syntax/syn_trans.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Syntax/syn_trans.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -567,14 +567,7 @@
           Term.list_comb (term_of ast, map term_of asts)
       | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
 
-    val free_fixed = Term.map_aterms
-      (fn t as Const (c, T) =>
-          (case try Lexicon.unmark_fixed c of
-            NONE => t
-          | SOME x => Free (x, T))
-        | t => t);
-
-    val exn_results = map (Exn.capture (term_of #> free_fixed)) asts;
+    val exn_results = map (Exn.capture term_of) asts;
     val exns = map_filter Exn.get_exn exn_results;
     val results = map_filter Exn.get_result exn_results
   in (case (results, exns) of ([], exn :: _) => reraise exn | _ => results) end;
--- a/src/Pure/Syntax/type_ext.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Pure/Syntax/type_ext.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -120,11 +120,14 @@
       | decode (Abs (x, T, t)) = Abs (x, T, decode t)
       | decode (t $ u) = decode t $ decode u
       | decode (Const (a, T)) =
-          let val c =
-            (case try Lexicon.unmark_const a of
-              SOME c => c
-            | NONE => snd (map_const a))
-          in Const (c, T) end
+          (case try Lexicon.unmark_fixed a of
+            SOME x => Free (x, T)
+          | NONE =>
+              let val c =
+                (case try Lexicon.unmark_const a of
+                  SOME c => c
+                | NONE => snd (map_const a))
+              in Const (c, T) end)
       | decode (Free (a, T)) =
           (case (map_free a, map_const a) of
             (SOME x, _) => Free (x, T)
--- a/src/Tools/Code/code_eval.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Tools/Code/code_eval.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      Tools/code/code_eval.ML_
+(*  Title:      Tools/code/code_eval.ML
     Author:     Florian Haftmann, TU Muenchen
 
 Runtime services building on code generation into implementation language SML.
@@ -97,19 +97,6 @@
 fun print_const const all_struct_name tycos_map consts_map =
   (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const;
 
-fun print_datatype tyco constrs all_struct_name tycos_map consts_map =
-  let
-    val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode;
-    fun check_base name name'' =
-      if upperize (Long_Name.base_name name) = upperize name''
-      then () else error ("Name as printed " ^ quote name''
-        ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry.");
-    val tyco'' = (the o AList.lookup (op =) tycos_map) tyco;
-    val constrs'' = map (the o AList.lookup (op =) consts_map) constrs;
-    val _ = check_base tyco tyco'';
-    val _ = map2 check_base constrs constrs'';
-  in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end;
-
 fun print_code is_first print_it ctxt =
   let
     val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt;
@@ -128,18 +115,6 @@
     val background' = register_const const background;
   in (print_code is_first (print_const const), background') end;
 
-fun ml_code_datatype_antiq (raw_tyco, raw_constrs) background =
-  let
-    val thy = ProofContext.theory_of background;
-    val tyco = Sign.intern_type thy raw_tyco;
-    val constrs = map (Code.check_const thy) raw_constrs;
-    val constrs' = (map fst o snd o Code.get_type thy) tyco;
-    val _ = if eq_set (op =) (constrs, constrs') then ()
-      else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
-    val is_first = is_first_occ background;
-    val background' = register_datatype tyco constrs background;
-  in (print_code is_first (print_datatype tyco constrs), background') end;
-
 end; (*local*)
 
 
@@ -171,10 +146,20 @@
     |> Code_Target.add_syntax_tyco target tyco (SOME (k, pr))
   end;
 
+fun add_eval_constr (const, const') thy =
+  let
+    val k = Code.args_number thy const;
+    fun pr pr' fxy ts = Code_Printer.brackify fxy
+      (const' :: the_list (Code_ML.print_tuple pr' Code_Printer.BR (map fst ts)));
+  in
+    thy
+    |> Code_Target.add_syntax_const target const (SOME (Code_Printer.simple_const_syntax (k, pr)))
+  end;
+
 fun add_eval_const (const, const') = Code_Target.add_syntax_const target
   const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const')));
 
-fun process (code_body, (tyco_map, const_map)) module_name NONE thy =
+fun process (code_body, (tyco_map, (constr_map, const_map))) module_name NONE thy =
       let
         val pr = Code_Printer.str o Long_Name.append module_name;
       in
@@ -182,6 +167,7 @@
         |> Code_Target.add_reserved target module_name
         |> Context.theory_map (ML_Context.exec (fn () => ML_Context.eval true Position.none code_body))
         |> fold (add_eval_tyco o apsnd pr) tyco_map
+        |> fold (add_eval_constr o apsnd pr) constr_map
         |> fold (add_eval_const o apsnd pr) const_map
       end
   | process (code_body, _) _ (SOME file_name) thy =
@@ -197,11 +183,15 @@
   let
     val datatypes = map (fn (raw_tyco, raw_cos) =>
       (prep_type thy raw_tyco, map (prep_const thy) raw_cos)) raw_datatypes;
+    val _ = map (uncurry (check_datatype thy)) datatypes;
+    val tycos = map fst datatypes;
+    val constrs = maps snd datatypes;
     val functions = map (prep_const thy) raw_functions;
-    val _ = map (uncurry (check_datatype thy)) datatypes;
+    val result = evaluation_code thy module_name tycos (constrs @ functions)
+      |> (apsnd o apsnd) (chop (length constrs));
   in
     thy
-    |> process (evaluation_code thy module_name (map fst datatypes) (maps snd datatypes @ functions)) module_name some_file
+    |> process result module_name some_file
   end;
 
 val code_reflect = gen_code_reflect Code_Target.cert_tyco Code.check_const;
@@ -211,10 +201,6 @@
 (** Isar setup **)
 
 val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq);
-val _ = ML_Context.add_antiq "code_datatype" (fn _ =>
-  (Args.type_name true --| Scan.lift (Args.$$$ "=")
-    -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term)))
-      >> ml_code_datatype_antiq);
 
 local
 
@@ -223,7 +209,6 @@
 
 val datatypesK = "datatypes";
 val functionsK = "functions";
-val module_nameK = "module_name";
 val fileK = "file";
 val andK = "and"
 
@@ -235,12 +220,11 @@
 
 val _ =
   OuterSyntax.command "code_reflect" "enrich runtime environment with generated code"
-    K.thy_decl (Scan.optional (P.$$$ datatypesK |-- (parse_datatype
+    K.thy_decl (P.name -- Scan.optional (P.$$$ datatypesK |-- (parse_datatype
       ::: Scan.repeat (P.$$$ andK |-- parse_datatype))) []
     -- Scan.optional (P.$$$ functionsK |-- Scan.repeat1 P.name) []
-    --| P.$$$ module_nameK -- P.name
     -- Scan.option (P.$$$ fileK |-- P.name)
-  >> (fn (((raw_datatypes, raw_functions), module_name), some_file) => Toplevel.theory
+  >> (fn (((module_name, raw_datatypes), raw_functions), some_file) => Toplevel.theory
     (code_reflect_cmd raw_datatypes raw_functions module_name some_file)));
 
 end; (*local*)
--- a/src/Tools/Code/code_haskell.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Tools/Code/code_haskell.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -309,10 +309,10 @@
 
 fun serialize_haskell module_prefix raw_module_name string_classes labelled_name
     raw_reserved includes raw_module_alias
-    syntax_class syntax_tyco syntax_const (code_of_pretty, code_writeln) program cs destination =
+    syntax_class syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
   let
-    val stmt_names = Code_Target.stmt_names_of_destination destination;
-    val module_name = if null stmt_names then raw_module_name else SOME "Code";
+    val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
+    val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
     val reserved = fold (insert (op =) o fst) includes raw_reserved;
     val (deresolver, hs_program) = haskell_program_of_program labelled_name
       module_name module_prefix reserved raw_module_alias program;
@@ -365,13 +365,13 @@
           );
       in print_module module_name' content end;
     fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks2 (map_filter
-        (fn (name, (_, SOME stmt)) => if null stmt_names
-              orelse member (op =) stmt_names name
+        (fn (name, (_, SOME stmt)) => if null presentation_stmt_names
+              orelse member (op =) presentation_stmt_names name
               then SOME (print_stmt false (name, stmt))
               else NONE
           | (_, (_, NONE)) => NONE) stmts);
     val serialize_module =
-      if null stmt_names then serialize_module1 else pair "" o serialize_module2;
+      if null presentation_stmt_names then serialize_module1 else pair "" o serialize_module2;
     fun check_destination destination =
       (File.check destination; destination);
     fun write_module destination (modlname, content) =
--- a/src/Tools/Code/code_ml.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Tools/Code/code_ml.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      Tools/code/code_ml.ML_
+(*  Title:      Tools/code/code_ml.ML
     Author:     Florian Haftmann, TU Muenchen
 
 Serializer for SML and OCaml.
@@ -9,6 +9,8 @@
   val target_SML: string
   val evaluation_code_of: theory -> string -> string
     -> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
+  val print_tuple: (Code_Printer.fixity -> 'a -> Pretty.T)
+    -> Code_Printer.fixity -> 'a list -> Pretty.T option
   val setup: theory -> theory
 end;
 
--- a/src/Tools/Code/code_scala.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Tools/Code/code_scala.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -340,10 +340,10 @@
 
 fun serialize_scala raw_module_name labelled_name
     raw_reserved includes raw_module_alias
-    _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program cs destination =
+    _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
   let
-    val stmt_names = Code_Target.stmt_names_of_destination destination;
-    val module_name = if null stmt_names then raw_module_name else SOME "Code";
+    val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
+    val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
     val reserved = fold (insert (op =) o fst) includes raw_reserved;
     val (deresolver, (the_module_name, sca_program)) = scala_program_of_program labelled_name
       module_name reserved raw_module_alias program;
--- a/src/Tools/Code/code_target.ML	Thu Apr 29 11:41:04 2010 -0700
+++ b/src/Tools/Code/code_target.ML	Thu Apr 29 11:42:34 2010 -0700
@@ -279,7 +279,7 @@
       (Symtab.lookup module_alias) (Symtab.lookup class')
       (Symtab.lookup tyco') (Symtab.lookup const')
       (Code_Printer.string_of_pretty width, Code_Printer.writeln_pretty width)
-      program4 names2
+      program4 names1
   end;
 
 fun mount_serializer thy alt_serializer target some_width module args naming program names =