merged
authorhuffman
Tue, 02 Mar 2010 04:35:44 -0800
changeset 35498 5c70de748522
parent 35497 979706bd5c16 (current diff)
parent 35440 bdf8ad377877 (diff)
child 35499 6acef0aea07d
merged
src/HOL/Rational.thy
src/HOLCF/Tools/Domain/domain_extender.ML
src/HOLCF/Tools/Domain/domain_isomorphism.ML
--- a/Admin/isatest/isatest-makedist	Tue Mar 02 04:31:50 2010 -0800
+++ b/Admin/isatest/isatest-makedist	Tue Mar 02 04:35:44 2010 -0800
@@ -91,8 +91,10 @@
 
 ## spawn test runs
 
+$SSH macbroy21 "$MAKEALL $HOME/settings/at-poly-test"
+# give test some time to copy settings and start
+sleep 15
 $SSH macbroy22 "$MAKEALL $HOME/settings/at-poly"
-# give test some time to copy settings and start
 sleep 15
 $SSH macbroy23 "$MAKEALL -l HOL HOL-ex $HOME/settings/at-sml-dev-e"
 sleep 15
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Admin/isatest/settings/at-poly-test	Tue Mar 02 04:35:44 2010 -0800
@@ -0,0 +1,27 @@
+# -*- shell-script -*- :mode=shellscript:
+
+  POLYML_HOME="/home/polyml/polyml-5.3.0"
+  ML_SYSTEM="polyml-5.3.0"
+  ML_PLATFORM="x86-linux"
+  ML_HOME="$POLYML_HOME/$ML_PLATFORM"
+  ML_OPTIONS="-H 500"
+
+ISABELLE_HOME_USER=~/isabelle-at-poly-test
+
+# Where to look for isabelle tools (multiple dirs separated by ':').
+ISABELLE_TOOLS="$ISABELLE_HOME/lib/Tools"
+
+# Location for temporary files (should be on a local file system).
+ISABELLE_TMP_PREFIX="/tmp/isabelle-$USER"
+
+
+# Heap input locations. ML system identifier is included in lookup.
+ISABELLE_PATH="$ISABELLE_HOME_USER/heaps:$ISABELLE_HOME/heaps"
+
+# Heap output location. ML system identifier is appended automatically later on.
+ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
+ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
+
+ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -t true"
+
+init_component /home/isabelle/contrib_devel/kodkodi
--- a/NEWS	Tue Mar 02 04:31:50 2010 -0800
+++ b/NEWS	Tue Mar 02 04:35:44 2010 -0800
@@ -38,6 +38,8 @@
 and ML_command "set Syntax.trace_ast" help to diagnose syntax
 problems.
 
+* Type constructors admit general mixfix syntax, not just infix.
+
 
 *** Pure ***
 
@@ -49,9 +51,16 @@
 datatype constructors have been renamed from InfixName to Infix etc.
 Minor INCOMPATIBILITY.
 
+* Commands 'type_notation' and 'no_type_notation' declare type syntax
+within a local theory context, with explicit checking of the
+constructors involved (in contrast to the raw 'syntax' versions).
+
 
 *** HOL ***
 
+* Theory "Rational" renamed to "Rat", for consistency with "Nat", "Int" etc.
+INCOMPATIBILITY.
+
 * New set of rules "ac_simps" provides combined assoc / commute rewrites
 for all interpretations of the appropriate generic locales.
 
@@ -172,6 +181,20 @@
 
 *** ML ***
 
+* Antiquotations for basic formal entities:
+
+    @{class NAME}         -- type class
+    @{class_syntax NAME}  -- syntax representation of the above
+
+    @{type_name NAME}     -- logical type
+    @{type_abbrev NAME}   -- type abbreviation
+    @{nonterminal NAME}   -- type of concrete syntactic category
+    @{type_syntax NAME}   -- syntax representation of any of the above
+
+    @{const_name NAME}    -- logical constant (INCOMPATIBILITY)
+    @{const_abbrev NAME}  -- abbreviated constant
+    @{const_syntax NAME}  -- syntax representation of any of the above
+
 * Antiquotation @{syntax_const NAME} ensures that NAME refers to a raw
 syntax constant (cf. 'syntax' command).
 
--- a/doc-src/IsarImplementation/Thy/Prelim.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -436,14 +436,14 @@
     val empty = [];
     val extend = I;
     fun merge (ts1, ts2) =
-      OrdList.union TermOrd.fast_term_ord ts1 ts2;
+      OrdList.union Term_Ord.fast_term_ord ts1 ts2;
   )
 
   val get = Terms.get;
 
   fun add raw_t thy =
     let val t = Sign.cert_term thy raw_t
-    in Terms.map (OrdList.insert TermOrd.fast_term_ord t) thy end;
+    in Terms.map (OrdList.insert Term_Ord.fast_term_ord t) thy end;
 
   end;
 *}
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex	Tue Mar 02 04:35:44 2010 -0800
@@ -544,14 +544,14 @@
 \ \ \ \ val\ empty\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline
 \ \ \ \ val\ extend\ {\isacharequal}\ I{\isacharsemicolon}\isanewline
 \ \ \ \ fun\ merge\ {\isacharparenleft}ts{\isadigit{1}}{\isacharcomma}\ ts{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\isanewline
-\ \ \ \ \ \ OrdList{\isachardot}union\ TermOrd{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ ts{\isadigit{1}}\ ts{\isadigit{2}}{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ OrdList{\isachardot}union\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ ts{\isadigit{1}}\ ts{\isadigit{2}}{\isacharsemicolon}\isanewline
 \ \ {\isacharparenright}\isanewline
 \isanewline
 \ \ val\ get\ {\isacharequal}\ Terms{\isachardot}get{\isacharsemicolon}\isanewline
 \isanewline
 \ \ fun\ add\ raw{\isacharunderscore}t\ thy\ {\isacharequal}\isanewline
 \ \ \ \ let\ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t\isanewline
-\ \ \ \ in\ Terms{\isachardot}map\ {\isacharparenleft}OrdList{\isachardot}insert\ TermOrd{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\ end{\isacharsemicolon}\isanewline
+\ \ \ \ in\ Terms{\isachardot}map\ {\isacharparenleft}OrdList{\isachardot}insert\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\ end{\isacharsemicolon}\isanewline
 \isanewline
 \ \ end{\isacharsemicolon}\isanewline
 {\isacharverbatimclose}%
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -13,14 +13,14 @@
   \end{matharray}
 
   \begin{rail}
-    'typedecl' typespec infix?
+    'typedecl' typespec mixfix?
     ;
     'typedef' altname? abstype '=' repset
     ;
 
     altname: '(' (name | 'open' | 'open' name) ')'
     ;
-    abstype: typespec infix?
+    abstype: typespec mixfix?
     ;
     repset: term ('morphisms' name name)?
     ;
@@ -367,7 +367,7 @@
     'rep\_datatype' ('(' (name +) ')')? (term +)
     ;
 
-    dtspec: parname? typespec infix? '=' (cons + '|')
+    dtspec: parname? typespec mixfix? '=' (cons + '|')
     ;
     cons: name ( type * ) mixfix?
   \end{rail}
--- a/doc-src/IsarRef/Thy/Inner_Syntax.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -244,11 +244,9 @@
 section {* Mixfix annotations *}
 
 text {* Mixfix annotations specify concrete \emph{inner syntax} of
-  Isabelle types and terms.  Some commands such as @{command
-  "typedecl"} admit infixes only, while @{command "definition"} etc.\
-  support the full range of general mixfixes and binders.  Fixed
-  parameters in toplevel theorem statements, locale specifications
-  also admit mixfix annotations.
+  Isabelle types and terms.  Locally fixed parameters in toplevel
+  theorem statements, locale specifications etc.\ also admit mixfix
+  annotations.
 
   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
   \begin{rail}
@@ -359,32 +357,47 @@
 *}
 
 
-section {* Explicit term notation *}
+section {* Explicit notation *}
 
 text {*
   \begin{matharray}{rcll}
+    @{command_def "type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+    @{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"} \\
   \end{matharray}
 
   \begin{rail}
+    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
+    ;
     ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
   \end{rail}
 
   \begin{description}
 
+  \item @{command "type_notation"}~@{text "c (mx)"} associates mixfix
+  syntax with an existing type constructor.  The arity of the
+  constructor is retrieved from the context.
+  
+  \item @{command "no_type_notation"} is similar to @{command
+  "type_notation"}, but removes the specified syntax annotation from
+  the present context.
+
   \item @{command "notation"}~@{text "c (mx)"} associates mixfix
-  syntax with an existing constant or fixed variable.  This is a
-  robust interface to the underlying @{command "syntax"} primitive
-  (\secref{sec:syn-trans}).  Type declaration and internal syntactic
-  representation of the given entity is retrieved from the context.
+  syntax with an existing constant or fixed variable.  The type
+  declaration of the given entity is retrieved from the context.
   
   \item @{command "no_notation"} is similar to @{command "notation"},
   but removes the specified syntax annotation from the present
   context.
 
   \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.
 *}
 
 
--- a/doc-src/IsarRef/Thy/Spec.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/Spec.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -959,9 +959,9 @@
   \end{matharray}
 
   \begin{rail}
-    'types' (typespec '=' type infix? +)
+    'types' (typespec '=' type mixfix? +)
     ;
-    'typedecl' typespec infix?
+    'typedecl' typespec mixfix?
     ;
     'arities' (nameref '::' arity +)
     ;
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Tue Mar 02 04:35:44 2010 -0800
@@ -33,14 +33,14 @@
   \end{matharray}
 
   \begin{rail}
-    'typedecl' typespec infix?
+    'typedecl' typespec mixfix?
     ;
     'typedef' altname? abstype '=' repset
     ;
 
     altname: '(' (name | 'open' | 'open' name) ')'
     ;
-    abstype: typespec infix?
+    abstype: typespec mixfix?
     ;
     repset: term ('morphisms' name name)?
     ;
@@ -372,7 +372,7 @@
     'rep\_datatype' ('(' (name +) ')')? (term +)
     ;
 
-    dtspec: parname? typespec infix? '=' (cons + '|')
+    dtspec: parname? typespec mixfix? '=' (cons + '|')
     ;
     cons: name ( type * ) mixfix?
   \end{rail}
@@ -906,6 +906,9 @@
       \item[iterations] sets how many sets of assignments are
         generated for each particular size.
 
+      \item[no\_assms] specifies whether assumptions in
+        structured proofs should be ignored.
+
     \end{description}
 
     These option can be given within square brackets.
--- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Tue Mar 02 04:35:44 2010 -0800
@@ -266,10 +266,9 @@
 %
 \begin{isamarkuptext}%
 Mixfix annotations specify concrete \emph{inner syntax} of
-  Isabelle types and terms.  Some commands such as \hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}} admit infixes only, while \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}} etc.\
-  support the full range of general mixfixes and binders.  Fixed
-  parameters in toplevel theorem statements, locale specifications
-  also admit mixfix annotations.
+  Isabelle types and terms.  Locally fixed parameters in toplevel
+  theorem statements, locale specifications etc.\ also admit mixfix
+  annotations.
 
   \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
   \begin{rail}
@@ -379,34 +378,47 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
-\isamarkupsection{Explicit term notation%
+\isamarkupsection{Explicit notation%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
 \begin{matharray}{rcll}
+    \indexdef{}{command}{type\_notation}\hypertarget{command.type-notation}{\hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
+    \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}} \\
   \end{matharray}
 
   \begin{rail}
+    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
+    ;
     ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
   \end{rail}
 
   \begin{description}
 
+  \item \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}~\isa{{\isachardoublequote}c\ {\isacharparenleft}mx{\isacharparenright}{\isachardoublequote}} associates mixfix
+  syntax with an existing type constructor.  The arity of the
+  constructor is retrieved from the context.
+  
+  \item \hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}type{\isacharunderscore}notation}}}} is similar to \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}, but removes the specified syntax annotation from
+  the present context.
+
   \item \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}~\isa{{\isachardoublequote}c\ {\isacharparenleft}mx{\isacharparenright}{\isachardoublequote}} associates mixfix
-  syntax with an existing constant or fixed variable.  This is a
-  robust interface to the underlying \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} primitive
-  (\secref{sec:syn-trans}).  Type declaration and internal syntactic
-  representation of the given entity is retrieved from the context.
+  syntax with an existing constant or fixed variable.  The type
+  declaration of the given entity is retrieved from the context.
   
   \item \hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}},
   but removes the specified syntax annotation from the present
   context.
 
-  \end{description}%
+  \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.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarRef/Thy/document/Spec.tex	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Tue Mar 02 04:35:44 2010 -0800
@@ -996,9 +996,9 @@
   \end{matharray}
 
   \begin{rail}
-    'types' (typespec '=' type infix? +)
+    'types' (typespec '=' type mixfix? +)
     ;
-    'typedecl' typespec infix?
+    'typedecl' typespec mixfix?
     ;
     'arities' (nameref '::' arity +)
     ;
--- a/doc-src/Nitpick/nitpick.tex	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/Nitpick/nitpick.tex	Tue Mar 02 04:35:44 2010 -0800
@@ -472,7 +472,7 @@
 \prew
 \textbf{lemma} ``$\forall n.\; \textit{Suc}~n \mathbin{\not=} n \,\Longrightarrow\, P$'' \\
 \textbf{nitpick} [\textit{card~nat}~= 100, \textit{check\_potential}] \\[2\smallskipamount]
-\slshape Warning: The conjecture either trivially holds for the given scopes or (more likely) lies outside Nitpick's supported
+\slshape Warning: The conjecture either trivially holds for the given scopes or lies outside Nitpick's supported
 fragment. Only potential counterexamples may be found. \\[2\smallskipamount]
 Nitpick found a potential counterexample: \\[2\smallskipamount]
 \hbox{}\qquad Free variable: \nopagebreak \\
@@ -2743,8 +2743,8 @@
 \item[$\bullet$] Although this has never been observed, arbitrary theorem
 morphisms could possibly confuse Nitpick, resulting in spurious counterexamples.
 
-%\item[$\bullet$] All constants and types whose names start with
-%\textit{Nitpick}{.} are reserved for internal use.
+\item[$\bullet$] All constants, types, free variables, and schematic variables
+whose names start with \textit{Nitpick}{.} are reserved for internal use.
 \end{enum}
 
 \let\em=\sl
--- a/doc-src/TutorialI/Advanced/Partial.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Advanced/Partial.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -34,7 +34,7 @@
 preconditions:
 *}
 
-constdefs subtract :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+definition subtract :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
 "n \<le> m \<Longrightarrow> subtract m n \<equiv> m - n"
 
 text{*
@@ -85,7 +85,7 @@
 Phrased differently, the relation
 *}
 
-constdefs step1 :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<times> 'a)set"
+definition step1 :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<times> 'a)set" where
   "step1 f \<equiv> {(y,x). y = f x \<and> y \<noteq> x}"
 
 text{*\noindent
@@ -160,7 +160,7 @@
 consider the following definition of function @{const find}:
 *}
 
-constdefs find2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+definition find2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
   "find2 f x \<equiv>
    fst(while (\<lambda>(x,x'). x' \<noteq> x) (\<lambda>(x,x'). (x',f x')) (x,f x))"
 
--- a/doc-src/TutorialI/CTL/CTL.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/CTL/CTL.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -365,8 +365,7 @@
 *}
 
 (*<*)
-constdefs
- eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set"
+definition eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set" where
 "eufix A B T \<equiv> B \<union> A \<inter> (M\<inverse> `` T)"
 
 lemma "lfp(eufix A B) \<subseteq> eusem A B"
@@ -397,8 +396,7 @@
 done
 
 (*
-constdefs
- eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set"
+definition eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set" where
 "eusem A B \<equiv> {s. \<exists>p\<in>Paths s. \<exists>j. p j \<in> B \<and> (\<forall>i < j. p i \<in> A)}"
 
 axioms
@@ -414,8 +412,7 @@
 apply(blast intro: M_total[THEN someI_ex])
 done
 
-constdefs
- pcons :: "state \<Rightarrow> (nat \<Rightarrow> state) \<Rightarrow> (nat \<Rightarrow> state)"
+definition pcons :: "state \<Rightarrow> (nat \<Rightarrow> state) \<Rightarrow> (nat \<Rightarrow> state)" where
 "pcons s p == \<lambda>i. case i of 0 \<Rightarrow> s | Suc j \<Rightarrow> p j"
 
 lemma pcons_PathI: "[| (s,t) : M; p \<in> Paths t |] ==> pcons s p \<in> Paths s";
--- a/doc-src/TutorialI/Misc/Option2.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Misc/Option2.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -24,8 +24,7 @@
 *}
 (*<*)
 (*
-constdefs
- infplus :: "nat option \<Rightarrow> nat option \<Rightarrow> nat option"
+definition infplus :: "nat option \<Rightarrow> nat option \<Rightarrow> nat option" where
 "infplus x y \<equiv> case x of None \<Rightarrow> None
                | Some m \<Rightarrow> (case y of None \<Rightarrow> None | Some n \<Rightarrow> Some(m+n))"
 
--- a/doc-src/TutorialI/Overview/LNCS/FP1.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Overview/LNCS/FP1.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -62,7 +62,7 @@
 consts xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
 defs xor_def: "xor x y \<equiv> x \<and> \<not>y \<or> \<not>x \<and> y"
 
-constdefs nand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+definition nand :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
          "nand x y \<equiv> \<not>(x \<and> y)"
 
 lemma "\<not> xor x x"
--- a/doc-src/TutorialI/Overview/LNCS/Ordinal.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Overview/LNCS/Ordinal.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,10 +9,10 @@
   "pred (Succ a) n = Some a"
   "pred (Limit f) n = Some (f n)"
 
-constdefs
-  OpLim :: "(nat \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"
+definition OpLim :: "(nat \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where
   "OpLim F a \<equiv> Limit (\<lambda>n. F n a)"
-  OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"    ("\<Squnion>")
+
+definition OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"    ("\<Squnion>") where
   "\<Squnion>f \<equiv> OpLim (power f)"
 
 consts
@@ -29,8 +29,7 @@
   "\<nabla>f (Succ a) = f (Succ (\<nabla>f a))"
   "\<nabla>f (Limit h) = Limit (\<lambda>n. \<nabla>f (h n))"
 
-constdefs
-  deriv :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"
+definition deriv :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where
   "deriv f \<equiv> \<nabla>(\<Squnion>f)"
 
 consts
@@ -40,12 +39,13 @@
   "veblen (Succ a) = \<nabla>(OpLim (power (veblen a)))"
   "veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))"
 
-constdefs
-  veb :: "ordinal \<Rightarrow> ordinal"
+definition veb :: "ordinal \<Rightarrow> ordinal" where
   "veb a \<equiv> veblen a Zero"
-  epsilon0 :: ordinal    ("\<epsilon>\<^sub>0")
+
+definition epsilon0 :: ordinal    ("\<epsilon>\<^sub>0") where
   "\<epsilon>\<^sub>0 \<equiv> veb Zero"
-  Gamma0 :: ordinal    ("\<Gamma>\<^sub>0")
+
+definition Gamma0 :: ordinal    ("\<Gamma>\<^sub>0") where
   "\<Gamma>\<^sub>0 \<equiv> Limit (\<lambda>n. (veb^n) Zero)"
 thm Gamma0_def
 
--- a/doc-src/TutorialI/Protocol/Message.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Protocol/Message.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -46,8 +46,7 @@
 text{*The inverse of a symmetric key is itself; that of a public key
       is the private key and vice versa*}
 
-constdefs
-  symKeys :: "key set"
+definition symKeys :: "key set" where
   "symKeys == {K. invKey K = K}"
 (*>*)
 
@@ -92,8 +91,7 @@
   "{|x, y|}"      == "CONST MPair x y"
 
 
-constdefs
-  keysFor :: "msg set => key set"
+definition keysFor :: "msg set => key set" where
     --{*Keys useful to decrypt elements of a message set*}
   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
 
--- a/doc-src/TutorialI/Rules/Primes.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Rules/Primes.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -99,8 +99,7 @@
 
 (**** The material below was omitted from the book ****)
 
-constdefs
-  is_gcd  :: "[nat,nat,nat] \<Rightarrow> bool"        (*gcd as a relation*)
+definition is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" where        (*gcd as a relation*)
     "is_gcd p m n == p dvd m  \<and>  p dvd n  \<and>
                      (ALL d. d dvd m \<and> d dvd n \<longrightarrow> d dvd p)"
 
--- a/doc-src/TutorialI/Sets/Examples.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/TutorialI/Sets/Examples.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -156,8 +156,7 @@
 lemma "{x. P x \<longrightarrow> Q x} = -{x. P x} \<union> {x. Q x}"
 by blast
 
-constdefs
-  prime   :: "nat set"
+definition prime :: "nat set" where
     "prime == {p. 1<p & (ALL m. m dvd p --> m=1 | m=p)}"
 
 lemma "{p*q | p q. p\<in>prime \<and> q\<in>prime} = 
--- a/doc-src/ZF/If.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/ZF/If.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,8 +8,7 @@
 
 theory If imports FOL begin
 
-constdefs
-  "if" :: "[o,o,o]=>o"
+definition "if" :: "[o,o,o]=>o" where
   "if(P,Q,R) == P&Q | ~P&R"
 
 lemma ifI:
--- a/doc-src/ZF/ZF_examples.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/doc-src/ZF/ZF_examples.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -64,7 +64,7 @@
      "t \<in> bt(A) ==> \<forall>k \<in> nat. n_nodes_aux(t)`k = n_nodes(t) #+ k"
   by (induct_tac t, simp_all) 
 
-constdefs  n_nodes_tail :: "i => i"
+definition n_nodes_tail :: "i => i" where
    "n_nodes_tail(t) == n_nodes_aux(t) ` 0"
 
 lemma "t \<in> bt(A) ==> n_nodes_tail(t) = n_nodes(t)"
--- a/etc/isar-keywords-ZF.el	Tue Mar 02 04:31:50 2010 -0800
+++ b/etc/isar-keywords-ZF.el	Tue Mar 02 04:35:44 2010 -0800
@@ -30,7 +30,6 @@
     "arities"
     "assume"
     "attribute_setup"
-    "axclass"
     "axiomatization"
     "axioms"
     "back"
@@ -108,6 +107,7 @@
     "no_notation"
     "no_syntax"
     "no_translations"
+    "no_type_notation"
     "nonterminals"
     "notation"
     "note"
@@ -189,6 +189,7 @@
     "txt"
     "txt_raw"
     "typ"
+    "type_notation"
     "typed_print_translation"
     "typedecl"
     "types"
@@ -348,7 +349,6 @@
     "abbreviation"
     "arities"
     "attribute_setup"
-    "axclass"
     "axiomatization"
     "axioms"
     "class"
@@ -385,6 +385,7 @@
     "no_notation"
     "no_syntax"
     "no_translations"
+    "no_type_notation"
     "nonterminals"
     "notation"
     "oracle"
@@ -404,6 +405,7 @@
     "text_raw"
     "theorems"
     "translations"
+    "type_notation"
     "typed_print_translation"
     "typedecl"
     "types"
--- a/etc/isar-keywords.el	Tue Mar 02 04:31:50 2010 -0800
+++ b/etc/isar-keywords.el	Tue Mar 02 04:35:44 2010 -0800
@@ -37,7 +37,6 @@
     "attribute_setup"
     "automaton"
     "ax_specification"
-    "axclass"
     "axiomatization"
     "axioms"
     "back"
@@ -145,6 +144,7 @@
     "no_notation"
     "no_syntax"
     "no_translations"
+    "no_type_notation"
     "nominal_datatype"
     "nominal_inductive"
     "nominal_inductive2"
@@ -252,6 +252,7 @@
     "txt"
     "txt_raw"
     "typ"
+    "type_notation"
     "typed_print_translation"
     "typedecl"
     "typedef"
@@ -451,7 +452,6 @@
     "atom_decl"
     "attribute_setup"
     "automaton"
-    "axclass"
     "axiomatization"
     "axioms"
     "boogie_end"
@@ -509,6 +509,7 @@
     "no_notation"
     "no_syntax"
     "no_translations"
+    "no_type_notation"
     "nominal_datatype"
     "nonterminals"
     "notation"
@@ -535,6 +536,7 @@
     "text_raw"
     "theorems"
     "translations"
+    "type_notation"
     "typed_print_translation"
     "typedecl"
     "types"
--- a/src/CCL/CCL.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/CCL/CCL.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -193,7 +193,7 @@
       fun mk_inj_lemmas r = [@{thm arg_cong}] RL [r RS (r RS @{thm eq_lemma})]
       val inj_lemmas = maps mk_inj_lemmas rews
     in
-      CHANGED (REPEAT (ares_tac [iffI, allI, conjI] i ORELSE
+      CHANGED (REPEAT (ares_tac [@{thm iffI}, @{thm allI}, @{thm conjI}] i ORELSE
         eresolve_tac inj_lemmas i ORELSE
         asm_simp_tac (simpset_of ctxt addsimps rews) i))
     end;
@@ -242,7 +242,7 @@
   val eq_lemma = thm "eq_lemma";
   val distinctness = thm "distinctness";
   fun mk_lemma (ra,rb) = [lemma] RL [ra RS (rb RS eq_lemma)] RL
-                           [distinctness RS notE,sym RS (distinctness RS notE)]
+                           [distinctness RS notE, @{thm sym} RS (distinctness RS notE)]
 in
   fun mk_lemmas rls = maps mk_lemma (mk_combs pair rls)
   fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs
@@ -258,7 +258,7 @@
   let
     fun mk_raw_dstnct_thm rls s =
       Goal.prove_global @{theory} [] [] (Syntax.read_prop_global @{theory} s)
-        (fn _=> rtac notI 1 THEN eresolve_tac rls 1)
+        (fn _=> rtac @{thm notI} 1 THEN eresolve_tac rls 1)
   in map (mk_raw_dstnct_thm caseB_lemmas)
     (mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end
 
@@ -406,9 +406,9 @@
   "~ false [= lam x. f(x)"
   "~ lam x. f(x) [= false"
   by (tactic {*
-    REPEAT (rtac notI 1 THEN
+    REPEAT (rtac @{thm notI} 1 THEN
       dtac @{thm case_pocong} 1 THEN
-      etac rev_mp 5 THEN
+      etac @{thm rev_mp} 5 THEN
       ALLGOALS (simp_tac @{simpset}) THEN
       REPEAT (resolve_tac [@{thm po_refl}, @{thm npo_lam_bot}] 1)) *})
 
--- a/src/CCL/Type.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/CCL/Type.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -127,7 +127,7 @@
 fun mk_ncanT_tac top_crls crls =
   SUBPROOF (fn {context = ctxt, prems = major :: prems, ...} =>
     resolve_tac ([major] RL top_crls) 1 THEN
-    REPEAT_SOME (eresolve_tac (crls @ [exE, @{thm bexE}, conjE, disjE])) THEN
+    REPEAT_SOME (eresolve_tac (crls @ [@{thm exE}, @{thm bexE}, @{thm conjE}, @{thm disjE}])) THEN
     ALLGOALS (asm_simp_tac (simpset_of ctxt)) THEN
     ALLGOALS (ares_tac (prems RL [@{thm lem}]) ORELSE' etac @{thm bspec}) THEN
     safe_tac (claset_of ctxt addSIs prems))
@@ -498,7 +498,7 @@
 fun EQgen_tac ctxt rews i =
  SELECT_GOAL
    (TRY (safe_tac (claset_of ctxt)) THEN
-    resolve_tac ((rews @ [refl]) RL ((rews @ [refl]) RL [@{thm ssubst_pair}])) i THEN
+    resolve_tac ((rews @ [@{thm refl}]) RL ((rews @ [@{thm refl}]) RL [@{thm ssubst_pair}])) i THEN
     ALLGOALS (simp_tac (simpset_of ctxt)) THEN
     ALLGOALS EQgen_raw_tac) i
 *}
--- a/src/FOL/IFOL.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/FOL/IFOL.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -760,8 +760,7 @@
 
 nonterminals letbinds letbind
 
-constdefs
-  Let :: "['a::{}, 'a => 'b] => ('b::{})"
+definition Let :: "['a::{}, 'a => 'b] => ('b::{})" where
     "Let(s, f) == f(s)"
 
 syntax
@@ -890,45 +889,4 @@
 lemma all_conj_distrib:
   "(ALL x. P(x) & Q(x)) <-> ((ALL x. P(x)) & (ALL x. Q(x)))" by iprover
 
-
-subsection {* Legacy ML bindings *}
-
-ML {*
-val refl = @{thm refl}
-val trans = @{thm trans}
-val sym = @{thm sym}
-val subst = @{thm subst}
-val ssubst = @{thm ssubst}
-val conjI = @{thm conjI}
-val conjE = @{thm conjE}
-val conjunct1 = @{thm conjunct1}
-val conjunct2 = @{thm conjunct2}
-val disjI1 = @{thm disjI1}
-val disjI2 = @{thm disjI2}
-val disjE = @{thm disjE}
-val impI = @{thm impI}
-val impE = @{thm impE}
-val mp = @{thm mp}
-val rev_mp = @{thm rev_mp}
-val TrueI = @{thm TrueI}
-val FalseE = @{thm FalseE}
-val iff_refl = @{thm iff_refl}
-val iff_trans = @{thm iff_trans}
-val iffI = @{thm iffI}
-val iffE = @{thm iffE}
-val iffD1 = @{thm iffD1}
-val iffD2 = @{thm iffD2}
-val notI = @{thm notI}
-val notE = @{thm notE}
-val allI = @{thm allI}
-val allE = @{thm allE}
-val spec = @{thm spec}
-val exI = @{thm exI}
-val exE = @{thm exE}
-val eq_reflection = @{thm eq_reflection}
-val iff_reflection = @{thm iff_reflection}
-val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq}
-val meta_eq_to_iff = @{thm meta_eq_to_iff}
-*}
-
 end
--- a/src/FOL/ex/If.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/FOL/ex/If.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,8 +7,7 @@
 
 theory If imports FOL begin
 
-constdefs
-  "if" :: "[o,o,o]=>o"
+definition "if" :: "[o,o,o]=>o" where
   "if(P,Q,R) == P&Q | ~P&R"
 
 lemma ifI:
--- a/src/FOLP/ex/If.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/FOLP/ex/If.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -4,8 +4,7 @@
 imports FOLP
 begin
 
-constdefs
-  "if" :: "[o,o,o]=>o"
+definition "if" :: "[o,o,o]=>o" where
   "if(P,Q,R) == P&Q | ~P&R"
 
 lemma ifI:
--- a/src/HOL/Algebra/AbelCoset.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/AbelCoset.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -38,15 +38,12 @@
                   ("racong\<index> _")
    "a_r_congruent G \<equiv> r_congruent \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
 
-constdefs
-  A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid"
-     (infixl "A'_Mod" 65)
+definition A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "A'_Mod" 65) where
     --{*Actually defined for groups rather than monoids*}
   "A_FactGroup G H \<equiv> FactGroup \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
 
-constdefs
-  a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow> 
-             ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" 
+definition a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow> 
+             ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" where 
     --{*the kernel of a homomorphism (additive)*}
   "a_kernel G H h \<equiv> kernel \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
                               \<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
--- a/src/HOL/Algebra/Bij.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Bij.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Algebra/Bij.thy
-    ID:         $Id$
     Author:     Florian Kammueller, with new proofs by L C Paulson
 *)
 
@@ -8,12 +7,11 @@
 
 section {* Bijections of a Set, Permutation and Automorphism Groups *}
 
-constdefs
-  Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set"
+definition Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set" where
     --{*Only extensional functions, since otherwise we get too many.*}
   "Bij S \<equiv> extensional S \<inter> {f. bij_betw f S S}"
 
-  BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
+definition BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid" where
   "BijGroup S \<equiv>
     \<lparr>carrier = Bij S,
      mult = \<lambda>g \<in> Bij S. \<lambda>f \<in> Bij S. compose S g f,
@@ -71,11 +69,10 @@
 done
 
 
-constdefs
-  auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set"
+definition auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set" where
   "auto G \<equiv> hom G G \<inter> Bij (carrier G)"
 
-  AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
+definition AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid" where
   "AutoGroup G \<equiv> BijGroup (carrier G) \<lparr>carrier := auto G\<rparr>"
 
 lemma (in group) id_in_auto: "(\<lambda>x \<in> carrier G. x) \<in> auto G"
--- a/src/HOL/Algebra/Congruence.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Congruence.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -35,15 +35,17 @@
   eq_is_closed :: "_ \<Rightarrow> 'a set \<Rightarrow> bool" ("is'_closed\<index> _")
   "is_closed A == (A \<subseteq> carrier S \<and> closure_of A = A)"
 
-syntax
+abbreviation
   not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".\<noteq>\<index>" 50)
-  not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
-  set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
+  where "x .\<noteq>\<^bsub>S\<^esub> y == ~(x .=\<^bsub>S\<^esub> y)"
 
-translations
-  "x .\<noteq>\<index> y" == "~(x .=\<index> y)"
-  "x .\<notin>\<index> A" == "~(x .\<in>\<index> A)"
-  "A {.\<noteq>}\<index> B" == "~(A {.=}\<index> B)"
+abbreviation
+  not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
+  where "x .\<notin>\<^bsub>S\<^esub> A == ~(x .\<in>\<^bsub>S\<^esub> A)"
+
+abbreviation
+  set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
+  where "A {.\<noteq>}\<^bsub>S\<^esub> B == ~(A {.=}\<^bsub>S\<^esub> B)"
 
 locale equivalence =
   fixes S (structure)
--- a/src/HOL/Algebra/Coset.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Coset.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -751,8 +751,7 @@
 
 subsection {*Order of a Group and Lagrange's Theorem*}
 
-constdefs
-  order :: "('a, 'b) monoid_scheme \<Rightarrow> nat"
+definition order :: "('a, 'b) monoid_scheme \<Rightarrow> nat" where
   "order S \<equiv> card (carrier S)"
 
 lemma (in group) rcosets_part_G:
@@ -822,9 +821,7 @@
 
 subsection {*Quotient Groups: Factorization of a Group*}
 
-constdefs
-  FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid"
-     (infixl "Mod" 65)
+definition FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "Mod" 65) where
     --{*Actually defined for groups rather than monoids*}
   "FactGroup G H \<equiv>
     \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
@@ -890,9 +887,8 @@
 text{*The quotient by the kernel of a homomorphism is isomorphic to the 
   range of that homomorphism.*}
 
-constdefs
-  kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow> 
-             ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" 
+definition kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow> 
+             ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" where 
     --{*the kernel of a homomorphism*}
   "kernel G H h \<equiv> {x. x \<in> carrier G & h x = \<one>\<^bsub>H\<^esub>}"
 
--- a/src/HOL/Algebra/Divisibility.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Divisibility.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -3630,8 +3630,7 @@
 
 text {* Number of factors for wellfoundedness *}
 
-constdefs
-  factorcount :: "_ \<Rightarrow> 'a \<Rightarrow> nat"
+definition factorcount :: "_ \<Rightarrow> 'a \<Rightarrow> nat" where
   "factorcount G a == THE c. (ALL as. set as \<subseteq> carrier G \<and> 
                                       wfactors G as a \<longrightarrow> c = length as)"
 
--- a/src/HOL/Algebra/FiniteProduct.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/FiniteProduct.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -26,8 +26,7 @@
 
 inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> foldSetD D f e"
 
-constdefs
-  foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a"
+definition foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a" where
   "foldD D f e A == THE x. (A, x) \<in> foldSetD D f e"
 
 lemma foldSetD_closed:
--- a/src/HOL/Algebra/Group.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Group.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -478,8 +478,7 @@
 
 subsection {* Direct Products *}
 
-constdefs
-  DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid"  (infixr "\<times>\<times>" 80)
+definition DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr "\<times>\<times>" 80) where
   "G \<times>\<times> H \<equiv> \<lparr>carrier = carrier G \<times> carrier H,
                 mult = (\<lambda>(g, h) (g', h'). (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')),
                 one = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>)\<rparr>"
@@ -545,8 +544,7 @@
   "[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I"
 by (fastsimp simp add: hom_def compose_def)
 
-constdefs
-  iso :: "_ => _ => ('a => 'b) set"  (infixr "\<cong>" 60)
+definition iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60) where
   "G \<cong> H == {h. h \<in> hom G H & bij_betw h (carrier G) (carrier H)}"
 
 lemma iso_refl: "(%x. x) \<in> G \<cong> G"
--- a/src/HOL/Algebra/IntRing.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/IntRing.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -22,8 +22,7 @@
 
 subsection {* @{text "\<Z>"}: The Set of Integers as Algebraic Structure *}
 
-constdefs
-  int_ring :: "int ring" ("\<Z>")
+definition int_ring :: "int ring" ("\<Z>") where
   "int_ring \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
 
 lemma int_Zcarr [intro!, simp]:
@@ -324,8 +323,7 @@
 
 subsection {* Ideals and the Modulus *}
 
-constdefs
-   ZMod :: "int => int => int set"
+definition ZMod :: "int => int => int set" where
   "ZMod k r == (Idl\<^bsub>\<Z>\<^esub> {k}) +>\<^bsub>\<Z>\<^esub> r"
 
 lemmas ZMod_defs =
@@ -407,8 +405,7 @@
 
 subsection {* Factorization *}
 
-constdefs
-  ZFact :: "int \<Rightarrow> int set ring"
+definition ZFact :: "int \<Rightarrow> int set ring" where
   "ZFact k == \<Z> Quot (Idl\<^bsub>\<Z>\<^esub> {k})"
 
 lemmas ZFact_defs = ZFact_def FactRing_def
--- a/src/HOL/Algebra/Ring.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/Ring.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -198,8 +198,7 @@
   This definition makes it easy to lift lemmas from @{term finprod}.
 *}
 
-constdefs
-  finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b"
+definition finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" where
   "finsum G f A == finprod (| carrier = carrier G,
      mult = add G, one = zero G |) f A"
 
--- a/src/HOL/Algebra/abstract/Ring2.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -207,7 +207,7 @@
         @{const_name Groups.minus}, @{const_name Groups.one}, @{const_name Groups.times}]
   | ring_ord _ = ~1;
 
-fun termless_ring (a, b) = (TermOrd.term_lpo ring_ord (a, b) = LESS);
+fun termless_ring (a, b) = (Term_Ord.term_lpo ring_ord (a, b) = LESS);
 
 val ring_ss = HOL_basic_ss settermless termless_ring addsimps
   [thm "a_assoc", thm "l_zero", thm "l_neg", thm "a_comm", thm "m_assoc",
--- a/src/HOL/Algebra/ringsimp.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Algebra/ringsimp.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -46,7 +46,7 @@
     val ops = map (fst o Term.strip_comb) ts;
     fun ord (Const (a, _)) = find_index (fn (Const (b, _)) => a=b | _ => false) ops
       | ord (Free (a, _)) = find_index (fn (Free (b, _)) => a=b | _ => false) ops;
-    fun less (a, b) = (TermOrd.term_lpo ord (a, b) = LESS);
+    fun less (a, b) = (Term_Ord.term_lpo ord (a, b) = LESS);
   in asm_full_simp_tac (HOL_ss settermless less addsimps simps) end;
 
 fun algebra_tac ctxt =
--- a/src/HOL/Auth/CertifiedEmail.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/CertifiedEmail.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -25,8 +25,7 @@
   BothAuth :: nat
 
 text{*We formalize a fixed way of computing responses.  Could be better.*}
-constdefs
-  "response"    :: "agent => agent => nat => msg"
+definition "response" :: "agent => agent => nat => msg" where
    "response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}"
 
 
--- a/src/HOL/Auth/Guard/Extensions.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Extensions.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -54,14 +54,12 @@
 
 subsubsection{*extract the agent number of an Agent message*}
 
-consts agt_nb :: "msg => agent"
-
-recdef agt_nb "measure size"
+primrec agt_nb :: "msg => agent" where
 "agt_nb (Agent A) = A"
 
 subsubsection{*messages that are pairs*}
 
-constdefs is_MPair :: "msg => bool"
+definition is_MPair :: "msg => bool" where
 "is_MPair X == EX Y Z. X = {|Y,Z|}"
 
 declare is_MPair_def [simp]
@@ -96,7 +94,7 @@
 
 declare is_MPair_def [simp del]
 
-constdefs has_no_pair :: "msg set => bool"
+definition has_no_pair :: "msg set => bool" where
 "has_no_pair H == ALL X Y. {|X,Y|} ~:H"
 
 declare has_no_pair_def [simp]
@@ -117,7 +115,7 @@
 
 subsubsection{*lemmas on keysFor*}
 
-constdefs usekeys :: "msg set => key set"
+definition usekeys :: "msg set => key set" where
 "usekeys G == {K. EX Y. Crypt K Y:G}"
 
 lemma finite_keysFor [intro]: "finite G ==> finite (keysFor G)"
@@ -224,20 +222,19 @@
 
 subsubsection{*greatest nonce used in a message*}
 
-consts greatest_msg :: "msg => nat"
-
-recdef greatest_msg "measure size"
-"greatest_msg (Nonce n) = n"
-"greatest_msg {|X,Y|} = max (greatest_msg X) (greatest_msg Y)"
-"greatest_msg (Crypt K X) = greatest_msg X"
-"greatest_msg other = 0"
+fun greatest_msg :: "msg => nat"
+where
+  "greatest_msg (Nonce n) = n"
+| "greatest_msg {|X,Y|} = max (greatest_msg X) (greatest_msg Y)"
+| "greatest_msg (Crypt K X) = greatest_msg X"
+| "greatest_msg other = 0"
 
 lemma greatest_msg_is_greatest: "Nonce n:parts {X} ==> n <= greatest_msg X"
 by (induct X, auto)
 
 subsubsection{*sets of keys*}
 
-constdefs keyset :: "msg set => bool"
+definition keyset :: "msg set => bool" where
 "keyset G == ALL X. X:G --> (EX K. X = Key K)"
 
 lemma keyset_in [dest]: "[| keyset G; X:G |] ==> EX K. X = Key K"
@@ -257,7 +254,7 @@
 
 subsubsection{*keys a priori necessary for decrypting the messages of G*}
 
-constdefs keysfor :: "msg set => msg set"
+definition keysfor :: "msg set => msg set" where
 "keysfor G == Key ` keysFor (parts G)"
 
 lemma keyset_keysfor [iff]: "keyset (keysfor G)"
@@ -295,7 +292,7 @@
 
 subsubsection{*general protocol properties*}
 
-constdefs is_Says :: "event => bool"
+definition is_Says :: "event => bool" where
 "is_Says ev == (EX A B X. ev = Says A B X)"
 
 lemma is_Says_Says [iff]: "is_Says (Says A B X)"
@@ -303,7 +300,7 @@
 
 (* one could also require that Gets occurs after Says
 but this is sufficient for our purpose *)
-constdefs Gets_correct :: "event list set => bool"
+definition Gets_correct :: "event list set => bool" where
 "Gets_correct p == ALL evs B X. evs:p --> Gets B X:set evs
 --> (EX A. Says A B X:set evs)"
 
@@ -312,7 +309,7 @@
 apply (simp add: Gets_correct_def)
 by (drule_tac x="Gets B X # evs" in spec, auto)
 
-constdefs one_step :: "event list set => bool"
+definition one_step :: "event list set => bool" where
 "one_step p == ALL evs ev. ev#evs:p --> evs:p"
 
 lemma one_step_Cons [dest]: "[| one_step p; ev#evs:p |] ==> evs:p"
@@ -324,7 +321,7 @@
 lemma trunc: "[| evs @ evs':p; one_step p |] ==> evs':p"
 by (induct evs, auto)
 
-constdefs has_only_Says :: "event list set => bool"
+definition has_only_Says :: "event list set => bool" where
 "has_only_Says p == ALL evs ev. evs:p --> ev:set evs
 --> (EX A B X. ev = Says A B X)"
 
@@ -450,7 +447,7 @@
       if A=A' then insert X (knows_max' A evs) else knows_max' A evs
   ))"
 
-constdefs knows_max :: "agent => event list => msg set"
+definition knows_max :: "agent => event list => msg set" where
 "knows_max A evs == knows_max' A evs Un initState A"
 
 abbreviation
@@ -512,7 +509,7 @@
     | Notes A X => parts {X} Un used' evs
   )"
 
-constdefs init :: "msg set"
+definition init :: "msg set" where
 "init == used []"
 
 lemma used_decomp: "used evs = init Un used' evs"
@@ -629,12 +626,11 @@
 
 subsubsection{*message of an event*}
 
-consts msg :: "event => msg"
-
-recdef msg "measure size"
-"msg (Says A B X) = X"
-"msg (Gets A X) = X"
-"msg (Notes A X) = X"
+primrec msg :: "event => msg"
+where
+  "msg (Says A B X) = X"
+| "msg (Gets A X) = X"
+| "msg (Notes A X) = X"
 
 lemma used_sub_parts_used: "X:used (ev # evs) ==> X:parts {msg ev} Un used evs"
 by (induct ev, auto)
--- a/src/HOL/Auth/Guard/Guard.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Guard.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -76,7 +76,7 @@
 
 subsection{*guarded sets*}
 
-constdefs Guard :: "nat => key set => msg set => bool"
+definition Guard :: "nat => key set => msg set => bool" where
 "Guard n Ks H == ALL X. X:H --> X:guard n Ks"
 
 subsection{*basic facts about @{term Guard}*}
@@ -178,12 +178,11 @@
 
 subsection{*number of Crypt's in a message*}
 
-consts crypt_nb :: "msg => nat"
-
-recdef crypt_nb "measure size"
-"crypt_nb (Crypt K X) = Suc (crypt_nb X)"
-"crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y"
-"crypt_nb X = 0" (* otherwise *)
+fun crypt_nb :: "msg => nat"
+where
+  "crypt_nb (Crypt K X) = Suc (crypt_nb X)"
+| "crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y"
+| "crypt_nb X = 0" (* otherwise *)
 
 subsection{*basic facts about @{term crypt_nb}*}
 
@@ -192,11 +191,10 @@
 
 subsection{*number of Crypt's in a message list*}
 
-consts cnb :: "msg list => nat"
-
-recdef cnb "measure size"
-"cnb [] = 0"
-"cnb (X#l) = crypt_nb X + cnb l"
+primrec cnb :: "msg list => nat"
+where
+  "cnb [] = 0"
+| "cnb (X#l) = crypt_nb X + cnb l"
 
 subsection{*basic facts about @{term cnb}*}
 
@@ -241,7 +239,7 @@
 
 subsection{*list corresponding to "decrypt"*}
 
-constdefs decrypt' :: "msg list => key => msg => msg list"
+definition decrypt' :: "msg list => key => msg => msg list" where
 "decrypt' l K Y == Y # remove l (Crypt K Y)"
 
 declare decrypt'_def [simp]
--- a/src/HOL/Auth/Guard/GuardK.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/GuardK.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -85,7 +85,7 @@
 
 subsection{*guarded sets*}
 
-constdefs GuardK :: "nat => key set => msg set => bool"
+definition GuardK :: "nat => key set => msg set => bool" where
 "GuardK n Ks H == ALL X. X:H --> X:guardK n Ks"
 
 subsection{*basic facts about @{term GuardK}*}
@@ -176,11 +176,9 @@
 
 subsection{*number of Crypt's in a message*}
 
-consts crypt_nb :: "msg => nat"
-
-recdef crypt_nb "measure size"
-"crypt_nb (Crypt K X) = Suc (crypt_nb X)"
-"crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y"
+fun crypt_nb :: "msg => nat" where
+"crypt_nb (Crypt K X) = Suc (crypt_nb X)" |
+"crypt_nb {|X,Y|} = crypt_nb X + crypt_nb Y" |
 "crypt_nb X = 0" (* otherwise *)
 
 subsection{*basic facts about @{term crypt_nb}*}
@@ -190,10 +188,8 @@
 
 subsection{*number of Crypt's in a message list*}
 
-consts cnb :: "msg list => nat"
-
-recdef cnb "measure size"
-"cnb [] = 0"
+primrec cnb :: "msg list => nat" where
+"cnb [] = 0" |
 "cnb (X#l) = crypt_nb X + cnb l"
 
 subsection{*basic facts about @{term cnb}*}
@@ -239,7 +235,7 @@
 
 subsection{*list corresponding to "decrypt"*}
 
-constdefs decrypt' :: "msg list => key => msg => msg list"
+definition decrypt' :: "msg list => key => msg => msg list" where
 "decrypt' l K Y == Y # remove l (Crypt K Y)"
 
 declare decrypt'_def [simp]
--- a/src/HOL/Auth/Guard/Guard_Public.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Guard_Public.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -19,7 +19,7 @@
 
 subsubsection{*signature*}
 
-constdefs sign :: "agent => msg => msg"
+definition sign :: "agent => msg => msg" where
 "sign A X == {|Agent A, X, Crypt (priK A) (Hash X)|}"
 
 lemma sign_inj [iff]: "(sign A X = sign A' X') = (A=A' & X=X')"
@@ -27,7 +27,7 @@
 
 subsubsection{*agent associated to a key*}
 
-constdefs agt :: "key => agent"
+definition agt :: "key => agent" where
 "agt K == @A. K = priK A | K = pubK A"
 
 lemma agt_priK [simp]: "agt (priK A) = A"
@@ -57,7 +57,7 @@
 
 subsubsection{*sets of private keys*}
 
-constdefs priK_set :: "key set => bool"
+definition priK_set :: "key set => bool" where
 "priK_set Ks == ALL K. K:Ks --> (EX A. K = priK A)"
 
 lemma in_priK_set: "[| priK_set Ks; K:Ks |] ==> EX A. K = priK A"
@@ -71,7 +71,7 @@
 
 subsubsection{*sets of good keys*}
 
-constdefs good :: "key set => bool"
+definition good :: "key set => bool" where
 "good Ks == ALL K. K:Ks --> agt K ~:bad"
 
 lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad"
@@ -85,11 +85,10 @@
 
 subsubsection{*greatest nonce used in a trace, 0 if there is no nonce*}
 
-consts greatest :: "event list => nat"
-
-recdef greatest "measure size"
-"greatest [] = 0"
-"greatest (ev # evs) = max (greatest_msg (msg ev)) (greatest evs)"
+primrec greatest :: "event list => nat"
+where
+  "greatest [] = 0"
+| "greatest (ev # evs) = max (greatest_msg (msg ev)) (greatest evs)"
 
 lemma greatest_is_greatest: "Nonce n:used evs ==> n <= greatest evs"
 apply (induct evs, auto simp: initState.simps)
@@ -99,7 +98,7 @@
 
 subsubsection{*function giving a new nonce*}
 
-constdefs new :: "event list => nat"
+definition new :: "event list => nat" where
 "new evs == Suc (greatest evs)"
 
 lemma new_isnt_used [iff]: "Nonce (new evs) ~:used evs"
@@ -151,7 +150,7 @@
 
 subsubsection{*regular protocols*}
 
-constdefs regular :: "event list set => bool"
+definition regular :: "event list set => bool" where
 "regular p == ALL evs A. evs:p --> (Key (priK A):parts (spies evs)) = (A:bad)"
 
 lemma priK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==>
--- a/src/HOL/Auth/Guard/Guard_Shared.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Guard_Shared.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -25,7 +25,7 @@
 
 subsubsection{*agent associated to a key*}
 
-constdefs agt :: "key => agent"
+definition agt :: "key => agent" where
 "agt K == @A. K = shrK A"
 
 lemma agt_shrK [simp]: "agt (shrK A) = A"
@@ -52,7 +52,7 @@
 
 subsubsection{*sets of symmetric keys*}
 
-constdefs shrK_set :: "key set => bool"
+definition shrK_set :: "key set => bool" where
 "shrK_set Ks == ALL K. K:Ks --> (EX A. K = shrK A)"
 
 lemma in_shrK_set: "[| shrK_set Ks; K:Ks |] ==> EX A. K = shrK A"
@@ -66,7 +66,7 @@
 
 subsubsection{*sets of good keys*}
 
-constdefs good :: "key set => bool"
+definition good :: "key set => bool" where
 "good Ks == ALL K. K:Ks --> agt K ~:bad"
 
 lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad"
@@ -154,7 +154,7 @@
 
 subsubsection{*regular protocols*}
 
-constdefs regular :: "event list set => bool"
+definition regular :: "event list set => bool" where
 "regular p == ALL evs A. evs:p --> (Key (shrK A):parts (spies evs)) = (A:bad)"
 
 lemma shrK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==>
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -198,7 +198,7 @@
 
 subsection{*guardedness of NB*}
 
-constdefs ya_keys :: "agent => agent => nat => nat => event list => key set"
+definition ya_keys :: "agent => agent => nat => nat => event list => key set" where
 "ya_keys A B NA NB evs == {shrK A,shrK B} Un {K. ya3 A B NA NB K:set evs}"
 
 lemma Guard_NB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==>
--- a/src/HOL/Auth/Guard/List_Msg.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/List_Msg.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -29,24 +29,18 @@
 
 subsubsection{*head*}
 
-consts head :: "msg => msg"
-
-recdef head "measure size"
+primrec head :: "msg => msg" where
 "head (cons x l) = x"
 
 subsubsection{*tail*}
 
-consts tail :: "msg => msg"
-
-recdef tail "measure size"
+primrec tail :: "msg => msg" where
 "tail (cons x l) = l"
 
 subsubsection{*length*}
 
-consts len :: "msg => nat"
-
-recdef len "measure size"
-"len (cons x l) = Suc (len l)"
+fun len :: "msg => nat" where
+"len (cons x l) = Suc (len l)" |
 "len other = 0"
 
 lemma len_not_empty: "n < len l ==> EX x l'. l = cons x l'"
@@ -54,18 +48,14 @@
 
 subsubsection{*membership*}
 
-consts isin :: "msg * msg => bool"
-
-recdef isin "measure (%(x,l). size l)"
-"isin (x, cons y l) = (x=y | isin (x,l))"
+fun isin :: "msg * msg => bool" where
+"isin (x, cons y l) = (x=y | isin (x,l))" |
 "isin (x, other) = False"
 
 subsubsection{*delete an element*}
 
-consts del :: "msg * msg => msg"
-
-recdef del "measure (%(x,l). size l)"
-"del (x, cons y l) = (if x=y then l else cons y (del (x,l)))"
+fun del :: "msg * msg => msg" where
+"del (x, cons y l) = (if x=y then l else cons y (del (x,l)))" |
 "del (x, other) = other"
 
 lemma notin_del [simp]: "~ isin (x,l) ==> del (x,l) = l"
@@ -76,10 +66,8 @@
 
 subsubsection{*concatenation*}
 
-consts app :: "msg * msg => msg"
-
-recdef app "measure (%(l,l'). size l)"
-"app (cons x l, l') = cons x (app (l,l'))"
+fun app :: "msg * msg => msg" where
+"app (cons x l, l') = cons x (app (l,l'))" |
 "app (other, l') = l'"
 
 lemma isin_app [iff]: "isin (x, app(l,l')) = (isin (x,l) | isin (x,l'))"
@@ -87,20 +75,16 @@
 
 subsubsection{*replacement*}
 
-consts repl :: "msg * nat * msg => msg"
-
-recdef repl "measure (%(l,i,x'). i)"
-"repl (cons x l, Suc i, x') = cons x (repl (l,i,x'))"
-"repl (cons x l, 0, x') = cons x' l"
+fun repl :: "msg * nat * msg => msg" where
+"repl (cons x l, Suc i, x') = cons x (repl (l,i,x'))" |
+"repl (cons x l, 0, x') = cons x' l" |
 "repl (other, i, M') = other"
 
 subsubsection{*ith element*}
 
-consts ith :: "msg * nat => msg"
-
-recdef ith "measure (%(l,i). i)"
-"ith (cons x l, Suc i) = ith (l,i)"
-"ith (cons x l, 0) = x"
+fun ith :: "msg * nat => msg" where
+"ith (cons x l, Suc i) = ith (l,i)" |
+"ith (cons x l, 0) = x" |
 "ith (other, i) = other"
 
 lemma ith_head: "0 < len l ==> ith (l,0) = head l"
@@ -108,10 +92,8 @@
 
 subsubsection{*insertion*}
 
-consts ins :: "msg * nat * msg => msg"
-
-recdef ins "measure (%(l,i,x). i)"
-"ins (cons x l, Suc i, y) = cons x (ins (l,i,y))"
+fun ins :: "msg * nat * msg => msg" where
+"ins (cons x l, Suc i, y) = cons x (ins (l,i,y))" |
 "ins (l, 0, y) = cons y l"
 
 lemma ins_head [simp]: "ins (l,0,y) = cons y l"
@@ -119,10 +101,8 @@
 
 subsubsection{*truncation*}
 
-consts trunc :: "msg * nat => msg"
-
-recdef trunc "measure (%(l,i). i)"
-"trunc (l,0) = l"
+fun trunc :: "msg * nat => msg" where
+"trunc (l,0) = l" |
 "trunc (cons x l, Suc i) = trunc (l,i)"
 
 lemma trunc_zero [simp]: "trunc (l,0) = l"
--- a/src/HOL/Auth/Guard/P1.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/P1.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -39,7 +39,7 @@
 subsubsection{*offer chaining:
 B chains his offer for A with the head offer of L for sending it to C*}
 
-constdefs chain :: "agent => nat => agent => msg => agent => msg"
+definition chain :: "agent => nat => agent => msg => agent => msg" where
 "chain B ofr A L C ==
 let m1= Crypt (pubK A) (Nonce ofr) in
 let m2= Hash {|head L, Agent C|} in
@@ -56,9 +56,7 @@
 
 subsubsection{*agent whose key is used to sign an offer*}
 
-consts shop :: "msg => msg"
-
-recdef shop "measure size"
+fun shop :: "msg => msg" where
 "shop {|B,X,Crypt K H|} = Agent (agt K)"
 
 lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
@@ -66,9 +64,7 @@
 
 subsubsection{*nonce used in an offer*}
 
-consts nonce :: "msg => msg"
-
-recdef nonce "measure size"
+fun nonce :: "msg => msg" where
 "nonce {|B,{|Crypt K ofr,m2|},CryptH|} = ofr"
 
 lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
@@ -76,9 +72,7 @@
 
 subsubsection{*next shop*}
 
-consts next_shop :: "msg => agent"
-
-recdef next_shop "measure size"
+fun next_shop :: "msg => agent" where
 "next_shop {|B,{|m1,Hash{|headL,Agent C|}|},CryptH|} = C"
 
 lemma next_shop_chain [iff]: "next_shop (chain B ofr A L C) = C"
@@ -86,7 +80,7 @@
 
 subsubsection{*anchor of the offer list*}
 
-constdefs anchor :: "agent => nat => agent => msg"
+definition anchor :: "agent => nat => agent => msg" where
 "anchor A n B == chain A n A (cons nil nil) B"
 
 lemma anchor_inj [iff]: "(anchor A n B = anchor A' n' B')
@@ -107,7 +101,7 @@
 
 subsubsection{*request event*}
 
-constdefs reqm :: "agent => nat => nat => msg => agent => msg"
+definition reqm :: "agent => nat => nat => msg => agent => msg" where
 "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
 cons (anchor A n B) nil|}"
 
@@ -118,7 +112,7 @@
 lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}"
 by (auto simp: reqm_def)
 
-constdefs req :: "agent => nat => nat => msg => agent => event"
+definition req :: "agent => nat => nat => msg => agent => event" where
 "req A r n I B == Says A B (reqm A r n I B)"
 
 lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B')
@@ -127,8 +121,8 @@
 
 subsubsection{*propose event*}
 
-constdefs prom :: "agent => nat => agent => nat => msg => msg =>
-msg => agent => msg"
+definition prom :: "agent => nat => agent => nat => msg => msg =>
+msg => agent => msg" where
 "prom B ofr A r I L J C == {|Agent A, Number r,
 app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}"
 
@@ -140,8 +134,8 @@
 lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}"
 by (auto simp: prom_def)
 
-constdefs pro :: "agent => nat => agent => nat => msg => msg =>
-msg => agent => event"
+definition pro :: "agent => nat => agent => nat => msg => msg =>
+msg => agent => event" where
 "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)"
 
 lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C'
@@ -198,7 +192,7 @@
 
 subsubsection{*offers of an offer list*}
 
-constdefs offer_nonces :: "msg => msg set"
+definition offer_nonces :: "msg => msg set" where
 "offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}"
 
 subsubsection{*the originator can get the offers*}
@@ -209,18 +203,14 @@
 
 subsubsection{*list of offers*}
 
-consts offers :: "msg => msg"
-
-recdef offers "measure size"
-"offers (cons M L) = cons {|shop M, nonce M|} (offers L)"
+fun offers :: "msg => msg" where
+"offers (cons M L) = cons {|shop M, nonce M|} (offers L)" |
 "offers other = nil"
 
 subsubsection{*list of agents whose keys are used to sign a list of offers*}
 
-consts shops :: "msg => msg"
-
-recdef shops "measure size"
-"shops (cons M L) = cons (shop M) (shops L)"
+fun shops :: "msg => msg" where
+"shops (cons M L) = cons (shop M) (shops L)" |
 "shops other = other"
 
 lemma shops_in_agl: "L:valid A n B ==> shops L:agl"
@@ -228,10 +218,8 @@
 
 subsubsection{*builds a trace from an itinerary*}
 
-consts offer_list :: "agent * nat * agent * msg * nat => msg"
-
-recdef offer_list "measure (%(A,n,B,I,ofr). size I)"
-"offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil"
+fun offer_list :: "agent * nat * agent * msg * nat => msg" where
+"offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil" |
 "offer_list (A,n,B,cons (Agent C) I,ofr) = (
 let L = offer_list (A,n,B,I,Suc ofr) in
 cons (chain (next_shop (head L)) ofr A L C) L)"
@@ -239,11 +227,9 @@
 lemma "I:agl ==> ALL ofr. offer_list (A,n,B,I,ofr):valid A n B"
 by (erule agl.induct, auto)
 
-consts trace :: "agent * nat * agent * nat * msg * msg * msg
-=> event list"
-
-recdef trace "measure (%(B,ofr,A,r,I,L,K). size K)"
-"trace (B,ofr,A,r,I,L,nil) = []"
+fun trace :: "agent * nat * agent * nat * msg * msg * msg
+=> event list" where
+"trace (B,ofr,A,r,I,L,nil) = []" |
 "trace (B,ofr,A,r,I,L,cons (Agent D) K) = (
 let C = (if K=nil then B else agt_nb (head K)) in
 let I' = (if K=nil then cons (Agent A) (cons (Agent B) I)
@@ -252,7 +238,7 @@
 pro C (Suc ofr) A r I' L nil D
 # trace (B,Suc ofr,A,r,I'',tail L,K))"
 
-constdefs trace' :: "agent => nat => nat => msg => agent => nat => event list"
+definition trace' :: "agent => nat => nat => msg => agent => nat => event list" where
 "trace' A r n I B ofr == (
 let AI = cons (Agent A) I in
 let L = offer_list (A,n,B,AI,ofr) in
--- a/src/HOL/Auth/Guard/P2.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/P2.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -26,7 +26,7 @@
 subsubsection{*offer chaining:
 B chains his offer for A with the head offer of L for sending it to C*}
 
-constdefs chain :: "agent => nat => agent => msg => agent => msg"
+definition chain :: "agent => nat => agent => msg => agent => msg" where
 "chain B ofr A L C ==
 let m1= sign B (Nonce ofr) in
 let m2= Hash {|head L, Agent C|} in
@@ -43,9 +43,7 @@
 
 subsubsection{*agent whose key is used to sign an offer*}
 
-consts shop :: "msg => msg"
-
-recdef shop "measure size"
+fun shop :: "msg => msg" where
 "shop {|Crypt K {|B,ofr,Crypt K' H|},m2|} = Agent (agt K')"
 
 lemma shop_chain [simp]: "shop (chain B ofr A L C) = Agent B"
@@ -53,9 +51,7 @@
 
 subsubsection{*nonce used in an offer*}
 
-consts nonce :: "msg => msg"
-
-recdef nonce "measure size"
+fun nonce :: "msg => msg" where
 "nonce {|Crypt K {|B,ofr,CryptH|},m2|} = ofr"
 
 lemma nonce_chain [simp]: "nonce (chain B ofr A L C) = Nonce ofr"
@@ -63,9 +59,7 @@
 
 subsubsection{*next shop*}
 
-consts next_shop :: "msg => agent"
-
-recdef next_shop "measure size"
+fun next_shop :: "msg => agent" where
 "next_shop {|m1,Hash {|headL,Agent C|}|} = C"
 
 lemma "next_shop (chain B ofr A L C) = C"
@@ -73,7 +67,7 @@
 
 subsubsection{*anchor of the offer list*}
 
-constdefs anchor :: "agent => nat => agent => msg"
+definition anchor :: "agent => nat => agent => msg" where
 "anchor A n B == chain A n A (cons nil nil) B"
 
 lemma anchor_inj [iff]:
@@ -88,7 +82,7 @@
 
 subsubsection{*request event*}
 
-constdefs reqm :: "agent => nat => nat => msg => agent => msg"
+definition reqm :: "agent => nat => nat => msg => agent => msg" where
 "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I),
 cons (anchor A n B) nil|}"
 
@@ -99,7 +93,7 @@
 lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}"
 by (auto simp: reqm_def)
 
-constdefs req :: "agent => nat => nat => msg => agent => event"
+definition req :: "agent => nat => nat => msg => agent => event" where
 "req A r n I B == Says A B (reqm A r n I B)"
 
 lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B')
@@ -108,8 +102,8 @@
 
 subsubsection{*propose event*}
 
-constdefs prom :: "agent => nat => agent => nat => msg => msg =>
-msg => agent => msg"
+definition prom :: "agent => nat => agent => nat => msg => msg =>
+msg => agent => msg" where
 "prom B ofr A r I L J C == {|Agent A, Number r,
 app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}"
 
@@ -120,8 +114,8 @@
 lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}"
 by (auto simp: prom_def)
 
-constdefs pro :: "agent => nat => agent => nat => msg => msg =>
-                  msg => agent => event"
+definition pro :: "agent => nat => agent => nat => msg => msg =>
+                  msg => agent => event" where
 "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)"
 
 lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C'
@@ -164,11 +158,10 @@
 
 subsubsection{*list of offers*}
 
-consts offers :: "msg => msg"
-
-recdef offers "measure size"
-"offers (cons M L) = cons {|shop M, nonce M|} (offers L)"
-"offers other = nil"
+fun offers :: "msg => msg"
+where
+  "offers (cons M L) = cons {|shop M, nonce M|} (offers L)"
+| "offers other = nil"
 
 
 subsection{*Properties of Protocol P2*}
--- a/src/HOL/Auth/Guard/Proto.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Guard/Proto.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -23,7 +23,7 @@
 
 types proto = "rule set"
 
-constdefs wdef :: "proto => bool"
+definition wdef :: "proto => bool" where
 "wdef p == ALL R k. R:p --> Number k:parts {msg' R}
 --> Number k:parts (msg`(fst R))"
 
@@ -35,19 +35,17 @@
   nb    :: "nat => msg"
   key   :: "key => key"
 
-consts apm :: "subs => msg => msg"
-
-primrec
-"apm s (Agent A) = Agent (agent s A)"
-"apm s (Nonce n) = Nonce (nonce s n)"
-"apm s (Number n) = nb s n"
-"apm s (Key K) = Key (key s K)"
-"apm s (Hash X) = Hash (apm s X)"
-"apm s (Crypt K X) = (
+primrec apm :: "subs => msg => msg" where
+  "apm s (Agent A) = Agent (agent s A)"
+| "apm s (Nonce n) = Nonce (nonce s n)"
+| "apm s (Number n) = nb s n"
+| "apm s (Key K) = Key (key s K)"
+| "apm s (Hash X) = Hash (apm s X)"
+| "apm s (Crypt K X) = (
 if (EX A. K = pubK A) then Crypt (pubK (agent s (agt K))) (apm s X)
 else if (EX A. K = priK A) then Crypt (priK (agent s (agt K))) (apm s X)
 else Crypt (key s K) (apm s X))"
-"apm s {|X,Y|} = {|apm s X, apm s Y|}"
+| "apm s {|X,Y|} = {|apm s X, apm s Y|}"
 
 lemma apm_parts: "X:parts {Y} ==> apm s X:parts {apm s Y}"
 apply (erule parts.induct, simp_all, blast)
@@ -69,12 +67,10 @@
 apply (drule_tac Y="msg x" and s=s in apm_parts, simp)
 by (blast dest: parts_parts)
 
-consts ap :: "subs => event => event"
-
-primrec
-"ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)"
-"ap s (Gets A X) = Gets (agent s A) (apm s X)"
-"ap s (Notes A X) = Notes (agent s A) (apm s X)"
+primrec ap :: "subs => event => event" where
+  "ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)"
+| "ap s (Gets A X) = Gets (agent s A) (apm s X)"
+| "ap s (Notes A X) = Notes (agent s A) (apm s X)"
 
 abbreviation
   ap' :: "subs => rule => event" where
@@ -94,7 +90,7 @@
 
 subsection{*nonces generated by a rule*}
 
-constdefs newn :: "rule => nat set"
+definition newn :: "rule => nat set" where
 "newn R == {n. Nonce n:parts {msg (snd R)} & Nonce n ~:parts (msg`(fst R))}"
 
 lemma newn_parts: "n:newn R ==> Nonce (nonce s n):parts {apm' s R}"
@@ -102,7 +98,7 @@
 
 subsection{*traces generated by a protocol*}
 
-constdefs ok :: "event list => rule => subs => bool"
+definition ok :: "event list => rule => subs => bool" where
 "ok evs R s == ((ALL x. x:fst R --> ap s x:set evs)
 & (ALL n. n:newn R --> Nonce (nonce s n) ~:used evs))"
 
@@ -124,7 +120,7 @@
 apply (unfold one_step_def, clarify)
 by (ind_cases "ev # evs:tr p" for ev evs, auto)
 
-constdefs has_only_Says' :: "proto => bool"
+definition has_only_Says' :: "proto => bool" where
 "has_only_Says' p == ALL R. R:p --> is_Says (snd R)"
 
 lemma has_only_Says'D: "[| R:p; has_only_Says' p |]
@@ -165,8 +161,8 @@
 
 subsection{*introduction of a fresh guarded nonce*}
 
-constdefs fresh :: "proto => rule => subs => nat => key set => event list
-=> bool"
+definition fresh :: "proto => rule => subs => nat => key set => event list
+=> bool" where
 "fresh p R s n Ks evs == (EX evs1 evs2. evs = evs2 @ ap' s R # evs1
 & Nonce n ~:used evs1 & R:p & ok evs1 R s & Nonce n:parts {apm' s R}
 & apm' s R:guard n Ks)"
@@ -226,7 +222,7 @@
 
 subsection{*safe keys*}
 
-constdefs safe :: "key set => msg set => bool"
+definition safe :: "key set => msg set => bool" where
 "safe Ks G == ALL K. K:Ks --> Key K ~:analz G"
 
 lemma safeD [dest]: "[| safe Ks G; K:Ks |] ==> Key K ~:analz G"
@@ -240,7 +236,7 @@
 
 subsection{*guardedness preservation*}
 
-constdefs preserv :: "proto => keyfun => nat => key set => bool"
+definition preserv :: "proto => keyfun => nat => key set => bool" where
 "preserv p keys n Ks == (ALL evs R' s' R s. evs:tr p -->
 Guard n Ks (spies evs) --> safe Ks (spies evs) --> fresh p R' s' n Ks evs -->
 keys R' s' n evs <= Ks --> R:p --> ok evs R s --> apm' s R:guard n Ks)"
@@ -257,7 +253,7 @@
 
 subsection{*monotonic keyfun*}
 
-constdefs monoton :: "proto => keyfun => bool"
+definition monoton :: "proto => keyfun => bool" where
 "monoton p keys == ALL R' s' n ev evs. ev # evs:tr p -->
 keys R' s' n evs <= keys R' s' n (ev # evs)"
 
@@ -323,7 +319,7 @@
 
 subsection{*unicity*}
 
-constdefs uniq :: "proto => secfun => bool"
+definition uniq :: "proto => secfun => bool" where
 "uniq p secret == ALL evs R R' n n' Ks s s'. R:p --> R':p -->
 n:newn R --> n':newn R' --> nonce s n = nonce s' n' -->
 Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} -->
@@ -340,13 +336,13 @@
 secret R n s Ks = secret R' n' s' Ks"
 by (unfold uniq_def, blast)
 
-constdefs ord :: "proto => (rule => rule => bool) => bool"
+definition ord :: "proto => (rule => rule => bool) => bool" where
 "ord p inff == ALL R R'. R:p --> R':p --> ~ inff R R' --> inff R' R"
 
 lemma ordD: "[| ord p inff; ~ inff R R'; R:p; R':p |] ==> inff R' R"
 by (unfold ord_def, blast)
 
-constdefs uniq' :: "proto => (rule => rule => bool) => secfun => bool"
+definition uniq' :: "proto => (rule => rule => bool) => secfun => bool" where
 "uniq' p inff secret == ALL evs R R' n n' Ks s s'. R:p --> R':p -->
 inff R R' --> n:newn R --> n':newn R' --> nonce s n = nonce s' n' -->
 Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} -->
@@ -372,13 +368,12 @@
 
 subsection{*Needham-Schroeder-Lowe*}
 
-constdefs
-a :: agent "a == Friend 0"
-b :: agent "b == Friend 1"
-a' :: agent "a' == Friend 2"
-b' :: agent "b' == Friend 3"
-Na :: nat "Na == 0"
-Nb :: nat "Nb == 1"
+definition a :: agent where "a == Friend 0"
+definition b :: agent where "b == Friend 1"
+definition a' :: agent where "a' == Friend 2"
+definition b' :: agent where "b' == Friend 3"
+definition Na :: nat where "Na == 0"
+definition Nb :: nat where "Nb == 1"
 
 abbreviation
   ns1 :: rule where
@@ -408,19 +403,19 @@
   ns3b :: event where
   "ns3b == Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})"
 
-constdefs keys :: "keyfun"
+definition keys :: "keyfun" where
 "keys R' s' n evs == {priK' s' a, priK' s' b}"
 
 lemma "monoton ns keys"
 by (simp add: keys_def monoton_def)
 
-constdefs secret :: "secfun"
+definition secret :: "secfun" where
 "secret R n s Ks ==
 (if R=ns1 then apm s (Crypt (pubK b) {|Nonce Na, Agent a|})
 else if R=ns2 then apm s (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})
 else Number 0)"
 
-constdefs inf :: "rule => rule => bool"
+definition inf :: "rule => rule => bool" where
 "inf R R' == (R=ns1 | (R=ns2 & R'~=ns1) | (R=ns3 & R'=ns3))"
 
 lemma inf_is_ord [iff]: "ord ns inf"
--- a/src/HOL/Auth/KerberosIV.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/KerberosIV.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -101,8 +101,7 @@
 
 
 (* Predicate formalising the association between authKeys and servKeys *)
-constdefs
-  AKcryptSK :: "[key, key, event list] => bool"
+definition AKcryptSK :: "[key, key, event list] => bool" where
   "AKcryptSK authK servK evs ==
      \<exists>A B Ts.
        Says Tgs A (Crypt authK
--- a/src/HOL/Auth/KerberosIV_Gets.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/KerberosIV_Gets.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -89,8 +89,7 @@
 
 
 (* Predicate formalising the association between authKeys and servKeys *)
-constdefs
-  AKcryptSK :: "[key, key, event list] => bool"
+definition AKcryptSK :: "[key, key, event list] => bool" where
   "AKcryptSK authK servK evs ==
      \<exists>A B Ts.
        Says Tgs A (Crypt authK
--- a/src/HOL/Auth/KerberosV.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/KerberosV.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -92,8 +92,7 @@
 
 
 (* Predicate formalising the association between authKeys and servKeys *)
-constdefs
-  AKcryptSK :: "[key, key, event list] => bool"
+definition AKcryptSK :: "[key, key, event list] => bool" where
   "AKcryptSK authK servK evs ==
      \<exists>A B tt.
        Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
--- a/src/HOL/Auth/Message.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Message.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -32,8 +32,7 @@
 text{*The inverse of a symmetric key is itself; that of a public key
       is the private key and vice versa*}
 
-constdefs
-  symKeys :: "key set"
+definition symKeys :: "key set" where
   "symKeys == {K. invKey K = K}"
 
 datatype  --{*We allow any number of friendly agents*}
@@ -61,12 +60,11 @@
   "{|x, y|}"      == "CONST MPair x y"
 
 
-constdefs
-  HPair :: "[msg,msg] => msg"                       ("(4Hash[_] /_)" [0, 1000])
+definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
     --{*Message Y paired with a MAC computed with the help of X*}
     "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
 
-  keysFor :: "msg set => key set"
+definition keysFor :: "msg set => key set" where
     --{*Keys useful to decrypt elements of a message set*}
   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
 
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
-(*  ID:         $Id$
-    Author:     Giampaolo Bella, Catania University
+(*  Author:     Giampaolo Bella, Catania University
 *)
 
 header{*Original Shoup-Rubin protocol*}
@@ -29,9 +28,7 @@
     between each agent and his smartcard*)
    shouprubin_assumes_securemeans [iff]: "evs \<in> sr \<Longrightarrow> secureM"
 
-constdefs
-
-  Unique :: "[event, event list] => bool" ("Unique _ on _")
+definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
    "Unique ev on evs == 
       ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
 
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
-(*  ID:         $Id$
-    Author:     Giampaolo Bella, Catania University
+(*  Author:     Giampaolo Bella, Catania University
 *)
 
 header{*Bella's modification of the Shoup-Rubin protocol*}
@@ -35,9 +34,7 @@
     between each agent and his smartcard*)
    shouprubin_assumes_securemeans [iff]: "evs \<in> srb \<Longrightarrow> secureM"
 
-constdefs
-
-  Unique :: "[event, event list] => bool" ("Unique _ on _")
+definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
    "Unique ev on evs == 
       ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
 
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -43,15 +43,11 @@
   shrK_disj_pin [iff]:  "shrK P \<noteq> pin Q"
   crdK_disj_pin [iff]:   "crdK C \<noteq> pin P"
 
-constdefs
-  legalUse :: "card => bool" ("legalUse (_)")
+definition legalUse :: "card => bool" ("legalUse (_)") where
   "legalUse C == C \<notin> stolen"
 
-consts  
-  illegalUse :: "card  => bool"
-primrec
-  illegalUse_def: 
-  "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad)  \<or>  Card A \<in> cloned )"
+primrec illegalUse :: "card  => bool" where
+  illegalUse_def: "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad)  \<or>  Card A \<in> cloned )"
 
 
 text{*initState must be defined with care*}
--- a/src/HOL/Auth/TLS.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/TLS.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -43,8 +43,7 @@
 
 theory TLS imports Public Nat_Int_Bij begin
 
-constdefs
-  certificate      :: "[agent,key] => msg"
+definition certificate :: "[agent,key] => msg" where
     "certificate A KA == Crypt (priSK Server) {|Agent A, Key KA|}"
 
 text{*TLS apparently does not require separate keypairs for encryption and
--- a/src/HOL/Auth/Yahalom.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/Yahalom.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -74,8 +74,7 @@
           ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> yahalom"
 
 
-constdefs 
-  KeyWithNonce :: "[key, nat, event list] => bool"
+definition KeyWithNonce :: "[key, nat, event list] => bool" where
   "KeyWithNonce K NB evs ==
      \<exists>A B na X. 
        Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} 
--- a/src/HOL/Auth/ZhouGollmann.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Auth/ZhouGollmann.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -21,8 +21,7 @@
 abbreviation f_con :: nat where "f_con == 4"
 
 
-constdefs
-  broken :: "agent set"    
+definition broken :: "agent set" where    
     --{*the compromised honest agents; TTP is included as it's not allowed to
         use the protocol*}
    "broken == bad - {Spy}"
--- a/src/HOL/Bali/AxCompl.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/AxCompl.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -20,9 +20,7 @@
            
 section "set of not yet initialzed classes"
 
-constdefs
-
-  nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set"
+definition nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set" where
  "nyinitcls G s \<equiv> {C. is_class G C \<and> \<not> initd C s}"
 
 lemma nyinitcls_subset_class: "nyinitcls G s \<subseteq> {C. is_class G C}"
@@ -115,8 +113,7 @@
 
 section "init-le"
 
-constdefs
-  init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool"            ("_\<turnstile>init\<le>_"  [51,51] 50)
+definition init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool" ("_\<turnstile>init\<le>_"  [51,51] 50) where
  "G\<turnstile>init\<le>n \<equiv> \<lambda>s. card (nyinitcls G s) \<le> n"
   
 lemma init_le_def2 [simp]: "(G\<turnstile>init\<le>n) s = (card (nyinitcls G s)\<le>n)"
@@ -135,9 +132,7 @@
 
 section "Most General Triples and Formulas"
 
-constdefs
-
-  remember_init_state :: "state assn"                ("\<doteq>")
+definition remember_init_state :: "state assn" ("\<doteq>") where
   "\<doteq> \<equiv> \<lambda>Y s Z. s = Z"
 
 lemma remember_init_state_def2 [simp]: "\<doteq> Y = op ="
@@ -579,8 +574,7 @@
 unroll the loop in iterated evaluations of the expression and evaluations of
 the loop body. *}
 
-constdefs
- unroll:: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times>  state) set"
+definition unroll :: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times>  state) set" where
 
  "unroll G l e c \<equiv> {(s,t). \<exists> v s1 s2.
                              G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1 \<and> the_Bool v \<and> normal s1 \<and>
--- a/src/HOL/Bali/AxExample.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/AxExample.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,8 +8,7 @@
 imports AxSem Example
 begin
 
-constdefs
-  arr_inv :: "st \<Rightarrow> bool"
+definition arr_inv :: "st \<Rightarrow> bool" where
  "arr_inv \<equiv> \<lambda>s. \<exists>obj a T el. globs s (Stat Base) = Some obj \<and>
                               values obj (Inl (arr, Base)) = Some (Addr a) \<and>
                               heap s a = Some \<lparr>tag=Arr T 2,values=el\<rparr>"
--- a/src/HOL/Bali/AxSem.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/AxSem.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -63,8 +63,7 @@
       "res"    <= (type) "AxSem.res"
       "a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> a \<Rightarrow> bool"
 
-constdefs
-  assn_imp   :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool"             (infixr "\<Rightarrow>" 25)
+definition assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr "\<Rightarrow>" 25) where
  "P \<Rightarrow> Q \<equiv> \<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z"
   
 lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
@@ -77,8 +76,7 @@
 
 subsection "peek-and"
 
-constdefs
-  peek_and   :: "'a assn \<Rightarrow> (state \<Rightarrow>  bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13)
+definition peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow>  bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13) where
  "P \<and>. p \<equiv> \<lambda>Y s Z. P Y s Z \<and> p s"
 
 lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))"
@@ -117,8 +115,7 @@
 
 subsection "assn-supd"
 
-constdefs
-  assn_supd  :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13)
+definition assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13) where
  "P ;. f \<equiv> \<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s"
 
 lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)"
@@ -128,8 +125,7 @@
 
 subsection "supd-assn"
 
-constdefs
-  supd_assn  :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13)
+definition supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13) where
  "f .; P \<equiv> \<lambda>Y s. P Y (f s)"
 
 
@@ -148,8 +144,7 @@
 
 subsection "subst-res"
 
-constdefs
-  subst_res   :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"              ("_\<leftarrow>_"  [60,61] 60)
+definition subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<leftarrow>_"  [60,61] 60) where
  "P\<leftarrow>w \<equiv> \<lambda>Y. P w"
 
 lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w"
@@ -184,8 +179,7 @@
 
 subsection "subst-Bool"
 
-constdefs
-  subst_Bool  :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn"             ("_\<leftarrow>=_" [60,61] 60)
+definition subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" ("_\<leftarrow>=_" [60,61] 60) where
  "P\<leftarrow>=b \<equiv> \<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b)"
 
 lemma subst_Bool_def2 [simp]: 
@@ -200,8 +194,7 @@
 
 subsection "peek-res"
 
-constdefs
-  peek_res    :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+definition peek_res :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" where
  "peek_res Pf \<equiv> \<lambda>Y. Pf Y Y"
 
 syntax
@@ -229,8 +222,7 @@
 
 subsection "ign-res"
 
-constdefs
-  ign_res    ::  "        'a assn \<Rightarrow> 'a assn"            ("_\<down>" [1000] 1000)
+definition ign_res :: "        'a assn \<Rightarrow> 'a assn" ("_\<down>" [1000] 1000) where
   "P\<down>        \<equiv> \<lambda>Y s Z. \<exists>Y. P Y s Z"
 
 lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)"
@@ -261,8 +253,7 @@
 
 subsection "peek-st"
 
-constdefs
-  peek_st    :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+definition peek_st :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" where
  "peek_st P \<equiv> \<lambda>Y s. P (store s) Y s"
 
 syntax
@@ -306,8 +297,7 @@
 
 subsection "ign-res-eq"
 
-constdefs
-  ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn"               ("_\<down>=_"  [60,61] 60)
+definition ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<down>=_"  [60,61] 60) where
  "P\<down>=w       \<equiv> \<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w)"
 
 lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)"
@@ -337,8 +327,7 @@
 
 subsection "RefVar"
 
-constdefs
-  RefVar    :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn"(infixr "..;" 13)
+definition RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr "..;" 13) where
  "vf ..; P \<equiv> \<lambda>Y s. let (v,s') = vf s in P (Var v) s'"
  
 lemma RefVar_def2 [simp]: "(vf ..; P) Y s =  
@@ -349,12 +338,11 @@
 
 subsection "allocation"
 
-constdefs
-  Alloc      :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition Alloc :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
  "Alloc G otag P \<equiv> \<lambda>Y s Z.
                    \<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z"
 
-  SXAlloc     :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition SXAlloc     :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
  "SXAlloc G P \<equiv> \<lambda>Y s Z. \<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z"
 
 
@@ -372,8 +360,7 @@
 
 section "validity"
 
-constdefs
-  type_ok  :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool"
+definition type_ok :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool" where
  "type_ok G t s \<equiv> 
     \<exists>L T C A. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> 
                             \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A )
@@ -419,10 +406,8 @@
 apply auto
 done
 
-constdefs
-  mtriples  :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow> 
-                ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times>  'sig) set \<Rightarrow> 'a triples"
-                                     ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75)
+definition mtriples :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow> 
+                ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times>  'sig) set \<Rightarrow> 'a triples" ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75) where
  "{{P} tf-\<succ> {Q} | ms} \<equiv> (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms"
   
 consts
@@ -641,8 +626,7 @@
 axioms 
 *)
 
-constdefs
- adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
 "adapt_pre P Q Q'\<equiv>\<lambda>Y s Z. \<forall>Y' s'. \<exists>Z'. P Y s Z' \<and> (Q Y' s' Z' \<longrightarrow> Q' Y' s' Z)"
 
 
--- a/src/HOL/Bali/Basis.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Basis.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -222,12 +222,12 @@
 section "quantifiers for option type"
 
 syntax
-  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
-  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
+  "_Oall" :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
+  "_Oex"  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
 
 syntax (symbols)
-  Oall :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
-  Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
+  "_Oall" :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
+  "_Oex"  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
 
 translations
   "! x:A: P"    == "! x:CONST Option.set A. P"
@@ -237,8 +237,7 @@
 
 text{* Deemed too special for theory Map. *}
 
-constdefs
-  chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)"
+definition chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" where
  "chg_map f a m == case m a of None => m | Some b => m(a|->f b)"
 
 lemma chg_map_new[simp]: "m a = None   ==> chg_map f a m = m"
@@ -253,8 +252,7 @@
 
 section "unique association lists"
 
-constdefs
-  unique   :: "('a \<times> 'b) list \<Rightarrow> bool"
+definition unique :: "('a \<times> 'b) list \<Rightarrow> bool" where
  "unique \<equiv> distinct \<circ> map fst"
 
 lemma uniqueD [rule_format (no_asm)]: 
--- a/src/HOL/Bali/Conform.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Conform.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -22,9 +22,7 @@
 section "extension of global store"
 
 
-constdefs
-
-  gext    :: "st \<Rightarrow> st \<Rightarrow> bool"                ("_\<le>|_"       [71,71]   70)
+definition gext :: "st \<Rightarrow> st \<Rightarrow> bool" ("_\<le>|_"       [71,71]   70) where
    "s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>globs s' r: tag obj'= tag obj"
 
 text {* For the the proof of type soundness we will need the 
@@ -98,9 +96,7 @@
 
 section "value conformance"
 
-constdefs
-
-  conf  :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool"    ("_,_\<turnstile>_\<Colon>\<preceq>_"   [71,71,71,71] 70)
+definition conf :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>_"   [71,71,71,71] 70) where
            "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
 
 lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -181,10 +177,7 @@
 
 section "value list conformance"
 
-constdefs
-
-  lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
-                                                ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70)
+definition lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70) where
            "G,s\<turnstile>vs[\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<exists>v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
 
 lemma lconfD: "\<lbrakk>G,s\<turnstile>vs[\<Colon>\<preceq>]Ts; Ts n = Some T\<rbrakk> \<Longrightarrow> G,s\<turnstile>(the (vs n))\<Colon>\<preceq>T"
@@ -267,10 +260,7 @@
 *}
 
   
-constdefs
-
-  wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
-                                          ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70)
+definition wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70) where
            "G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<forall> v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
 
 lemma wlconfD: "\<lbrakk>G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts; Ts n = Some T; vs n = Some v\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -348,9 +338,7 @@
 
 section "object conformance"
 
-constdefs
-
-  oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool"  ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_"  [71,71,71,71] 70)
+definition oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_"  [71,71,71,71] 70) where
            "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and> 
                            (case r of 
                               Heap a \<Rightarrow> is_type G (obj_ty obj) 
@@ -385,9 +373,7 @@
 
 section "state conformance"
 
-constdefs
-
-  conforms :: "state \<Rightarrow> env' \<Rightarrow> bool"          (     "_\<Colon>\<preceq>_"   [71,71]      70)
+definition conforms :: "state \<Rightarrow> env' \<Rightarrow> bool"   ("_\<Colon>\<preceq>_"   [71,71]      70)  where
    "xs\<Colon>\<preceq>E \<equiv> let (G, L) = E; s = snd xs; l = locals s in
     (\<forall>r. \<forall>obj\<in>globs s r:           G,s\<turnstile>obj   \<Colon>\<preceq>\<surd>r) \<and>
                 \<spacespace>                   G,s\<turnstile>l    [\<sim>\<Colon>\<preceq>]L\<spacespace> \<and>
--- a/src/HOL/Bali/Decl.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Decl.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -206,8 +206,7 @@
   "mdecl" <= (type) "sig \<times> methd"
 
 
-constdefs 
-  mhead::"methd \<Rightarrow> mhead"
+definition mhead :: "methd \<Rightarrow> mhead" where
   "mhead m \<equiv> \<lparr>access=access m, static=static m, pars=pars m, resT=resT m\<rparr>"
 
 lemma access_mhead [simp]:"access (mhead m) = access m"
@@ -275,7 +274,7 @@
 lemma memberid_pair_simp1: "memberid p  = memberid (snd p)"
 by (simp add: pair_memberid_def)
 
-constdefs is_field :: "qtname \<times> memberdecl \<Rightarrow> bool"
+definition is_field :: "qtname \<times> memberdecl \<Rightarrow> bool" where
 "is_field m \<equiv> \<exists> declC f. m=(declC,fdecl f)"
   
 lemma is_fieldD: "is_field m \<Longrightarrow> \<exists> declC f. m=(declC,fdecl f)"
@@ -284,7 +283,7 @@
 lemma is_fieldI: "is_field (C,fdecl f)"
 by (simp add: is_field_def)
 
-constdefs is_method :: "qtname \<times> memberdecl \<Rightarrow> bool"
+definition is_method :: "qtname \<times> memberdecl \<Rightarrow> bool" where
 "is_method membr \<equiv> \<exists> declC m. membr=(declC,mdecl m)"
   
 lemma is_methodD: "is_method membr \<Longrightarrow> \<exists> declC m. membr=(declC,mdecl m)"
@@ -315,8 +314,7 @@
                       isuperIfs::qtname list,\<dots>::'a\<rparr>"
   "idecl" <= (type) "qtname \<times> iface"
 
-constdefs
-  ibody :: "iface \<Rightarrow> ibody"
+definition ibody :: "iface \<Rightarrow> ibody" where
   "ibody i \<equiv> \<lparr>access=access i,imethods=imethods i\<rparr>"
 
 lemma access_ibody [simp]: "(access (ibody i)) = access i"
@@ -351,8 +349,7 @@
                       super::qtname,superIfs::qtname list,\<dots>::'a\<rparr>"
   "cdecl" <= (type) "qtname \<times> class"
 
-constdefs
-  cbody :: "class \<Rightarrow> cbody"
+definition cbody :: "class \<Rightarrow> cbody" where
   "cbody c \<equiv> \<lparr>access=access c, cfields=cfields c,methods=methods c,init=init c\<rparr>"
 
 lemma access_cbody [simp]:"access (cbody c) = access c"
@@ -394,7 +391,7 @@
 lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)"
 by (simp add: SXcptC_def)
 
-constdefs standard_classes :: "cdecl list"
+definition standard_classes :: "cdecl list" where
          "standard_classes \<equiv> [ObjectC, SXcptC Throwable,
                 SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
                 SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
@@ -470,7 +467,7 @@
   where "G|-C <:C D == (C,D) \<in>(subcls1 G)^+"
 
 notation (xsymbols)
-  subcls1_syntax  ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_"  [71,71,71] 70) and
+  subcls1_syntax  ("_\<turnstile>_\<prec>\<^sub>C1_"  [71,71,71] 70) and
   subclseq_syntax  ("_\<turnstile>_\<preceq>\<^sub>C _"  [71,71,71] 70) and
   subcls_syntax  ("_\<turnstile>_\<prec>\<^sub>C _"  [71,71,71] 70)
 
@@ -510,7 +507,7 @@
 "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D\<rbrakk> \<Longrightarrow> \<exists> c. class G C = Some c"
 by (auto simp add: subcls1_def dest: tranclD)
 
-lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C\<^sub>1 D \<Longrightarrow> P"
+lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C1 D \<Longrightarrow> P"
 by (auto simp add: subcls1_def)
 
 lemma no_subcls_Object: "G\<turnstile>Object\<prec>\<^sub>C D \<Longrightarrow> P"
@@ -520,14 +517,13 @@
 
 section "well-structured programs"
 
-constdefs
-  ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool"
+definition ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool" where
  "ws_idecl G I si \<equiv> \<forall>J\<in>set si.  is_iface G J   \<and> (J,I)\<notin>(subint1 G)^+"
   
-  ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+definition ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" where
  "ws_cdecl G C sc \<equiv> C\<noteq>Object \<longrightarrow> is_class G sc \<and> (sc,C)\<notin>(subcls1 G)^+"
   
-  ws_prog  :: "prog \<Rightarrow> bool"
+definition ws_prog  :: "prog \<Rightarrow> bool" where
  "ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces  G). ws_idecl G I (isuperIfs i)) \<and> 
               (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
 
@@ -680,7 +676,7 @@
   then have "is_class G C \<Longrightarrow> P C"  
   proof (induct rule: subcls1_induct)
     fix C
-    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
+    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
        and iscls:"is_class G C"
     show "P C"
     proof (cases "C=Object")
@@ -715,7 +711,7 @@
   then have "\<And> c. class G C = Some c\<Longrightarrow> P C c"  
   proof (induct rule: subcls1_induct)
     fix C c
-    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> P S s)"
+    assume   hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> P S s)"
        and iscls:"class G C = Some c"
     show "P C c"
     proof (cases "C=Object")
@@ -725,7 +721,7 @@
       with ws iscls obtain sc where
         sc: "class G (super c) = Some sc"
         by (auto dest: ws_prog_cdeclD)
-      from iscls False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
+      from iscls False have "G\<turnstile>C \<prec>\<^sub>C1 (super c)" by (rule subcls1I)
       with False ws step hyp iscls sc
       show "P C c" 
         by (auto)  
@@ -767,52 +763,57 @@
 
 section "general recursion operators for the interface and class hiearchies"
 
-consts
-  iface_rec  :: "prog \<times> qtname \<Rightarrow>   \<spacespace>  (qtname \<Rightarrow> iface \<Rightarrow> 'a set \<Rightarrow> 'a) \<Rightarrow> 'a"
-  class_rec  :: "prog \<times> qtname \<Rightarrow> 'a \<Rightarrow> (qtname \<Rightarrow> class \<Rightarrow> 'a     \<Rightarrow> 'a) \<Rightarrow> 'a"
-
-recdef iface_rec "same_fst ws_prog (\<lambda>G. (subint1 G)^-1)" 
-"iface_rec (G,I) = 
-  (\<lambda>f. case iface G I of 
+function
+  iface_rec  :: "prog \<Rightarrow> qtname \<Rightarrow>   \<spacespace>(qtname \<Rightarrow> iface \<Rightarrow> 'a set \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+[simp del]: "iface_rec G I f = 
+  (case iface G I of 
          None \<Rightarrow> undefined 
        | Some i \<Rightarrow> if ws_prog G 
                       then f I i 
-                               ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))
+                               ((\<lambda>J. iface_rec G J f)`set (isuperIfs i))
                       else undefined)"
-(hints recdef_wf: wf_subint1 intro: subint1I)
-declare iface_rec.simps [simp del]
+by auto
+termination
+by (relation "inv_image (same_fst ws_prog (\<lambda>G. (subint1 G)^-1)) (%(x,y,z). (x,y))")
+ (auto simp: wf_subint1 subint1I wf_same_fst)
 
 lemma iface_rec: 
 "\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow> 
- iface_rec (G,I) f = f I i ((\<lambda>J. iface_rec (G,J) f)`set (isuperIfs i))"
+ iface_rec G I f = f I i ((\<lambda>J. iface_rec G J f)`set (isuperIfs i))"
 apply (subst iface_rec.simps)
 apply simp
 done
 
-recdef class_rec "same_fst ws_prog (\<lambda>G. (subcls1 G)^-1)"
-"class_rec(G,C) = 
-  (\<lambda>t f. case class G C of 
+
+function
+  class_rec  :: "prog \<Rightarrow> qtname \<Rightarrow> 'a \<Rightarrow> (qtname \<Rightarrow> class \<Rightarrow> 'a     \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+[simp del]: "class_rec G C t f = 
+  (case class G C of 
            None \<Rightarrow> undefined 
          | Some c \<Rightarrow> if ws_prog G 
                         then f C c 
                                  (if C = Object then t 
-                                                else class_rec (G,super c) t f)
+                                                else class_rec G (super c) t f)
                         else undefined)"
-(hints recdef_wf: wf_subcls1 intro: subcls1I)
-declare class_rec.simps [simp del]
+
+by auto
+termination
+by (relation "inv_image (same_fst ws_prog (\<lambda>G. (subcls1 G)^-1)) (%(x,y,z,w). (x,y))")
+ (auto simp: wf_subcls1 subcls1I wf_same_fst)
 
 lemma class_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>  
- class_rec (G,C) t f = 
-   f C c (if C = Object then t else class_rec (G,super c) t f)"
-apply (rule class_rec.simps [THEN trans [THEN fun_cong [THEN fun_cong]]])
+ class_rec G C t f = 
+   f C c (if C = Object then t else class_rec G (super c) t f)"
+apply (subst class_rec.simps)
 apply simp
 done
 
-constdefs
-imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
   --{* methods of an interface, with overriding and inheritance, cf. 9.2 *}
 "imethds G I 
-  \<equiv> iface_rec (G,I)  
+  \<equiv> iface_rec G I  
               (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
                         (Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
         
--- a/src/HOL/Bali/DeclConcepts.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/DeclConcepts.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,8 +8,7 @@
 
 section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)"
 
-constdefs
-is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
+definition is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool" where
 "is_public G qn \<equiv> (case class G qn of
                      None       \<Rightarrow> (case iface G qn of
                                       None       \<Rightarrow> False
@@ -38,14 +37,16 @@
 
 declare accessible_in_RefT_simp [simp del]
 
-constdefs
-  is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+definition is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool" where
     "is_acc_class G P C \<equiv> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
-  is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+
+definition is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool" where
     "is_acc_iface G P I \<equiv> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
-  is_acc_type  :: "prog \<Rightarrow> pname \<Rightarrow> ty     \<Rightarrow> bool"
+
+definition is_acc_type  :: "prog \<Rightarrow> pname \<Rightarrow> ty     \<Rightarrow> bool" where
     "is_acc_type  G P T \<equiv> is_type G T  \<and> G\<turnstile>T accessible_in P"
-  is_acc_reftype  :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool"
+
+definition is_acc_reftype  :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool" where
   "is_acc_reftype  G P T \<equiv> isrtype G T  \<and> G\<turnstile>T accessible_in' P"
 
 lemma is_acc_classD:
@@ -336,8 +337,7 @@
 text {* Convert a qualified method declaration (qualified with its declaring 
 class) to a qualified member declaration:  @{text methdMembr}  *}
 
-constdefs
-methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)"
+definition methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)" where
  "methdMembr m \<equiv> (fst m,mdecl (snd m))"
 
 lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)"
@@ -355,8 +355,7 @@
 text {* Convert a qualified method (qualified with its declaring 
 class) to a qualified member declaration:  @{text method}  *}
 
-constdefs
-method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)" 
+definition method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)" where 
 "method sig m \<equiv> (declclass m, mdecl (sig, mthd m))"
 
 lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))"
@@ -377,8 +376,7 @@
 lemma memberid_method_simp[simp]:  "memberid (method sig m) = mid sig"
   by (simp add: method_def) 
 
-constdefs
-fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)" 
+definition fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)" where 
 "fieldm n f \<equiv> (declclass f, fdecl (n, fld f))"
 
 lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))"
@@ -402,7 +400,7 @@
 text {* Select the signature out of a qualified method declaration:
  @{text msig} *}
 
-constdefs msig:: "(qtname \<times> mdecl) \<Rightarrow> sig"
+definition msig :: "(qtname \<times> mdecl) \<Rightarrow> sig" where
 "msig m \<equiv> fst (snd m)"
 
 lemma msig_simp[simp]: "msig (c,(s,m)) = s"
@@ -411,7 +409,7 @@
 text {* Convert a qualified method (qualified with its declaring 
 class) to a qualified method declaration:  @{text qmdecl}  *}
 
-constdefs qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
+definition qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)" where
 "qmdecl sig m \<equiv> (declclass m, (sig,mthd m))"
 
 lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))"
@@ -504,10 +502,8 @@
       it is not accessible for inheritance at all.
 \end{itemize}
 *}
-constdefs
-inheritable_in:: 
- "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool"
-                  ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
+definition inheritable_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60) where
+                  
 "G\<turnstile>membr inheritable_in pack 
   \<equiv> (case (accmodi membr) of
        Private   \<Rightarrow> False
@@ -529,25 +525,21 @@
 
 subsubsection "declared-in/undeclared-in"
 
-constdefs cdeclaredmethd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table"
+definition cdeclaredmethd :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table" where
 "cdeclaredmethd G C 
   \<equiv> (case class G C of
        None \<Rightarrow> \<lambda> sig. None
      | Some c \<Rightarrow> table_of (methods c)
     )"
 
-constdefs
-cdeclaredfield:: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table"
+definition cdeclaredfield :: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table" where
 "cdeclaredfield G C 
   \<equiv> (case class G C of
        None \<Rightarrow> \<lambda> sig. None
      | Some c \<Rightarrow> table_of (cfields c)
     )"
 
-
-constdefs
-declared_in:: "prog  \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool"
-                                 ("_\<turnstile> _ declared'_in _" [61,61,61] 60)
+definition declared_in :: "prog  \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ declared'_in _" [61,61,61] 60) where
 "G\<turnstile>m declared_in C \<equiv> (case m of
                         fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn  = Some f
                       | mdecl (sig,m) \<Rightarrow> cdeclaredmethd G C sig = Some m)"
@@ -567,10 +559,7 @@
 by (cases m) 
    (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
 
-constdefs
-undeclared_in:: "prog  \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool"
-                                 ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
-
+definition undeclared_in :: "prog  \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60) where
 "G\<turnstile>m undeclared_in C \<equiv> (case m of
                             fid fn  \<Rightarrow> cdeclaredfield G C fn  = None
                           | mid sig \<Rightarrow> cdeclaredmethd G C sig = None)"
@@ -591,7 +580,7 @@
 
   Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
 | Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C; 
-               G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
+               G\<turnstile>C \<prec>\<^sub>C1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S 
               \<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
 text {* Note that in the case of an inherited member only the members of the
 direct superclass are concerned. If a member of a superclass of the direct
@@ -617,19 +606,16 @@
                            ("_ \<turnstile>Field _  _ member'_of _" [61,61,61] 60)
  where "G\<turnstile>Field n f member_of C == G\<turnstile>fieldm n f member_of C"
 
-constdefs
-inherits:: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool"
-                           ("_ \<turnstile> _ inherits _" [61,61,61] 60)
+definition inherits :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool" ("_ \<turnstile> _ inherits _" [61,61,61] 60) where
 "G\<turnstile>C inherits m 
   \<equiv> G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and> 
-    (\<exists> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
+    (\<exists> S. G\<turnstile>C \<prec>\<^sub>C1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
 
 lemma inherits_member: "G\<turnstile>C inherits m \<Longrightarrow> G\<turnstile>m member_of C"
 by (auto simp add: inherits_def intro: members.Inherited)
 
 
-constdefs member_in::"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
-                           ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
+definition member_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ member'_in _" [61,61,61] 60) where
 "G\<turnstile>m member_in C \<equiv> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
 text {* A member is in a class if it is member of the class or a superclass.
 If a member is in a class we can select this member. This additional notion
@@ -676,7 +662,7 @@
            G\<turnstile>Method new declared_in (declclass new);  
            G\<turnstile>Method old declared_in (declclass old); 
            G\<turnstile>Method old inheritable_in pid (declclass new);
-           G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
+           G\<turnstile>(declclass new) \<prec>\<^sub>C1 superNew;
            G \<turnstile>Method old member_of superNew
            \<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
 
@@ -716,9 +702,7 @@
 
 subsubsection "Hiding"
 
-constdefs hides::
-"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" 
-                                ("_\<turnstile> _ hides _" [61,61,61] 60)
+definition hides :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" ("_\<turnstile> _ hides _" [61,61,61] 60) where 
 "G\<turnstile>new hides old
   \<equiv> is_static new \<and> msig new = msig old \<and>
     G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
@@ -777,11 +761,7 @@
 by (auto simp add: hides_def)
 
 subsubsection "permits access" 
-constdefs 
-permits_acc:: 
- "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
-                   ("_ \<turnstile> _ in _ permits'_acc'_from _" [61,61,61,61] 60)
-
+definition permits_acc :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ in _ permits'_acc'_from _" [61,61,61,61] 60) where
 "G\<turnstile>membr in cls permits_acc_from accclass 
   \<equiv> (case (accmodi membr) of
        Private   \<Rightarrow> (declclass membr = accclass)
@@ -980,7 +960,7 @@
   next
     case (Inherited n C S)
     assume undecl: "G\<turnstile> memberid n undeclared_in C"
-    assume  super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
+    assume  super: "G\<turnstile>C\<prec>\<^sub>C1S"
     assume    hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
     assume   eqid: "memberid n = memberid m"
     assume "G \<turnstile> m member_of C"
@@ -1011,7 +991,7 @@
        (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
 next
   case (Inherited m C S)  
-  assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
+  assume "G\<turnstile>C\<prec>\<^sub>C1S" and "is_class G S"
   then show "is_class G C"
     by - (rule subcls_is_class2,auto)
 qed
@@ -1043,7 +1023,7 @@
         intro: rtrancl_trans)
 
 lemma stat_override_declclasses_relation: 
-"\<lbrakk>G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew; G \<turnstile>Method old member_of superNew \<rbrakk>
+"\<lbrakk>G\<turnstile>(declclass new) \<prec>\<^sub>C1 superNew; G \<turnstile>Method old member_of superNew \<rbrakk>
 \<Longrightarrow> G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old)"
 apply (rule trancl_rtrancl_trancl)
 apply (erule r_into_trancl)
@@ -1257,7 +1237,7 @@
             "G\<turnstile> memberid m undeclared_in D"  
             "G \<turnstile> Class S accessible_in pid D" 
             "G \<turnstile> m member_of S"
-    assume super: "G\<turnstile>D\<prec>\<^sub>C\<^sub>1S"
+    assume super: "G\<turnstile>D\<prec>\<^sub>C1S"
     assume hyp: "\<lbrakk>G\<turnstile>S\<preceq>\<^sub>C C; G\<turnstile>C\<preceq>\<^sub>C declclass m\<rbrakk> \<Longrightarrow> G \<turnstile> m member_of C"
     assume subclseq_C_m: "G\<turnstile>C\<preceq>\<^sub>C declclass m"
     assume "G\<turnstile>D\<preceq>\<^sub>C C"
@@ -1399,27 +1379,24 @@
 translations 
   "fspec" <= (type) "vname \<times> qtname" 
 
-constdefs
-imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
 "imethds G I 
-  \<equiv> iface_rec (G,I)  
+  \<equiv> iface_rec G I  
               (\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> 
                         (Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
 text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
 
-constdefs
-accimethds:: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition accimethds :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
 "accimethds G pack I
   \<equiv> if G\<turnstile>Iface I accessible_in pack 
        then imethds G I
        else \<lambda> k. {}"
 text {* only returns imethds if the interface is accessible *}
 
-constdefs
-methd:: "prog \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
+definition methd :: "prog \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table" where
 
 "methd G C 
- \<equiv> class_rec (G,C) empty
+ \<equiv> class_rec G C empty
              (\<lambda>C c subcls_mthds. 
                filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
                           subcls_mthds 
@@ -1431,8 +1408,7 @@
      Every new method with the same signature coalesces the
      method of a superclass. *}
 
-constdefs                      
-accmethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table"
+definition accmethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table" where
 "accmethd G S C 
  \<equiv> filter_tab (\<lambda>sig m. G\<turnstile>method sig m of C accessible_from S) 
               (methd G C)"
@@ -1446,15 +1422,14 @@
     So we must test accessibility of method @{term m} of class @{term C} 
     (not @{term "declclass m"}) *}
 
-constdefs 
-dynmethd:: "prog  \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynmethd :: "prog  \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
 "dynmethd G statC dynC  
   \<equiv> \<lambda> sig. 
      (if G\<turnstile>dynC \<preceq>\<^sub>C statC
         then (case methd G statC sig of
                 None \<Rightarrow> None
               | Some statM 
-                  \<Rightarrow> (class_rec (G,dynC) empty
+                  \<Rightarrow> (class_rec G dynC empty
                        (\<lambda>C c subcls_mthds. 
                           subcls_mthds
                           ++
@@ -1473,8 +1448,7 @@
         filters the new methods (to get only those methods which actually
         override the methods of the static class) *}
 
-constdefs 
-dynimethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynimethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
 "dynimethd G I dynC 
   \<equiv> \<lambda> sig. if imethds G I sig \<noteq> {}
                then methd G dynC sig
@@ -1493,8 +1467,7 @@
    down to the actual dynamic class.
  *}
 
-constdefs
-dynlookup::"prog  \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynlookup :: "prog  \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
 "dynlookup G statT dynC
   \<equiv> (case statT of
        NullT        \<Rightarrow> empty
@@ -1506,17 +1479,15 @@
     In a wellformd context statT will not be NullT and in case
     statT is an array type, dynC=Object *}
 
-constdefs
-fields:: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list"
+definition fields :: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list" where
 "fields G C 
-  \<equiv> class_rec (G,C) [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
+  \<equiv> class_rec G C [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
 text {* @{term "fields G C"} 
      list of fields of a class, including all the fields of the superclasses
      (private, inherited and hidden ones) not only the accessible ones
      (an instance of a object allocates all these fields *}
 
-constdefs
-accfield:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname  \<times>  field) table"
+definition accfield :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname  \<times>  field) table" where
 "accfield G S C
   \<equiv> let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
     in filter_tab (\<lambda>n (declC,f). G\<turnstile> (declC,fdecl (n,f)) of C accessible_from S)
@@ -1531,12 +1502,10 @@
    inheritance, too. So we must test accessibility of field @{term f} of class 
    @{term C} (not @{term "declclass f"}) *} 
 
-constdefs
-
-  is_methd :: "prog \<Rightarrow> qtname  \<Rightarrow> sig \<Rightarrow> bool"
+definition is_methd :: "prog \<Rightarrow> qtname  \<Rightarrow> sig \<Rightarrow> bool" where
  "is_methd G \<equiv> \<lambda>C sig. is_class G C \<and> methd G C sig \<noteq> None"
 
-constdefs efname:: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)"
+definition efname :: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)" where
 "efname \<equiv> fst"
 
 lemma efname_simp[simp]:"efname (n,f) = n"
@@ -1836,7 +1805,7 @@
                 (\<lambda>_ dynM. G,sig \<turnstile> dynM overrides statM \<or> dynM = statM)
                 (methd G C)"
         let "?class_rec C" =
-              "(class_rec (G, C) empty
+              "(class_rec G C empty
                            (\<lambda>C c subcls_mthds. subcls_mthds ++ (?filter C)))"
         from statM Subcls ws subclseq_dynC_statC
         have dynmethd_dynC_def:
@@ -2300,9 +2269,8 @@
 
 section "calculation of the superclasses of a class"
 
-constdefs 
- superclasses:: "prog \<Rightarrow> qtname \<Rightarrow> qtname set"
- "superclasses G C \<equiv> class_rec (G,C) {} 
+definition superclasses :: "prog \<Rightarrow> qtname \<Rightarrow> qtname set" where
+ "superclasses G C \<equiv> class_rec G C {} 
                        (\<lambda> C c superclss. (if C=Object 
                                             then {} 
                                             else insert (super c) superclss))"
--- a/src/HOL/Bali/DefiniteAssignment.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/DefiniteAssignment.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -74,7 +74,7 @@
 "jumpNestingOkS jmps (FinA a c) = False"
 
 
-constdefs jumpNestingOk :: "jump set \<Rightarrow> term \<Rightarrow> bool"
+definition jumpNestingOk :: "jump set \<Rightarrow> term \<Rightarrow> bool" where
 "jumpNestingOk jmps t \<equiv> (case t of
                       In1 se \<Rightarrow> (case se of
                                    Inl e \<Rightarrow> True
@@ -156,7 +156,7 @@
 "assignsEs     [] = {}"
 "assignsEs (e#es) = assignsE e \<union> assignsEs es"
 
-constdefs assigns:: "term \<Rightarrow> lname set"
+definition assigns :: "term \<Rightarrow> lname set" where
 "assigns t \<equiv> (case t of
                 In1 se \<Rightarrow> (case se of
                              Inl e \<Rightarrow> assignsE e
@@ -429,20 +429,14 @@
 
 subsection {* Lifting set operations to range of tables (map to a set) *}
 
-constdefs 
- union_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
-                    ("_ \<Rightarrow>\<union> _" [67,67] 65)
+definition union_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<union> _" [67,67] 65) where
  "A \<Rightarrow>\<union> B \<equiv> \<lambda> k. A k \<union> B k"
 
-constdefs
- intersect_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
-                    ("_ \<Rightarrow>\<inter>  _" [72,72] 71)
+definition intersect_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<inter>  _" [72,72] 71) where
  "A \<Rightarrow>\<inter>  B \<equiv> \<lambda> k. A k \<inter> B k"
 
-constdefs
- all_union_ts:: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" 
-                                                     (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40)
-"A \<Rightarrow>\<union>\<^sub>\<forall> B \<equiv> \<lambda> k. A k \<union> B"
+definition all_union_ts :: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40) where 
+ "A \<Rightarrow>\<union>\<^sub>\<forall> B \<equiv> \<lambda> k. A k \<union> B"
   
 subsubsection {* Binary union of tables *}
 
@@ -513,15 +507,15 @@
          brk :: "breakass" --{* Definetly assigned variables for 
                                 abrupt completion with a break *}
 
-constdefs rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
+definition rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" where
 "rmlab k A \<equiv> \<lambda> x. if x=k then UNIV else A x"
  
 (*
-constdefs setbrk :: "breakass \<Rightarrow> assigned \<Rightarrow> breakass set"
+definition setbrk :: "breakass \<Rightarrow> assigned \<Rightarrow> breakass set" where
 "setbrk b A \<equiv> {b} \<union> {a| a. a\<in> brk A \<and> lab a \<noteq> lab b}"
 *)
 
-constdefs range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80)
+definition range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80) where 
  "\<Rightarrow>\<Inter>A \<equiv> {x |x. \<forall> k. x \<in> A k}"
 
 text {*
--- a/src/HOL/Bali/Eval.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Eval.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -140,8 +140,7 @@
   lst_inj_vals  ("\<lfloor>_\<rfloor>\<^sub>l" 1000)
   where "\<lfloor>es\<rfloor>\<^sub>l == In3 es"
 
-constdefs
-  undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
+definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals" where
  "undefined3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. undefined) (\<lambda>x. Unit))
                      (\<lambda>x. In2 undefined) (\<lambda>x. In3 undefined)"
 
@@ -160,8 +159,7 @@
 
 section "exception throwing and catching"
 
-constdefs
-  throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt"
+definition throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt" where
  "throw a' x \<equiv> abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
 
 lemma throw_def2: 
@@ -170,8 +168,7 @@
 apply (simp (no_asm))
 done
 
-constdefs
-  fits    :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
+definition fits :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
  "G,s\<turnstile>a' fits T  \<equiv> (\<exists>rt. T=RefT rt) \<longrightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
 
 lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
@@ -195,8 +192,7 @@
 apply iprover
 done
 
-constdefs
-  catch ::"prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool"      ("_,_\<turnstile>catch _"[61,61,61]60)
+definition catch :: "prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
  "G,s\<turnstile>catch C\<equiv>\<exists>xc. abrupt s=Some (Xcpt xc) \<and> 
                     G,store s\<turnstile>Addr (the_Loc xc) fits Class C"
 
@@ -221,8 +217,7 @@
 apply (simp (no_asm))
 done
 
-constdefs
-  new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
+definition new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state" where
  "new_xcpt_var vn \<equiv> 
      \<lambda>(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
 
@@ -237,9 +232,7 @@
 
 section "misc"
 
-constdefs
-
-  assign     :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
+definition assign :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state" where
  "assign f v \<equiv> \<lambda>(x,s). let (x',s') = (if x = None then f v else id) (x,s)
                    in  (x',if x' = None then s' else s)"
 
@@ -300,9 +293,7 @@
 done
 *)
 
-constdefs
-
-  init_comp_ty :: "ty \<Rightarrow> stmt"
+definition init_comp_ty :: "ty \<Rightarrow> stmt" where
  "init_comp_ty T \<equiv> if (\<exists>C. T = Class C) then Init (the_Class T) else Skip"
 
 lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip"
@@ -310,9 +301,7 @@
 apply (simp (no_asm))
 done
 
-constdefs
-
- invocation_class  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
+definition invocation_class :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname" where
  "invocation_class m s a' statT 
     \<equiv> (case m of
          Static \<Rightarrow> if (\<exists> statC. statT = ClassT statC) 
@@ -321,7 +310,7 @@
        | SuperM \<Rightarrow> the_Class (RefT statT)
        | IntVir \<Rightarrow> obj_class (lookup_obj s a'))"
 
-invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname"
+definition invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname" where
 "invocation_declclass G m s a' statT sig 
    \<equiv> declclass (the (dynlookup G statT 
                                 (invocation_class m s a' statT)
@@ -341,9 +330,8 @@
                                             else Object)"
 by (simp add: invocation_class_def)
 
-constdefs
-  init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
-                   state \<Rightarrow> state"
+definition init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
+                   state \<Rightarrow> state" where
  "init_lvars G C sig mode a' pvs 
    \<equiv> \<lambda> (x,s). 
       let m = mthd (the (methd G C sig));
@@ -376,8 +364,7 @@
 apply (simp (no_asm) add: Let_def)
 done
 
-constdefs
-  body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr"
+definition body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr" where
  "body G C sig \<equiv> let m = the (methd G C sig) 
                  in Body (declclass m) (stmt (mbody (mthd m)))"
 
@@ -390,12 +377,10 @@
 
 section "variables"
 
-constdefs
-
-  lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar"
+definition lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar" where
  "lvar vn s \<equiv> (the (locals s vn), \<lambda>v. supd (lupd(vn\<mapsto>v)))"
 
-  fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+definition fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state" where
  "fvar C stat fn a' s 
     \<equiv> let (oref,xf) = if stat then (Stat C,id)
                               else (Heap (the_Addr a'),np a');
@@ -403,7 +388,7 @@
                   f = (\<lambda>v. supd (upd_gobj oref n v)) 
       in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
 
-  avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+definition avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state" where
  "avar G i' a' s 
     \<equiv> let   oref = Heap (the_Addr a'); 
                i = the_Intg i'; 
@@ -446,9 +431,7 @@
 apply (simp (no_asm) add: Let_def split_beta)
 done
 
-constdefs
-check_field_access::
-"prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+definition check_field_access :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" where
 "check_field_access G accC statDeclC fn stat a' s
  \<equiv> let oref = if stat then Stat statDeclC
                       else Heap (the_Addr a');
@@ -461,9 +444,7 @@
                   AccessViolation)
         s"
 
-constdefs
-check_method_access:: 
-  "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow>  sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+definition check_method_access :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow>  sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" where
 "check_method_access G accC statT mode sig  a' s
  \<equiv> let invC = invocation_class mode (store s) a' statT;
        dynM = the (dynlookup G statT invC sig)
--- a/src/HOL/Bali/Example.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Example.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -153,23 +153,18 @@
   
   foo    :: mname
 
-constdefs
-  
-  foo_sig   :: sig
- "foo_sig   \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
+definition foo_sig :: sig
+ where "foo_sig   \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
   
-  foo_mhead :: mhead
- "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
+definition foo_mhead :: mhead
+ where "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
 
-constdefs
-  
-  Base_foo :: mdecl
- "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
+definition Base_foo :: mdecl
+ where "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
                         mbody=\<lparr>lcls=[],stmt=Return (!!z)\<rparr>\<rparr>)"
 
-constdefs
-  Ext_foo  :: mdecl
- "Ext_foo  \<equiv> (foo_sig, 
+definition Ext_foo :: mdecl
+ where "Ext_foo  \<equiv> (foo_sig, 
               \<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
                mbody=\<lparr>lcls=[]
                      ,stmt=Expr({Ext,Ext,False}Cast (Class Ext) (!!z)..vee := 
@@ -177,12 +172,10 @@
                                 Return (Lit Null)\<rparr>
               \<rparr>)"
 
-constdefs
-  
-arr_viewed_from :: "qtname \<Rightarrow> qtname \<Rightarrow> var"
+definition arr_viewed_from :: "qtname \<Rightarrow> qtname \<Rightarrow> var" where
 "arr_viewed_from accC C \<equiv> {accC,Base,True}StatRef (ClassT C)..arr"
 
-BaseCl :: "class"
+definition BaseCl :: "class" where
 "BaseCl \<equiv> \<lparr>access=Public,
            cfields=[(arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
                     (vee, \<lparr>access=Public,static=False,type=Iface HasFoo    \<rparr>)],
@@ -192,7 +185,7 @@
            super=Object,
            superIfs=[HasFoo]\<rparr>"
   
-ExtCl  :: "class"
+definition ExtCl  :: "class" where
 "ExtCl  \<equiv> \<lparr>access=Public,
            cfields=[(vee, \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)], 
            methods=[Ext_foo],
@@ -200,7 +193,7 @@
            super=Base,
            superIfs=[]\<rparr>"
 
-MainCl :: "class"
+definition MainCl :: "class" where
 "MainCl \<equiv> \<lparr>access=Public,
            cfields=[], 
            methods=[], 
@@ -209,16 +202,14 @@
            superIfs=[]\<rparr>"
 (* The "main" method is modeled seperately (see tprg) *)
 
-constdefs
-  
-  HasFooInt :: iface
- "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
+definition HasFooInt :: iface
+ where "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
 
-  Ifaces ::"idecl list"
- "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
+definition Ifaces ::"idecl list"
+ where "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
 
- "Classes" ::"cdecl list"
- "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes"
+definition "Classes" ::"cdecl list"
+ where "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes"
 
 lemmas table_classes_defs = 
      Classes_def standard_classes_def ObjectC_def SXcptC_def
@@ -273,8 +264,7 @@
   tprg :: prog where
   "tprg == \<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
 
-constdefs
-  test    :: "(ty)list \<Rightarrow> stmt"
+definition test :: "(ty)list \<Rightarrow> stmt" where
  "test pTs \<equiv> e:==NewC Ext;; 
            \<spacespace> Try Expr({Main,ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))
            \<spacespace> Catch((SXcpt NullPointer) z)
--- a/src/HOL/Bali/State.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/State.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -38,9 +38,7 @@
   "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option\<rparr>"
   "obj"   <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option,\<dots>::'a\<rparr>"
 
-constdefs
-  
-  the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table"
+definition the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table" where
  "the_Arr obj \<equiv> SOME (T,k,t). obj = Some \<lparr>tag=Arr T k,values=t\<rparr>"
 
 lemma the_Arr_Arr [simp]: "the_Arr (Some \<lparr>tag=Arr T k,values=cs\<rparr>) = (T,k,cs)"
@@ -52,9 +50,7 @@
 apply (auto simp add: the_Arr_def)
 done
 
-constdefs
-
-  upd_obj       :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" 
+definition upd_obj :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" where 
  "upd_obj n v \<equiv> \<lambda> obj . obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>"
 
 lemma upd_obj_def2 [simp]: 
@@ -62,8 +58,7 @@
 apply (auto simp: upd_obj_def)
 done
 
-constdefs
-  obj_ty        :: "obj \<Rightarrow> ty"
+definition obj_ty :: "obj \<Rightarrow> ty" where
  "obj_ty obj    \<equiv> case tag obj of 
                     CInst C \<Rightarrow> Class C 
                   | Arr T k \<Rightarrow> T.[]"
@@ -102,9 +97,7 @@
 apply (auto split add: obj_tag.split_asm)
 done
 
-constdefs
-
-  obj_class :: "obj \<Rightarrow> qtname"
+definition obj_class :: "obj \<Rightarrow> qtname" where
  "obj_class obj \<equiv> case tag obj of 
                     CInst C \<Rightarrow> C 
                   | Arr T k \<Rightarrow> Object"
@@ -143,9 +136,7 @@
   "Stat" => "CONST Inr"
   "oref" <= (type) "loc + qtname"
 
-constdefs
-  fields_table::
-    "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool)  \<Rightarrow> (fspec, ty) table"
+definition fields_table :: "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool)  \<Rightarrow> (fspec, ty) table" where
  "fields_table G C P 
     \<equiv> Option.map type \<circ> table_of (filter (split P) (DeclConcepts.fields G C))"
 
@@ -182,14 +173,13 @@
 apply simp
 done
 
-constdefs
-  in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool"            ("(_/ in'_bounds _)" [50, 51] 50)
+definition in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool" ("(_/ in'_bounds _)" [50, 51] 50) where
  "i in_bounds k \<equiv> 0 \<le> i \<and> i < k"
 
-  arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option"
+definition arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option" where
  "arr_comps T k \<equiv> \<lambda>i. if i in_bounds k then Some T else None"
   
-  var_tys       :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table"
+definition var_tys       :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table" where
 "var_tys G oi r 
   \<equiv> case r of 
       Heap a \<Rightarrow> (case oi of 
@@ -232,15 +222,13 @@
 
 subsection "access"
 
-constdefs
-
-  globs  :: "st \<Rightarrow> globs"
+definition globs :: "st \<Rightarrow> globs" where
  "globs  \<equiv> st_case (\<lambda>g l. g)"
   
-  locals :: "st \<Rightarrow> locals"
+definition locals :: "st \<Rightarrow> locals" where
  "locals \<equiv> st_case (\<lambda>g l. l)"
 
-  heap   :: "st \<Rightarrow> heap"
+definition heap   :: "st \<Rightarrow> heap" where
  "heap s \<equiv> globs s \<circ> Heap"
 
 
@@ -262,8 +250,7 @@
 
 subsection "memory allocation"
 
-constdefs
-  new_Addr     :: "heap \<Rightarrow> loc option"
+definition new_Addr :: "heap \<Rightarrow> loc option" where
  "new_Addr h   \<equiv> if (\<forall>a. h a \<noteq> None) then None else Some (SOME a. h a = None)"
 
 lemma new_AddrD: "new_Addr h = Some a \<Longrightarrow> h a = None"
@@ -303,20 +290,19 @@
 
 subsection "update"
 
-constdefs
-  gupd       :: "oref  \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st"        ("gupd'(_\<mapsto>_')"[10,10]1000)
+definition gupd :: "oref  \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')"[10,10]1000) where
  "gupd r obj  \<equiv> st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
 
-  lupd       :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"        ("lupd'(_\<mapsto>_')"[10,10]1000)
+definition lupd       :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')"[10,10]1000) where
  "lupd vn v   \<equiv> st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))"
 
-  upd_gobj   :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"
+definition upd_gobj   :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" where
  "upd_gobj r n v \<equiv> st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
 
-  set_locals  :: "locals \<Rightarrow> st \<Rightarrow> st"
+definition set_locals  :: "locals \<Rightarrow> st \<Rightarrow> st" where
  "set_locals l \<equiv> st_case (\<lambda>g l'. st g l)"
 
-  init_obj    :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st"
+definition init_obj    :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st" where
  "init_obj G oi r \<equiv> gupd(r\<mapsto>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>)"
 
 abbreviation
@@ -476,8 +462,7 @@
 
         
 
-constdefs
-  abrupt_if    :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt"
+definition abrupt_if :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt" where
  "abrupt_if c x' x \<equiv> if c \<and> (x = None) then x' else x"
 
 lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x"
@@ -557,8 +542,7 @@
 apply auto
 done
 
-constdefs
-   absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt"
+definition absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt" where
   "absorb j a \<equiv> if a=Some (Jump j) then None else a"
 
 lemma absorb_SomeD [dest!]: "absorb j a = Some x \<Longrightarrow> a = Some x"
@@ -611,9 +595,7 @@
 apply clarsimp
 done
 
-constdefs
-
-  normal     :: "state \<Rightarrow> bool"
+definition normal :: "state \<Rightarrow> bool" where
  "normal \<equiv> \<lambda>s. abrupt s = None"
 
 lemma normal_def2 [simp]: "normal s = (abrupt s = None)"
@@ -621,8 +603,7 @@
 apply (simp (no_asm))
 done
 
-constdefs
-  heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool"
+definition heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool" where
  "heap_free n \<equiv> \<lambda>s. atleast_free (heap (store s)) n"
 
 lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n"
@@ -632,12 +613,10 @@
 
 subsection "update"
 
-constdefs
- 
-  abupd     :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state"
+definition abupd :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state" where
  "abupd f \<equiv> prod_fun f id"
 
-  supd     :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" 
+definition supd     :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" where
  "supd \<equiv> prod_fun id"
   
 lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)"
@@ -692,12 +671,10 @@
 
 section "initialisation test"
 
-constdefs
-
-  inited   :: "qtname \<Rightarrow> globs \<Rightarrow> bool"
+definition inited :: "qtname \<Rightarrow> globs \<Rightarrow> bool" where
  "inited C g \<equiv> g (Stat C) \<noteq> None"
 
-  initd    :: "qtname \<Rightarrow> state \<Rightarrow> bool"
+definition initd    :: "qtname \<Rightarrow> state \<Rightarrow> bool" where
  "initd C \<equiv> inited C \<circ> globs \<circ> store"
 
 lemma not_inited_empty [simp]: "\<not>inited C empty"
@@ -731,7 +708,7 @@
 done
 
 section {* @{text error_free} *}
-constdefs error_free:: "state \<Rightarrow> bool"
+definition error_free :: "state \<Rightarrow> bool" where
 "error_free s \<equiv> \<not> (\<exists> err. abrupt s = Some (Error err))"
 
 lemma error_free_Norm [simp,intro]: "error_free (Norm s)"
--- a/src/HOL/Bali/Table.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Table.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -37,11 +37,9 @@
 
 section "map of / table of"
 
-syntax
-  table_of      :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"   --{* concrete table *}
-  
 abbreviation
-  "table_of \<equiv> map_of"
+  table_of :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) table"   --{* concrete table *}
+  where "table_of \<equiv> map_of"
 
 translations
   (type)"'a \<rightharpoonup> 'b"       <= (type)"'a \<Rightarrow> 'b Datatype.option"
@@ -53,9 +51,7 @@
 by (simp add: map_add_def)
 
 section {* Conditional Override *}
-constdefs
-cond_override:: 
-  "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table"
+definition cond_override :: "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table" where
 
 --{* when merging tables old and new, only override an entry of table old when  
    the condition cond holds *}
@@ -100,8 +96,7 @@
 
 section {* Filter on Tables *}
 
-constdefs
-filter_tab:: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table"
+definition filter_tab :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table" where
 "filter_tab c t \<equiv> \<lambda> k. (case t k of 
                            None   \<Rightarrow> None
                          | Some x \<Rightarrow> if c k x then Some x else None)"
--- a/src/HOL/Bali/Term.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Term.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -261,9 +261,7 @@
   StatRef :: "ref_ty \<Rightarrow> expr"
   where "StatRef rt == Cast (RefT rt) (Lit Null)"
   
-constdefs
-
-  is_stmt :: "term \<Rightarrow> bool"
+definition is_stmt :: "term \<Rightarrow> bool" where
  "is_stmt t \<equiv> \<exists>c. t=In1r c"
 
 ML {* bind_thms ("is_stmt_rews", sum3_instantiate @{context} @{thm is_stmt_def}) *}
@@ -467,7 +465,7 @@
 "eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
 "eval_binop CondOr  v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
 
-constdefs need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool"
+definition need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool" where
 "need_second_arg binop v1 \<equiv> \<not> ((binop=CondAnd \<and>  \<not> the_Bool v1) \<or>
                                (binop=CondOr  \<and> the_Bool v1))"
 text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument
--- a/src/HOL/Bali/Trans.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Trans.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,7 +9,7 @@
 
 theory Trans imports Evaln begin
 
-constdefs groundVar:: "var \<Rightarrow> bool"
+definition groundVar :: "var \<Rightarrow> bool" where
 "groundVar v \<equiv> (case v of
                    LVar ln \<Rightarrow> True
                  | {accC,statDeclC,stat}e..fn \<Rightarrow> \<exists> a. e=Lit a
@@ -34,7 +34,7 @@
     done
 qed
 
-constdefs groundExprs:: "expr list \<Rightarrow> bool"
+definition groundExprs :: "expr list \<Rightarrow> bool" where
 "groundExprs es \<equiv> list_all (\<lambda> e. \<exists> v. e=Lit v) es"
   
 consts the_val:: "expr \<Rightarrow> val"
--- a/src/HOL/Bali/Type.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/Type.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -41,8 +41,7 @@
 abbreviation Array :: "ty \<Rightarrow> ty"  ("_.[]" [90] 90)
   where "T.[] == RefT (ArrayT T)"
 
-constdefs
-  the_Class :: "ty \<Rightarrow> qtname"
+definition the_Class :: "ty \<Rightarrow> qtname" where
  "the_Class T \<equiv> SOME C. T = Class C" (** primrec not possible here **)
  
 lemma the_Class_eq [simp]: "the_Class (Class C)= C"
--- a/src/HOL/Bali/TypeRel.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/TypeRel.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -65,7 +65,7 @@
 done
 
 lemma subcls1I1:
- "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1 D"
+ "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C1 D"
 apply (auto dest: subcls1I)
 done
 
@@ -126,7 +126,7 @@
 done
 
 lemma single_inheritance: 
-"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C\<^sub>1 B; G\<turnstile>A \<prec>\<^sub>C\<^sub>1 C\<rbrakk> \<Longrightarrow> B = C"
+"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C1 B; G\<turnstile>A \<prec>\<^sub>C1 C\<rbrakk> \<Longrightarrow> B = C"
 by (auto simp add: subcls1_def)
   
 lemma subcls_compareable:
@@ -134,11 +134,11 @@
  \<rbrakk> \<Longrightarrow> G\<turnstile>X \<preceq>\<^sub>C Y \<or> G\<turnstile>Y \<preceq>\<^sub>C X"
 by (rule triangle_lemma)  (auto intro: single_inheritance) 
 
-lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D; ws_prog G \<rbrakk>
+lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C1 D; ws_prog G \<rbrakk>
  \<Longrightarrow> C \<noteq> D"
 proof 
   assume ws: "ws_prog G" and
-    subcls1: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" and
+    subcls1: "G\<turnstile>C \<prec>\<^sub>C1 D" and
      eq_C_D: "C=D"
   from subcls1 obtain c 
     where
@@ -167,7 +167,7 @@
   then show ?thesis
   proof (induct rule: converse_trancl_induct)
     fix C
-    assume subcls1_C_D: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+    assume subcls1_C_D: "G\<turnstile>C \<prec>\<^sub>C1 D"
     then obtain c  where
         "C\<noteq>Object" and
         "class G C = Some c" and
@@ -178,7 +178,7 @@
       by (auto dest: ws_prog_cdeclD)
   next
     fix C Z
-    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C1 Z" and
             subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
            nsubcls_D_Z: "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
     show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
@@ -213,13 +213,13 @@
   then show ?thesis
   proof (induct rule: converse_trancl_induct)
     fix C
-    assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+    assume "G\<turnstile>C \<prec>\<^sub>C1 D"
     with ws 
     show "C\<noteq>D" 
       by (blast dest: subcls1_irrefl)
   next
     fix C Z
-    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+    assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C1 Z" and
             subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
                neq_Z_D: "Z \<noteq> D"
     show "C\<noteq>D"
@@ -298,7 +298,7 @@
   assume       clsC: "class G C = Some c"
   assume subcls_C_C: "G\<turnstile>C \<prec>\<^sub>C D"
   then obtain S where 
-                  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" and
+                  "G\<turnstile>C \<prec>\<^sub>C1 S" and
     subclseq_S_D: "G\<turnstile>S \<preceq>\<^sub>C D"
     by (blast dest: tranclD)
   with clsC 
@@ -341,9 +341,9 @@
 where
   direct:         "G\<turnstile>C\<leadsto>1J     \<spacespace>\<spacespace>     \<Longrightarrow> G\<turnstile>C\<leadsto>J"
 | subint:        "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
-| subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+| subcls1:       "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C1D; G\<turnstile>D\<leadsto>J \<rbrakk>  \<Longrightarrow> G\<turnstile>C\<leadsto>J"
 
-lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C\<^sub>1D \<and> G\<turnstile>D\<leadsto>J)" 
+lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C1D \<and> G\<turnstile>D\<leadsto>J)" 
 apply (erule implmt.induct)
 apply fast+
 done
@@ -568,8 +568,7 @@
 apply (fast dest: widen_Class_Class widen_Class_Iface)
 done
 
-constdefs
-  widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70)
+definition widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70) where
  "G\<turnstile>Ts[\<preceq>]Ts' \<equiv> list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts'"
 
 lemma widens_Nil [simp]: "G\<turnstile>[][\<preceq>][]"
--- a/src/HOL/Bali/TypeSafe.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/TypeSafe.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -95,17 +95,13 @@
 
 section "result conformance"
 
-constdefs
-  assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool"
-          ("_\<le>|_\<preceq>_\<Colon>\<preceq>_"                                        [71,71,71,71] 70)
+definition assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool" ("_\<le>|_\<preceq>_\<Colon>\<preceq>_" [71,71,71,71] 70) where
 "s\<le>|f\<preceq>T\<Colon>\<preceq>E \<equiv>
  (\<forall>s' w. Norm s'\<Colon>\<preceq>E \<longrightarrow> fst E,s'\<turnstile>w\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> assign f w (Norm s')\<Colon>\<preceq>E) \<and>
  (\<forall>s' w. error_free s' \<longrightarrow> (error_free (assign f w s')))"      
 
 
-constdefs
-  rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
-          ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_"                               [71,71,71,71,71,71] 70)
+definition rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool" ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_" [71,71,71,71,71,71] 70) where
   "G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T 
     \<equiv> case T of
         Inl T  \<Rightarrow> if (\<exists> var. t=In2 var)
@@ -330,11 +326,8 @@
 
 declare fun_upd_apply [simp del]
 
-
-constdefs
-  DynT_prop::"[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" 
-                                              ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
- "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and> 
+definition DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70) where
+  "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and> 
                      (if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t)"
 
 lemma DynT_propI: 
--- a/src/HOL/Bali/WellForm.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/WellForm.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -31,8 +31,7 @@
 text  {* well-formed field declaration (common part for classes and interfaces),
         cf. 8.3 and (9.3) *}
 
-constdefs
-  wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
+definition wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool" where
  "wf_fdecl G P \<equiv> \<lambda>(fn,f). is_acc_type G P (type f)"
 
 lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
@@ -55,8 +54,7 @@
 \item the parameter names are unique
 \end{itemize} 
 *}
-constdefs
-  wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
+definition wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool" where
  "wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
                             \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and> 
                             is_acc_type G P (resTy mh) \<and>
@@ -78,7 +76,7 @@
 \end{itemize}
 *}
 
-constdefs callee_lcl:: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv"
+definition callee_lcl :: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv" where
 "callee_lcl C sig m 
  \<equiv> \<lambda> k. (case k of
             EName e 
@@ -88,12 +86,11 @@
                 | Res \<Rightarrow> Some (resTy m))
           | This \<Rightarrow> if is_static m then None else Some (Class C))"
 
-constdefs parameters :: "methd \<Rightarrow> lname set"
+definition parameters :: "methd \<Rightarrow> lname set" where
 "parameters m \<equiv>  set (map (EName \<circ> VNam) (pars m)) 
                   \<union> (if (static m) then {} else {This})"
 
-constdefs
-  wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
+definition wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool" where
  "wf_mdecl G C \<equiv> 
       \<lambda>(sig,m).
           wf_mhead G (pid C) sig (mhead m) \<and> 
@@ -219,8 +216,7 @@
       superinterfaces widens to each of the corresponding result types
 \end{itemize}
 *}
-constdefs
-  wf_idecl :: "prog  \<Rightarrow> idecl \<Rightarrow> bool"
+definition wf_idecl :: "prog  \<Rightarrow> idecl \<Rightarrow> bool" where
  "wf_idecl G \<equiv> 
     \<lambda>(I,i). 
         ws_idecl G I (isuperIfs i) \<and> 
@@ -321,8 +317,7 @@
 \end{itemize}
 *}
 (* to Table *)
-constdefs entails:: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
-                                 ("_ entails _" 20)
+definition entails :: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" ("_ entails _" 20) where
 "t entails P \<equiv> \<forall>k. \<forall> x \<in> t k: P x"
 
 lemma entailsD:
@@ -332,8 +327,7 @@
 lemma empty_entails[simp]: "empty entails P"
 by (simp add: entails_def)
 
-constdefs
- wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
+definition wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool" where
 "wf_cdecl G \<equiv> 
    \<lambda>(C,c).
       \<not>is_iface G C \<and>
@@ -361,8 +355,7 @@
             ))"
 
 (*
-constdefs
- wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
+definition wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool" where
 "wf_cdecl G \<equiv> 
    \<lambda>(C,c).
       \<not>is_iface G C \<and>
@@ -518,8 +511,7 @@
 \item all defined classes are wellformed
 \end{itemize}
 *}
-constdefs
-  wf_prog  :: "prog \<Rightarrow> bool"
+definition wf_prog :: "prog \<Rightarrow> bool" where
  "wf_prog G \<equiv> let is = ifaces G; cs = classes G in
                  ObjectC \<in> set cs \<and> 
                 (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
@@ -738,13 +730,15 @@
  \<Longrightarrow> \<not>is_static im \<and> accmodi im = Public"
 proof -
   assume asm: "wf_prog G" "is_iface G I" "im \<in> imethds G I sig"
+
+  note iface_rec_induct' = iface_rec.induct[of "(%x y z. P x y)", standard]
   have "wf_prog G \<longrightarrow> 
          (\<forall> i im. iface G I = Some i \<longrightarrow> im \<in> imethds G I sig
                   \<longrightarrow> \<not>is_static im \<and> accmodi im = Public)" (is "?P G I")
-  proof (rule iface_rec.induct,intro allI impI)
+  proof (induct G I rule: iface_rec_induct', intro allI impI)
     fix G I i im
-    assume hyp: "\<forall> J i. J \<in> set (isuperIfs i) \<and> ws_prog G \<and> iface G I = Some i
-                 \<longrightarrow> ?P G J"
+    assume hyp: "\<And> i J. iface G I = Some i \<Longrightarrow> ws_prog G \<Longrightarrow> J \<in> set (isuperIfs i)
+                 \<Longrightarrow> ?P G J"
     assume wf: "wf_prog G" and if_I: "iface G I = Some i" and 
            im: "im \<in> imethds G I sig" 
     show "\<not>is_static im \<and> accmodi im = Public" 
@@ -919,7 +913,7 @@
      inheritable: "G \<turnstile>Method old inheritable_in pid C" and
          subclsC: "G\<turnstile>C\<prec>\<^sub>C declclass old"
     from cls_C neq_C_Obj  
-    have super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" 
+    have super: "G\<turnstile>C \<prec>\<^sub>C1 super c" 
       by (rule subcls1I)
     from wf cls_C neq_C_Obj
     have accessible_super: "G\<turnstile>(Class (super c)) accessible_in (pid C)" 
@@ -1353,14 +1347,16 @@
   qed
 qed
 
+lemmas class_rec_induct' = class_rec.induct[of "%x y z w. P x y", standard]
+
 lemma declclass_widen[rule_format]: 
  "wf_prog G 
  \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m 
  \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
-proof (rule class_rec.induct,intro allI impI)
+proof (induct G C rule: class_rec_induct', intro allI impI)
   fix G C c m
-  assume Hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c 
-               \<longrightarrow> ?P G (super c)"
+  assume Hyp: "\<And>c. class G C = Some c \<Longrightarrow> ws_prog G \<Longrightarrow> C \<noteq> Object
+               \<Longrightarrow> ?P G (super c)"
   assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
          m:  "methd G C sig = Some m"
   show "G\<turnstile>C\<preceq>\<^sub>C declclass m" 
@@ -1385,7 +1381,7 @@
       moreover note wf False cls_C  
       ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"  
         by (auto intro: Hyp [rule_format])
-      moreover from cls_C False have  "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
+      moreover from cls_C False have  "G\<turnstile>C \<prec>\<^sub>C1 super c" by (rule subcls1I)
       ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
     next
       case Some
@@ -1539,7 +1535,7 @@
     by (auto intro: method_declared_inI)
   note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *)
   from clsC neq_C_Obj
-  have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+  have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C1 super c"
     by (rule subcls1I)
   then have "G\<turnstile>C \<prec>\<^sub>C super c" ..
   also from old wf is_cls_super
@@ -1609,7 +1605,7 @@
       by (auto dest: ws_prog_cdeclD)
     from clsC wf neq_C_Obj 
     have superAccessible: "G\<turnstile>(Class (super c)) accessible_in (pid C)" and
-         subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+         subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C1 super c"
       by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
               intro: subcls1I)
     show "\<exists>new. ?Constraint C new old"
@@ -1984,27 +1980,6 @@
   qed
 qed
 
-(* Tactical version *)
-(*
-lemma declclassD[rule_format]:
- "wf_prog G \<longrightarrow>  
- (\<forall> c d m. class G C = Some c \<longrightarrow> methd G C sig = Some m \<longrightarrow> 
-  class G (declclass m) = Some d
- \<longrightarrow> table_of (methods d) sig  = Some (mthd m))"
-apply (rule class_rec.induct)
-apply (rule impI)
-apply (rule allI)+
-apply (rule impI)
-apply (case_tac "C=Object")
-apply   (force simp add: methd_rec)
-
-apply   (subst methd_rec)
-apply     (blast dest: wf_ws_prog)+
-apply   (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
-apply     (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
-done
-*)
-
 lemma dynmethd_Object:
   assumes statM: "methd G Object sig = Some statM" and
         private: "accmodi statM = Private" and 
@@ -2363,9 +2338,9 @@
   have "wf_prog G  \<longrightarrow> 
            (\<forall> c m. class G C = Some c \<longrightarrow>  methd G C sig = Some m 
                    \<longrightarrow>  methd G (declclass m) sig = Some m)"      (is "?P G C") 
-  proof (rule class_rec.induct,intro allI impI)
+  proof (induct G C rule: class_rec_induct', intro allI impI)
     fix G C c m
-    assume hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c \<longrightarrow>
+    assume hyp: "\<And>c. class G C = Some c \<Longrightarrow> ws_prog G \<Longrightarrow> C \<noteq> Object \<Longrightarrow>
                      ?P G (super c)"
     assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
             m: "methd G C sig = Some m"
--- a/src/HOL/Bali/WellType.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Bali/WellType.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -53,11 +53,10 @@
   emhead = "ref_ty \<times> mhead"
 
 --{* Some mnemotic selectors for emhead *}
-constdefs 
-  "declrefT" :: "emhead \<Rightarrow> ref_ty"
+definition "declrefT" :: "emhead \<Rightarrow> ref_ty" where
   "declrefT \<equiv> fst"
 
-  "mhd"     :: "emhead \<Rightarrow> mhead"
+definition "mhd"     :: "emhead \<Rightarrow> mhead" where
   "mhd \<equiv> snd"
 
 lemma declrefT_simp[simp]:"declrefT (r,m) = r"
@@ -138,11 +137,10 @@
 done
 
 
-constdefs
-  empty_dt :: "dyn_ty"
+definition empty_dt :: "dyn_ty" where
  "empty_dt \<equiv> \<lambda>a. None"
 
-  invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode"
+definition invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode" where
 "invmode m e \<equiv> if is_static m 
                   then Static 
                   else if e=Super then SuperM else IntVir"
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -30,7 +30,8 @@
   VC_Take of int list * (bool * string list) option |
   VC_Only of string list |
   VC_Without of string list |
-  VC_Examine of string list 
+  VC_Examine of string list |
+  VC_Single of string
 
 fun get_vc thy vc_name =
   (case Boogie_VCs.lookup thy vc_name of
@@ -42,11 +43,37 @@
       | _ => error ("There is no verification condition " ^
           quote vc_name ^ ".")))
 
+local
+  fun split_goal t =
+    (case Boogie_Tactics.split t of
+      [tp] => tp
+    | _ => error "Multiple goals.")
+
+  fun single_prep t =
+    let
+      val (us, u) = split_goal t
+      val assms = [((@{binding vc_trace}, []), map (rpair []) us)]
+    in
+      pair [u] o snd o ProofContext.add_assms_i Assumption.assume_export assms
+    end
+
+  fun single_prove goal ctxt thm =
+    Goal.prove ctxt [] [] goal (fn {context, ...} => HEADGOAL (
+      Boogie_Tactics.split_tac
+      THEN' Boogie_Tactics.drop_assert_at_tac
+      THEN' SUBPROOF (fn _ => Tactic.rtac thm 1) context))
+in
 fun boogie_vc (vc_name, vc_opts) lthy =
   let
     val thy = ProofContext.theory_of lthy
     val vc = get_vc thy vc_name
 
+    fun extract vc l =
+      (case Boogie_VCs.extract vc l of
+        SOME vc' => vc'
+      | NONE => error ("There is no assertion to be proved with label " ^
+          quote l ^ "."))
+
     val vcs =
       (case vc_opts of
         VC_Complete => [vc]
@@ -55,18 +82,29 @@
       | VC_Take (ps, SOME (false, ls)) => [Boogie_VCs.paths_without ps ls vc]
       | VC_Only ls => [Boogie_VCs.only ls vc]
       | VC_Without ls => [Boogie_VCs.without ls vc]
-      | VC_Examine ls => map_filter (Boogie_VCs.extract vc) ls)
+      | VC_Examine ls => map (extract vc) ls
+      | VC_Single l => [extract vc l])
     val ts = map Boogie_VCs.prop_of_vc vcs
 
+    val (prepare, finish) =
+      (case vc_opts of
+         VC_Single _ => (single_prep (hd ts), single_prove (hd ts))
+      | _ => (pair ts, K I))
+
     val discharge = fold (Boogie_VCs.discharge o pair vc_name)
     fun after_qed [thms] = Local_Theory.theory (discharge (vcs ~~ thms))
       | after_qed _ = I
   in
     lthy
     |> fold Variable.auto_fixes ts
-    |> Proof.theorem_i NONE after_qed [map (rpair []) ts]
+    |> (fn lthy1 => lthy1
+    |> prepare
+    |-> (fn us => fn lthy2 => lthy2
+    |> Proof.theorem_i NONE (fn thmss => fn ctxt =>
+         let val export = map (finish lthy1) o ProofContext.export ctxt lthy2
+         in after_qed (map export thmss) ctxt end) [map (rpair []) us]))
   end
-
+end
 
 fun write_list head =
   map Pretty.str o sort (dict_ord string_ord o pairself explode) #>
@@ -239,7 +277,8 @@
 
 val vc_name = OuterParse.name
 
-val vc_labels = Scan.repeat1 OuterParse.name
+val vc_label = OuterParse.name
+val vc_labels = Scan.repeat1 vc_label
 
 val vc_paths =
   OuterParse.nat -- (Args.$$$ "-" |-- OuterParse.nat) >> (op upto) ||
@@ -247,13 +286,14 @@
 
 val vc_opts =
   scan_arg
-   (scan_val "examine" vc_labels >> VC_Examine ||
+   (scan_val "assertion" vc_label >> VC_Single ||
+    scan_val "examine" vc_labels >> VC_Examine ||
     scan_val "take" ((OuterParse.list vc_paths >> flat) -- Scan.option (
       scan_val "without" vc_labels >> pair false ||
       scan_val "and_also" vc_labels >> pair true) >> VC_Take) ||
     scan_val "only" vc_labels >> VC_Only ||
     scan_val "without" vc_labels >> VC_Without) ||
-  Scan.succeed VC_Complete  
+  Scan.succeed VC_Complete
 
 val _ =
   OuterSyntax.command "boogie_vc"
--- a/src/HOL/Boogie/Tools/boogie_loader.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -49,6 +49,15 @@
   if line = 0 orelse col = 0 then no_label_name
   else label_prefix ^ string_of_int line ^ "_" ^ string_of_int col
 
+fun mk_syntax name i =
+  let
+    val syn = Syntax.escape name
+    val args = space_implode ",/ " (replicate i "_")
+  in
+    if i = 0 then Mixfix (syn, [], 1000)
+    else Mixfix (syn ^ "()'(/" ^ args ^ "')", replicate i 0, 1000)
+  end
+
 
 datatype attribute_value = StringValue of string | TermValue of term
 
@@ -78,8 +87,8 @@
       | NONE =>
           let
             val args = Name.variant_list [] (replicate arity "'")
-            val (T, thy') =
-              ObjectLogic.typedecl (Binding.name isa_name, args, NoSyn) thy
+            val (T, thy') = ObjectLogic.typedecl (Binding.name isa_name, args,
+              mk_syntax name arity) thy
             val type_name = fst (Term.dest_Type T)
           in (((name, type_name), log_new name type_name), thy') end)
     end
@@ -93,24 +102,6 @@
 
 
 local
-  val quote_mixfix =
-    Symbol.explode #> map
-      (fn "_" => "'_"
-        | "(" => "'("
-        | ")" => "')"
-        | "/" => "'/"
-        | s => s) #>
-    implode
-
-  fun mk_syntax name i =
-    let
-      val syn = quote_mixfix name
-      val args = concat (separate ",/ " (replicate i "_"))
-    in
-      if i = 0 then Mixfix (syn, [], 1000)
-      else Mixfix (syn ^ "()'(/" ^ args ^ "')", replicate i 0, 1000)
-    end
-
   fun maybe_builtin T =
     let
       fun const name = SOME (Const (name, T))
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -9,6 +9,9 @@
   val unfold_labels_tac: Proof.context -> int -> tactic
   val boogie_tac: Proof.context -> thm list -> int -> tactic
   val boogie_all_tac: Proof.context -> thm list -> tactic
+  val split: term -> (term list * term) list
+  val split_tac: int -> tactic
+  val drop_assert_at_tac: int -> tactic
   val setup: theory -> theory
 end
 
@@ -47,25 +50,38 @@
 
 
 local
-  fun prep_tac facts =
-    Method.insert_tac facts
-    THEN' REPEAT_ALL_NEW
-      (Tactic.resolve_tac [@{thm impI}, @{thm conjI}, @{thm TrueI}]
-       ORELSE' Tactic.etac @{thm conjE})
+  fun explode_conj (@{term "op &"} $ t $ u) = explode_conj t @ explode_conj u
+    | explode_conj t = [t] 
 
-  val drop_assert_at = Conv.concl_conv ~1 (Conv.try_conv (Conv.arg_conv
-    (Conv.rewr_conv assert_at_def)))
+  fun splt (ts, @{term "op -->"} $ t $ u) = splt (ts @ explode_conj t, u)
+    | splt (ts, @{term "op &"} $ t $ u) = splt (ts, t) @ splt (ts, u)
+    | splt (ts, @{term assert_at} $ _ $ t) = [(ts, t)]
+    | splt (_, @{term True}) = []
+    | splt tp = [tp]
+in
+fun split t =
+  splt ([], HOLogic.dest_Trueprop t)
+  |> map (fn (us, u) => (map HOLogic.mk_Trueprop us, HOLogic.mk_Trueprop u))
+end
 
+val split_tac = REPEAT_ALL_NEW (
+  Tactic.resolve_tac [@{thm impI}, @{thm conjI}, @{thm TrueI}]
+  ORELSE' Tactic.etac @{thm conjE})
+
+val drop_assert_at_tac = CONVERSION (Conv.concl_conv ~1 (Conv.try_conv (
+  Conv.arg_conv (Conv.rewr_conv assert_at_def))))
+
+local
   fun case_name_of t =
     (case HOLogic.dest_Trueprop (Logic.strip_imp_concl t) of
       @{term assert_at} $ Free (n, _) $ _ => n
     | _ => raise TERM ("case_name_of", [t]))
 
   fun boogie_cases ctxt = METHOD_CASES (fn facts =>
-    ALLGOALS (prep_tac facts) #>
+    ALLGOALS (Method.insert_tac facts THEN' split_tac) #>
     Seq.maps (fn st =>
       st
-      |> ALLGOALS (CONVERSION drop_assert_at)
+      |> ALLGOALS drop_assert_at_tac
       |> Seq.map (pair (map (rpair [] o case_name_of) (Thm.prems_of st)))) #>
     Seq.maps (fn (names, st) =>
       CASES
--- a/src/HOL/Code_Evaluation.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Code_Evaluation.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -119,6 +119,44 @@
 end
 *}
 
+setup {*
+let
+  fun mk_term_of_eq thy ty vs tyco abs ty_rep proj =
+    let
+      val arg = Var (("x", 0), ty);
+      val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $
+        (HOLogic.mk_term_of ty_rep (Const (proj, ty --> ty_rep) $ arg))
+        |> Thm.cterm_of thy;
+      val cty = Thm.ctyp_of thy ty;
+    in
+      @{thm term_of_anything}
+      |> Drule.instantiate' [SOME cty] [SOME (Thm.cterm_of thy arg), SOME rhs]
+      |> Thm.varifyT
+    end;
+  fun add_term_of_code tyco raw_vs abs raw_ty_rep proj thy =
+    let
+      val algebra = Sign.classes_of thy;
+      val vs = map (fn (v, sort) =>
+        (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
+      val ty = Type (tyco, map TFree vs);
+      val ty_rep = map_atyps
+        (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_ty_rep;
+      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
+      val eq = mk_term_of_eq thy ty vs tyco abs ty_rep proj;
+   in
+      thy
+      |> Code.del_eqns const
+      |> Code.add_eqn eq
+    end;
+  fun ensure_term_of_code (tyco, (raw_vs, ((abs, ty), (proj, _)))) thy =
+    let
+      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
+    in if has_inst then add_term_of_code tyco raw_vs abs ty proj thy else thy end;
+in
+  Code.abstype_interpretation ensure_term_of_code
+end
+*}
+
 
 subsubsection {* Code generator setup *}
 
--- a/src/HOL/Decision_Procs/Cooper.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Decision_Procs/Cooper.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -293,10 +293,10 @@
 by (induct p, simp_all)
 
 
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
   "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
   (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
   "evaldjf f ps \<equiv> foldr (djf f) ps F"
 
 lemma djf_Or: "Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)"
@@ -340,7 +340,7 @@
   thus ?thesis by (simp only: list_all_iff)
 qed
 
-constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "DJ f p \<equiv> evaldjf f (disjuncts p)"
 
 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
@@ -395,7 +395,7 @@
   "lex_ns ([], ms) = True"
   "lex_ns (ns, []) = False"
   "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
-constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
+definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
   "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
 
 consts
@@ -455,10 +455,10 @@
 lemma nummul_nb: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul i t)"
 by (induct t rule: nummul.induct, auto simp add: numadd_nb)
 
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
   "numneg t \<equiv> nummul (- 1) t"
 
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
   "numsub s t \<equiv> (if s = t then C 0 else numadd (s, numneg t))"
 
 lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)"
@@ -505,7 +505,7 @@
 lemma not_bn: "bound0 p \<Longrightarrow> bound0 (not p)"
 by (cases p, auto)
 
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else And p q)"
 lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)"
 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
@@ -515,7 +515,7 @@
 lemma conj_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
 using conj_def by auto 
 
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p else Or p q)"
 
 lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)"
@@ -525,7 +525,7 @@
 lemma disj_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
 using disj_def by auto 
 
-constdefs   imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "imp p q \<equiv> (if (p = F \<or> q=T) then T else if p=T then q else if q=F then not p else Imp p q)"
 lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)"
 by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) (simp_all add: not)
@@ -534,7 +534,7 @@
 lemma imp_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
 using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) simp_all
 
-constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else 
        if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
   Iff p q)"
@@ -1749,7 +1749,7 @@
   shows "(\<exists> x. Ifm bbs (x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (i#bs)) ` set (\<beta> p). Ifm bbs ((b+j)#bs) p))"
   using cp_thm[OF lp up dd dp,where i="i"] by auto
 
-constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
+definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
   "unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q;
              B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
@@ -1814,7 +1814,7 @@
 qed
     (* Cooper's Algorithm *)
 
-constdefs cooper :: "fm \<Rightarrow> fm"
+definition cooper :: "fm \<Rightarrow> fm" where
   "cooper p \<equiv> 
   (let (q,B,d) = unit p; js = iupt 1 d;
        mq = simpfm (minusinf q);
--- a/src/HOL/Decision_Procs/Ferrack.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -169,26 +169,26 @@
 lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
 by (cases p) auto
 
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else 
    if p = q then p else And p q)"
 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
 
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
        else if p=q then p else Or p q)"
 
 lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
 
-constdefs  imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
     else Imp p q)"
 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
 by (cases "p=F \<or> q=T",simp_all add: imp_def) 
 
-constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> NOT p = q) then F else 
        if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
   Iff p q)"
@@ -369,10 +369,10 @@
 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
 by (induct p, simp_all)
 
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
   "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
   (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
   "evaldjf f ps \<equiv> foldr (djf f) ps F"
 
 lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
@@ -423,7 +423,7 @@
   thus ?thesis by (simp only: list_all_iff)
 qed
 
-constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "DJ f p \<equiv> evaldjf f (disjuncts p)"
 
 lemma DJ: assumes fdj: "\<forall> p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))"
@@ -653,10 +653,10 @@
 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
 by (induct t rule: nummul.induct, auto )
 
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
   "numneg t \<equiv> nummul t (- 1)"
 
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
   "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
 
 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
@@ -724,7 +724,7 @@
   from maxcoeff_nz[OF nz th] show ?thesis .
 qed
 
-constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
+definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
   "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
    (let t' = simpnum t ; g = numgcd t' in 
       if g > 1 then (let g' = gcd n g in 
@@ -1779,7 +1779,7 @@
 
 
     (* Implement the right hand side of Ferrante and Rackoff's Theorem. *)
-constdefs ferrack:: "fm \<Rightarrow> fm"
+definition ferrack :: "fm \<Rightarrow> fm" where
   "ferrack p \<equiv> (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p'
                 in if (mp = T \<or> pp = T) then T else 
                    (let U = remdps(map simp_num_pair 
--- a/src/HOL/Decision_Procs/MIR.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Decision_Procs/MIR.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -566,7 +566,7 @@
   thus ?thesis by (simp only: list_all_iff)
 qed
 
-constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "DJ f p \<equiv> evaldjf f (disjuncts p)"
 
 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
@@ -623,7 +623,7 @@
   "lex_ns ([], ms) = True"
   "lex_ns (ns, []) = False"
   "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
-constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
+definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
   "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
 
 consts 
@@ -873,10 +873,10 @@
 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
 by (induct t rule: nummul.induct, auto)
 
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
   "numneg t \<equiv> nummul t (- 1)"
 
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
   "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
 
 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
@@ -1038,7 +1038,7 @@
   from maxcoeff_nz[OF nz th] show ?thesis .
 qed
 
-constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
+definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
   "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
    (let t' = simpnum t ; g = numgcd t' in 
       if g > 1 then (let g' = gcd n g in 
@@ -1137,7 +1137,7 @@
 lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
 by (induct p, auto)
 
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else 
    if p = q then p else And p q)"
 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
@@ -1148,7 +1148,7 @@
 lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
 using conj_def by auto 
 
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
        else if p=q then p else Or p q)"
 
@@ -1159,7 +1159,7 @@
 lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
 using disj_def by auto 
 
-constdefs   imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
     else Imp p q)"
 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
@@ -1169,7 +1169,7 @@
 lemma imp_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
 using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def) 
 
-constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else 
        if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
   Iff p q)"
@@ -1216,7 +1216,7 @@
   thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gnz gd] real_of_int_div[OF gnz gc] by simp
 qed
 
-constdefs simpdvd:: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)"
+definition simpdvd :: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" where
   "simpdvd d t \<equiv> 
    (let g = numgcd t in 
       if g > 1 then (let g' = gcd d g in 
@@ -1479,7 +1479,7 @@
 
   (* Generic quantifier elimination *)
 
-constdefs list_conj :: "fm list \<Rightarrow> fm"
+definition list_conj :: "fm list \<Rightarrow> fm" where
   "list_conj ps \<equiv> foldr conj ps T"
 lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)"
   by (induct ps, auto simp add: list_conj_def)
@@ -1487,7 +1487,7 @@
   by (induct ps, auto simp add: list_conj_def)
 lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
   by (induct ps, auto simp add: list_conj_def)
-constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs
                    in conj (decr (list_conj yes)) (f (list_conj no)))"
 
@@ -2954,7 +2954,7 @@
                                             else (NDvd (i*k) (CN 0 c (Mul k e))))"
   "a\<rho> p = (\<lambda> k. p)"
 
-constdefs \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
+definition \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
   "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>\<rho> p (t,k))"
 
 lemma \<sigma>\<rho>:
@@ -3517,7 +3517,7 @@
   "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
   "isrlfm p = (isatom p \<and> (bound0 p))"
 
-constdefs fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm"
+definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
   "fp p n s j \<equiv> (if n > 0 then 
             (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j)))))
                         (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
@@ -3836,7 +3836,7 @@
 
     (* Linearize a formula where Bound 0 ranges over [0,1) *)
 
-constdefs rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm"
+definition rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm" where
   "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
 
 lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))"
@@ -5103,7 +5103,7 @@
 
     (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
 
-constdefs ferrack01:: "fm \<Rightarrow> fm"
+definition ferrack01 :: "fm \<Rightarrow> fm" where
   "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
                     U = remdups(map simp_num_pair 
                      (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
@@ -5350,7 +5350,7 @@
   shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (real i#bs)) ` set (\<beta> p). Ifm ((b+real j)#bs) p))"
   using cp_thm[OF lp up dd dp] by auto
 
-constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
+definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
   "unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q;
              B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
@@ -5417,7 +5417,7 @@
 qed
     (* Cooper's Algorithm *)
 
-constdefs cooper :: "fm \<Rightarrow> fm"
+definition cooper :: "fm \<Rightarrow> fm" where
   "cooper p \<equiv> 
   (let (q,B,d) = unit p; js = iupt (1,d);
        mq = simpfm (minusinf q);
@@ -5561,7 +5561,7 @@
   shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
   using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp 
 
-constdefs chooset:: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int"
+definition chooset :: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int" where
   "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
              B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ; 
              a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>\<rho> q))
@@ -5621,7 +5621,7 @@
   ultimately show ?thes by blast
 qed
 
-constdefs stage:: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm"
+definition stage :: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" where
   "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) (iupt (1,c*d)))"
 lemma stage:
   shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))"
@@ -5641,7 +5641,7 @@
   from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp
 qed
 
-constdefs redlove:: "fm \<Rightarrow> fm"
+definition redlove :: "fm \<Rightarrow> fm" where
   "redlove p \<equiv> 
   (let (q,B,d) = chooset p;
        mq = simpfm (minusinf q);
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -273,10 +273,10 @@
   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
   shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
 
-constdefs tmneg :: "tm \<Rightarrow> tm"
+definition tmneg :: "tm \<Rightarrow> tm" where
   "tmneg t \<equiv> tmmul t (C (- 1,1))"
 
-constdefs tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
+definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm" where
   "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))"
 
 lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
@@ -477,26 +477,26 @@
 lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
 by (induct p rule: not.induct) auto
 
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else 
    if p = q then p else And p q)"
 lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)"
 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
 
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
        else if p=q then p else Or p q)"
 
 lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)"
 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
 
-constdefs  imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
     else Imp p q)"
 lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)"
 by (cases "p=F \<or> q=T",simp_all add: imp_def) 
 
-constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
   "iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> NOT p = q) then F else 
        if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
   Iff p q)"
@@ -776,10 +776,10 @@
 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
 by (induct p, simp_all)
 
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
   "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
   (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
   "evaldjf f ps \<equiv> foldr (djf f) ps F"
 
 lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
@@ -823,7 +823,7 @@
   thus ?thesis by (simp only: list_all_iff)
 qed
 
-constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "DJ f p \<equiv> evaldjf f (disjuncts p)"
 
 lemma DJ: assumes fdj: "\<forall> p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))"
@@ -869,10 +869,10 @@
   "conjuncts T = []"
   "conjuncts p = [p]"
 
-constdefs list_conj :: "fm list \<Rightarrow> fm"
+definition list_conj :: "fm list \<Rightarrow> fm" where
   "list_conj ps \<equiv> foldr conj ps T"
 
-constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
   "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
                    in conj (decr0 (list_conj yes)) (f (list_conj no)))"
 
@@ -1158,7 +1158,7 @@
   "conjs p = [p]"
 lemma conjs_ci: "(\<forall> q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
 by (induct p rule: conjs.induct, auto)
-constdefs list_disj :: "fm list \<Rightarrow> fm"
+definition list_disj :: "fm list \<Rightarrow> fm" where
   "list_disj ps \<equiv> foldr disj ps F"
 
 lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -188,12 +188,12 @@
 | "poly_cmul y (CN c n p) = CN (poly_cmul y c) n (poly_cmul y p)"
 | "poly_cmul y p = C y *\<^sub>p p"
 
-constdefs monic:: "poly \<Rightarrow> (poly \<times> bool)"
+definition monic :: "poly \<Rightarrow> (poly \<times> bool)" where
   "monic p \<equiv> (let h = headconst p in if h = 0\<^sub>N then (p,False) else ((C (Ninv h)) *\<^sub>p p, 0>\<^sub>N h))"
 
 subsection{* Pseudo-division *}
 
-constdefs shift1:: "poly \<Rightarrow> poly"
+definition shift1 :: "poly \<Rightarrow> poly" where
   "shift1 p \<equiv> CN 0\<^sub>p 0 p"
 consts funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
 
@@ -212,7 +212,7 @@
   by pat_completeness auto
 
 
-constdefs polydivide:: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)"
+definition polydivide :: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)" where
   "polydivide s p \<equiv> polydivide_aux (head p,degree p,p,0, s)"
 
 fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly" where
@@ -262,7 +262,7 @@
 lemma isnpolyh_mono: "\<lbrakk>n' \<le> n ; isnpolyh p n\<rbrakk> \<Longrightarrow> isnpolyh p n'"
 by (induct p rule: isnpolyh.induct, auto)
 
-constdefs isnpoly:: "poly \<Rightarrow> bool"
+definition isnpoly :: "poly \<Rightarrow> bool" where
   "isnpoly p \<equiv> isnpolyh p 0"
 
 text{* polyadd preserves normal forms *}
--- a/src/HOL/Divides.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Divides.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -309,6 +309,15 @@
   apply (simp add: no_zero_divisors)
   done
 
+lemma div_mult_swap:
+  assumes "c dvd b"
+  shows "a * (b div c) = (a * b) div c"
+proof -
+  from assms have "b div c * (a div 1) = b * a div (c * 1)"
+    by (simp only: div_mult_div_if_dvd one_dvd)
+  then show ?thesis by (simp add: mult_commute)
+qed
+   
 lemma div_mult_mult2 [simp]:
   "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b"
   by (drule div_mult_mult1) (simp add: mult_commute)
@@ -347,6 +356,24 @@
 apply(simp add: div_mult_div_if_dvd dvd_power_same)
 done
 
+lemma dvd_div_eq_mult:
+  assumes "a \<noteq> 0" and "a dvd b"  
+  shows "b div a = c \<longleftrightarrow> b = c * a"
+proof
+  assume "b = c * a"
+  then show "b div a = c" by (simp add: assms)
+next
+  assume "b div a = c"
+  then have "b div a * a = c * a" by simp
+  moreover from `a dvd b` have "b div a * a = b" by (simp add: dvd_div_mult_self)
+  ultimately show "b = c * a" by simp
+qed
+   
+lemma dvd_div_div_eq_mult:
+  assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d"
+  shows "b div a = d div c \<longleftrightarrow> b * c = a * d"
+  using assms by (auto simp add: mult_commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym)
+
 end
 
 class ring_div = semiring_div + idom
--- a/src/HOL/Finite_Set.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Finite_Set.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -2528,8 +2528,7 @@
   fold1Set_insertI [intro]:
    "\<lbrakk> fold_graph f a A x; a \<notin> A \<rbrakk> \<Longrightarrow> fold1Set f (insert a A) x"
 
-constdefs
-  fold1 :: "('a => 'a => 'a) => 'a set => 'a"
+definition fold1 :: "('a => 'a => 'a) => 'a set => 'a" where
   "fold1 f A == THE x. fold1Set f A x"
 
 lemma fold1Set_nonempty:
--- a/src/HOL/Fun.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Fun.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -119,8 +119,9 @@
 
 subsection {* Injectivity and Surjectivity *}
 
-constdefs
-  inj_on :: "['a => 'b, 'a set] => bool"  -- "injective"
+definition
+  inj_on :: "['a => 'b, 'a set] => bool" where
+  -- "injective"
   "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y"
 
 text{*A common special case: functions injective over the entire domain type.*}
@@ -132,11 +133,14 @@
   bij_betw :: "('a => 'b) => 'a set => 'b set => bool" where -- "bijective"
   [code del]: "bij_betw f A B \<longleftrightarrow> inj_on f A & f ` A = B"
 
-constdefs
-  surj :: "('a => 'b) => bool"                   (*surjective*)
+definition
+  surj :: "('a => 'b) => bool" where
+  -- "surjective"
   "surj f == ! y. ? x. y=f(x)"
 
-  bij :: "('a => 'b) => bool"                    (*bijective*)
+definition
+  bij :: "('a => 'b) => bool" where
+  -- "bijective"
   "bij f == inj f & surj f"
 
 lemma injI:
@@ -377,8 +381,8 @@
 
 subsection{*Function Updating*}
 
-constdefs
-  fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)"
+definition
+  fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
   "fun_upd f a b == % x. if x=a then b else f x"
 
 nonterminals
--- a/src/HOL/GCD.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/GCD.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -412,6 +412,33 @@
 apply (rule gcd_mult_cancel_nat [transferred], auto)
 done
 
+lemma coprime_crossproduct_nat:
+  fixes a b c d :: nat
+  assumes "coprime a d" and "coprime b c"
+  shows "a * c = b * d \<longleftrightarrow> a = b \<and> c = d" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs then show ?lhs by simp
+next
+  assume ?lhs
+  from `?lhs` have "a dvd b * d" by (auto intro: dvdI dest: sym)
+  with `coprime a d` have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat)
+  from `?lhs` have "b dvd a * c" by (auto intro: dvdI dest: sym)
+  with `coprime b c` have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat)
+  from `?lhs` have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult_commute)
+  with `coprime b c` have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
+  from `?lhs` have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult_commute)
+  with `coprime a d` have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
+  from `a dvd b` `b dvd a` have "a = b" by (rule Nat.dvd.antisym)
+  moreover from `c dvd d` `d dvd c` have "c = d" by (rule Nat.dvd.antisym)
+  ultimately show ?rhs ..
+qed
+
+lemma coprime_crossproduct_int:
+  fixes a b c d :: int
+  assumes "coprime a d" and "coprime b c"
+  shows "\<bar>a\<bar> * \<bar>c\<bar> = \<bar>b\<bar> * \<bar>d\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>b\<bar> \<and> \<bar>c\<bar> = \<bar>d\<bar>"
+  using assms by (intro coprime_crossproduct_nat [transferred]) auto
+
 text {* \medskip Addition laws *}
 
 lemma gcd_add1_nat [simp]: "gcd ((m::nat) + n) n = gcd m n"
--- a/src/HOL/Groebner_Basis.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Groebner_Basis.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -510,7 +510,7 @@
     let
       val z = instantiate_cterm ([(zT,T)],[]) zr
       val eq = instantiate_cterm ([(eqT,T)],[]) geq
-      val th = Simplifier.rewrite (ss addsimps simp_thms)
+      val th = Simplifier.rewrite (ss addsimps @{thms simp_thms})
            (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"}
                   (Thm.capply (Thm.capply eq t) z)))
     in equal_elim (symmetric th) TrueI
@@ -640,7 +640,7 @@
 
 val comp_conv = (Simplifier.rewrite
 (HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"}
-              addsimps ths addsimps simp_thms
+              addsimps ths addsimps @{thms simp_thms}
               addsimprocs Numeral_Simprocs.field_cancel_numeral_factors
                addsimprocs [add_frac_frac_simproc, add_frac_num_simproc,
                             ord_frac_simproc]
--- a/src/HOL/Groups.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Groups.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1216,7 +1216,7 @@
         @{const_name Groups.uminus}, @{const_name Groups.minus}]
   | agrp_ord _ = ~1;
 
-fun termless_agrp (a, b) = (TermOrd.term_lpo agrp_ord (a, b) = LESS);
+fun termless_agrp (a, b) = (Term_Ord.term_lpo agrp_ord (a, b) = LESS);
 
 local
   val ac1 = mk_meta_eq @{thm add_assoc};
@@ -1250,7 +1250,7 @@
 val eqI_rules = [@{thm less_eqI}, @{thm le_eqI}, @{thm eq_eqI}];
 
 val dest_eqI = 
-  fst o HOLogic.dest_bin "op =" HOLogic.boolT o HOLogic.dest_Trueprop o concl_of;
+  fst o HOLogic.dest_bin @{const_name "op ="} HOLogic.boolT o HOLogic.dest_Trueprop o concl_of;
 
 );
 *}
--- a/src/HOL/HOL.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/HOL.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -845,12 +845,16 @@
 
 setup {*
 let
-  (*prevent substitution on bool*)
-  fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso
-    Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("bool", []) | _ => false)
-      (nth (Thm.prems_of thm) (i - 1)) then Hypsubst.hyp_subst_tac i thm else no_tac thm;
+  fun non_bool_eq (@{const_name "op ="}, Type (_, [T, _])) = T <> @{typ bool}
+    | non_bool_eq _ = false;
+  val hyp_subst_tac' =
+    SUBGOAL (fn (goal, i) =>
+      if Term.exists_Const non_bool_eq goal
+      then Hypsubst.hyp_subst_tac i
+      else no_tac);
 in
   Hypsubst.hypsubst_setup
+  (*prevent substitution on bool*)
   #> Context_Rules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)
 end
 *}
@@ -1118,8 +1122,7 @@
   its premise.
 *}
 
-constdefs
-  simp_implies :: "[prop, prop] => prop"  (infixr "=simp=>" 1)
+definition simp_implies :: "[prop, prop] => prop"  (infixr "=simp=>" 1) where
   [code del]: "simp_implies \<equiv> op ==>"
 
 lemma simp_impliesI:
@@ -1392,13 +1395,23 @@
 )
 *}
 
-constdefs
-  induct_forall where "induct_forall P == \<forall>x. P x"
-  induct_implies where "induct_implies A B == A \<longrightarrow> B"
-  induct_equal where "induct_equal x y == x = y"
-  induct_conj where "induct_conj A B == A \<and> B"
-  induct_true where "induct_true == True"
-  induct_false where "induct_false == False"
+definition induct_forall where
+  "induct_forall P == \<forall>x. P x"
+
+definition induct_implies where
+  "induct_implies A B == A \<longrightarrow> B"
+
+definition induct_equal where
+  "induct_equal x y == x = y"
+
+definition induct_conj where
+  "induct_conj A B == A \<and> B"
+
+definition induct_true where
+  "induct_true == True"
+
+definition induct_false where
+  "induct_false == False"
 
 lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))"
   by (unfold atomize_all induct_forall_def)
@@ -1721,8 +1734,8 @@
 
 fun eq_codegen thy defs dep thyname b t gr =
     (case strip_comb t of
-       (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE
-     | (Const ("op =", _), [t, u]) =>
+       (Const (@{const_name "op ="}, Type (_, [Type ("fun", _), _])), _) => NONE
+     | (Const (@{const_name "op ="}, _), [t, u]) =>
           let
             val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr;
             val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr';
@@ -1731,7 +1744,7 @@
             SOME (Codegen.parens
               (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''')
           end
-     | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen
+     | (t as Const (@{const_name "op ="}, _), ts) => SOME (Codegen.invoke_codegen
          thy defs dep thyname b (Codegen.eta_expand t ts 2) gr)
      | _ => NONE);
 
@@ -2050,7 +2063,7 @@
 
 (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *)
 local
-  fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t
+  fun wrong_prem (Const (@{const_name All}, _) $ Abs (_, _, t)) = wrong_prem t
     | wrong_prem (Bound _) = true
     | wrong_prem _ = false;
   val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of);
--- a/src/HOL/Hilbert_Choice.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hilbert_Choice.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -307,8 +307,8 @@
 
 subsection {* Least value operator *}
 
-constdefs
-  LeastM :: "['a => 'b::ord, 'a => bool] => 'a"
+definition
+  LeastM :: "['a => 'b::ord, 'a => bool] => 'a" where
   "LeastM m P == SOME x. P x & (\<forall>y. P y --> m x <= m y)"
 
 syntax
@@ -360,11 +360,12 @@
 
 subsection {* Greatest value operator *}
 
-constdefs
-  GreatestM :: "['a => 'b::ord, 'a => bool] => 'a"
+definition
+  GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" where
   "GreatestM m P == SOME x. P x & (\<forall>y. P y --> m y <= m x)"
 
-  Greatest :: "('a::ord => bool) => 'a"    (binder "GREATEST " 10)
+definition
+  Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) where
   "Greatest == GreatestM (%x. x)"
 
 syntax
--- a/src/HOL/Hoare/Arith2.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Arith2.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Hoare/Arith2.thy
-    ID:         $Id$
     Author:     Norbert Galm
     Copyright   1995 TUM
 
@@ -10,11 +9,10 @@
 imports Main
 begin
 
-constdefs
-  "cd"    :: "[nat, nat, nat] => bool"
+definition "cd" :: "[nat, nat, nat] => bool" where
   "cd x m n  == x dvd m & x dvd n"
 
-  gcd     :: "[nat, nat] => nat"
+definition gcd     :: "[nat, nat] => nat" where
   "gcd m n     == @x.(cd x m n) & (!y.(cd y m n) --> y<=x)"
 
 consts fac     :: "nat => nat"
--- a/src/HOL/Hoare/Heap.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Heap.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Hoare/Heap.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2002 TUM
 
@@ -19,19 +18,17 @@
 lemma not_Ref_eq [iff]: "(ALL y. x ~= Ref y) = (x = Null)"
   by (induct x) auto
 
-consts addr :: "'a ref \<Rightarrow> 'a"
-primrec "addr(Ref a) = a"
+primrec addr :: "'a ref \<Rightarrow> 'a" where
+  "addr (Ref a) = a"
 
 
 section "The heap"
 
 subsection "Paths in the heap"
 
-consts
- Path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
-primrec
-"Path h x [] y = (x = y)"
-"Path h x (a#as) y = (x = Ref a \<and> Path h (h a) as y)"
+primrec Path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool" where
+  "Path h x [] y \<longleftrightarrow> x = y"
+| "Path h x (a#as) y \<longleftrightarrow> x = Ref a \<and> Path h (h a) as y"
 
 lemma [iff]: "Path h Null xs y = (xs = [] \<and> y = Null)"
 apply(case_tac xs)
@@ -60,8 +57,7 @@
 
 subsection "Non-repeating paths"
 
-constdefs
-  distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
+definition distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool" where
   "distPath h x as y   \<equiv>   Path h x as y  \<and>  distinct as"
 
 text{* The term @{term"distPath h x as y"} expresses the fact that a
@@ -86,8 +82,7 @@
 
 subsubsection "Relational abstraction"
 
-constdefs
- List :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> bool"
+definition List :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> bool" where
 "List h x as == Path h x as Null"
 
 lemma [simp]: "List h x [] = (x = Null)"
@@ -138,10 +133,10 @@
 
 subsection "Functional abstraction"
 
-constdefs
- islist :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> bool"
+definition islist :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> bool" where
 "islist h p == \<exists>as. List h p as"
- list :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list"
+
+definition list :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list" where
 "list h p == SOME as. List h p as"
 
 lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"
--- a/src/HOL/Hoare/Hoare_Logic.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Hoare_Logic.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -40,7 +40,7 @@
                                       (s ~: b --> Sem c2 s s'))"
 "Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
 
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" where
   "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
 
 
--- a/src/HOL/Hoare/Hoare_Logic_Abort.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -42,7 +42,7 @@
 "Sem(While b x c) s s' =
  (if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
 
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" where
   "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
 
 
@@ -256,8 +256,8 @@
 (* Special syntax for guarded statements and guarded array updates: *)
 
 syntax
-  guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
-  array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
+  "_guarded_com" :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
+  "_array_update" :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
 translations
   "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
   "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
--- a/src/HOL/Hoare/Pointer_Examples.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Pointer_Examples.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -216,19 +216,17 @@
 
 text"This is still a bit rough, especially the proof."
 
-constdefs
- cor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+definition cor :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "cor P Q == if P then True else Q"
- cand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+
+definition cand :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
 "cand P Q == if P then Q else False"
 
-consts merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
-
-recdef merge "measure(%(xs,ys,f). size xs + size ys)"
+fun merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list" where
 "merge(x#xs,y#ys,f) = (if f x y then x # merge(xs,y#ys,f)
-                                else y # merge(x#xs,ys,f))"
-"merge(x#xs,[],f) = x # merge(xs,[],f)"
-"merge([],y#ys,f) = y # merge([],ys,f)"
+                                else y # merge(x#xs,ys,f))" |
+"merge(x#xs,[],f) = x # merge(xs,[],f)" |
+"merge([],y#ys,f) = y # merge([],ys,f)" |
 "merge([],[],f) = []"
 
 text{* Simplifies the proof a little: *}
@@ -481,7 +479,7 @@
 
 subsection "Storage allocation"
 
-constdefs new :: "'a set \<Rightarrow> 'a"
+definition new :: "'a set \<Rightarrow> 'a" where
 "new A == SOME a. a \<notin> A"
 
 
--- a/src/HOL/Hoare/Pointers0.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Pointers0.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -73,8 +73,7 @@
 
 subsubsection "Relational abstraction"
 
-constdefs
- List :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> bool"
+definition List :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> bool" where
 "List h x as == Path h x as Null"
 
 lemma [simp]: "List h x [] = (x = Null)"
@@ -122,10 +121,10 @@
 
 subsection "Functional abstraction"
 
-constdefs
- islist :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool"
+definition islist :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where
 "islist h p == \<exists>as. List h p as"
- list :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list"
+
+definition list :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list" where
 "list h p == SOME as. List h p as"
 
 lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"
@@ -321,13 +320,11 @@
 
 text"This is still a bit rough, especially the proof."
 
-consts merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
-
-recdef merge "measure(%(xs,ys,f). size xs + size ys)"
+fun merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list" where
 "merge(x#xs,y#ys,f) = (if f x y then x # merge(xs,y#ys,f)
-                                else y # merge(x#xs,ys,f))"
-"merge(x#xs,[],f) = x # merge(xs,[],f)"
-"merge([],y#ys,f) = y # merge([],ys,f)"
+                                else y # merge(x#xs,ys,f))" |
+"merge(x#xs,[],f) = x # merge(xs,[],f)" |
+"merge([],y#ys,f) = y # merge([],ys,f)" |
 "merge([],[],f) = []"
 
 lemma imp_disjCL: "(P|Q \<longrightarrow> R) = ((P \<longrightarrow> R) \<and> (~P \<longrightarrow> Q \<longrightarrow> R))"
@@ -407,7 +404,7 @@
 
 subsection "Storage allocation"
 
-constdefs new :: "'a set \<Rightarrow> 'a::ref"
+definition new :: "'a set \<Rightarrow> 'a::ref" where
 "new A == SOME a. a \<notin> A & a \<noteq> Null"
 
 
--- a/src/HOL/Hoare/SepLogHeap.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/SepLogHeap.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -41,8 +41,7 @@
 
 subsection "Lists on the heap"
 
-constdefs
- List :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
+definition List :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool" where
 "List h x as == Path h x as 0"
 
 lemma [simp]: "List h x [] = (x = 0)"
--- a/src/HOL/Hoare/Separation.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare/Separation.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -16,20 +16,19 @@
 
 text{* The semantic definition of a few connectives: *}
 
-constdefs
- ortho:: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55)
+definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55) where
 "h1 \<bottom> h2 == dom h1 \<inter> dom h2 = {}"
 
- is_empty :: "heap \<Rightarrow> bool"
+definition is_empty :: "heap \<Rightarrow> bool" where
 "is_empty h == h = empty"
 
- singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+definition singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
 "singl h x y == dom h = {x} & h x = Some y"
 
- star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
+definition star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" where
 "star P Q == \<lambda>h. \<exists>h1 h2. h = h1++h2 \<and> h1 \<bottom> h2 \<and> P h1 \<and> Q h2"
 
- wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
+definition wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" where
 "wand P Q == \<lambda>h. \<forall>h'. h' \<bottom> h \<and> P h' \<longrightarrow> Q(h++h')"
 
 text{*This is what assertions look like without any syntactic sugar: *}
--- a/src/HOL/Hoare_Parallel/Gar_Coll.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -31,8 +31,7 @@
 under which the selected edge @{text "R"} and node @{text "T"} are
 valid: *}
 
-constdefs
-  Mut_init :: "gar_coll_state \<Rightarrow> bool"
+definition Mut_init :: "gar_coll_state \<Rightarrow> bool" where
   "Mut_init \<equiv> \<guillemotleft> T \<in> Reach \<acute>E \<and> R < length \<acute>E \<and> T < length \<acute>M \<guillemotright>"
 
 text {* \noindent For the mutator we
@@ -40,14 +39,13 @@
 @{text "\<acute>z"} is set to false if the mutator has already redirected an
 edge but has not yet colored the new target.   *}
 
-constdefs
-  Redirect_Edge :: "gar_coll_state ann_com"
+definition Redirect_Edge :: "gar_coll_state ann_com" where
   "Redirect_Edge \<equiv> .{\<acute>Mut_init \<and> \<acute>z}. \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
 
-  Color_Target :: "gar_coll_state ann_com"
+definition Color_Target :: "gar_coll_state ann_com" where
   "Color_Target \<equiv> .{\<acute>Mut_init \<and> \<not>\<acute>z}. \<langle>\<acute>M:=\<acute>M[T:=Black],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
 
-  Mutator :: "gar_coll_state ann_com"
+definition Mutator :: "gar_coll_state ann_com" where
   "Mutator \<equiv>
   .{\<acute>Mut_init \<and> \<acute>z}. 
   WHILE True INV .{\<acute>Mut_init \<and> \<acute>z}. 
@@ -88,22 +86,20 @@
 
 consts  M_init :: nodes
 
-constdefs
-  Proper_M_init :: "gar_coll_state \<Rightarrow> bool"
+definition Proper_M_init :: "gar_coll_state \<Rightarrow> bool" where
   "Proper_M_init \<equiv>  \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
  
-  Proper :: "gar_coll_state \<Rightarrow> bool"
+definition Proper :: "gar_coll_state \<Rightarrow> bool" where
   "Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>"
 
-  Safe :: "gar_coll_state \<Rightarrow> bool"
+definition Safe :: "gar_coll_state \<Rightarrow> bool" where
   "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
 
 lemmas collector_defs = Proper_M_init_def Proper_def Safe_def
 
 subsubsection {* Blackening the roots *}
 
-constdefs
-  Blacken_Roots :: " gar_coll_state ann_com"
+definition Blacken_Roots :: " gar_coll_state ann_com" where
   "Blacken_Roots \<equiv> 
   .{\<acute>Proper}.
   \<acute>ind:=0;;
@@ -133,13 +129,11 @@
 
 subsubsection {* Propagating black *}
 
-constdefs
-  PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or>
    (\<not>\<acute>z \<and> i=R \<and> (snd(\<acute>E!R)) = T \<and> (\<exists>r. ind \<le> r \<and> r < length \<acute>E \<and> BtoW(\<acute>E!r,\<acute>M))))\<guillemotright>"
 
-constdefs  
-  Propagate_Black_aux :: "gar_coll_state ann_com"
+definition Propagate_Black_aux :: "gar_coll_state ann_com" where
   "Propagate_Black_aux \<equiv>
   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
   \<acute>ind:=0;;
@@ -214,14 +208,12 @@
 
 subsubsection {* Refining propagating black *}
 
-constdefs
-  Auxk :: "gar_coll_state \<Rightarrow> bool"
+definition Auxk :: "gar_coll_state \<Rightarrow> bool" where
   "Auxk \<equiv> \<guillemotleft>\<acute>k<length \<acute>M \<and> (\<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> 
           \<acute>obc<Blacks \<acute>M \<or> (\<not>\<acute>z \<and> \<acute>ind=R \<and> snd(\<acute>E!R)=T  
           \<and> (\<exists>r. \<acute>ind<r \<and> r<length \<acute>E \<and> BtoW(\<acute>E!r, \<acute>M))))\<guillemotright>"
 
-constdefs  
-  Propagate_Black :: " gar_coll_state ann_com"
+definition Propagate_Black :: " gar_coll_state ann_com" where
   "Propagate_Black \<equiv>
   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
   \<acute>ind:=0;;
@@ -328,12 +320,10 @@
 
 subsubsection {* Counting black nodes *}
 
-constdefs
-  CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
 
-constdefs
-  Count :: " gar_coll_state ann_com"
+definition Count :: " gar_coll_state ann_com" where
   "Count \<equiv>
   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
     \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
@@ -398,12 +388,10 @@
   Append_to_free2: "i \<notin> Reach e 
      \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
 
-constdefs
-  AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "AppendInv \<equiv> \<guillemotleft>\<lambda>ind. \<forall>i<length \<acute>M. ind\<le>i \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black\<guillemotright>"
 
-constdefs
-  Append :: " gar_coll_state ann_com"
+definition Append :: " gar_coll_state ann_com" where
    "Append \<equiv>
   .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
   \<acute>ind:=0;;
@@ -444,8 +432,7 @@
 
 subsubsection {* Correctness of the Collector *}
 
-constdefs 
-  Collector :: " gar_coll_state ann_com"
+definition Collector :: " gar_coll_state ann_com" where
   "Collector \<equiv>
 .{\<acute>Proper}.  
  WHILE True INV .{\<acute>Proper}. 
--- a/src/HOL/Hoare_Parallel/Graph.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -13,20 +13,19 @@
 
 consts Roots :: "nat set"
 
-constdefs
-  Proper_Roots :: "nodes \<Rightarrow> bool"
+definition Proper_Roots :: "nodes \<Rightarrow> bool" where
   "Proper_Roots M \<equiv> Roots\<noteq>{} \<and> Roots \<subseteq> {i. i<length M}"
 
-  Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool"
+definition Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool" where
   "Proper_Edges \<equiv> (\<lambda>(M,E). \<forall>i<length E. fst(E!i)<length M \<and> snd(E!i)<length M)"
 
-  BtoW :: "(edge \<times> nodes) \<Rightarrow> bool"
+definition BtoW :: "(edge \<times> nodes) \<Rightarrow> bool" where
   "BtoW \<equiv> (\<lambda>(e,M). (M!fst e)=Black \<and> (M!snd e)\<noteq>Black)"
 
-  Blacks :: "nodes \<Rightarrow> nat set"
+definition Blacks :: "nodes \<Rightarrow> nat set" where
   "Blacks M \<equiv> {i. i<length M \<and> M!i=Black}"
 
-  Reach :: "edges \<Rightarrow> nat set"
+definition Reach :: "edges \<Rightarrow> nat set" where
   "Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
               \<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))
               \<or> x\<in>Roots}"
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -26,24 +26,23 @@
 
 subsection {* The Mutators *}
 
-constdefs 
-  Mul_mut_init :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_mut_init :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "Mul_mut_init \<equiv> \<guillemotleft> \<lambda>n. n=length \<acute>Muts \<and> (\<forall>i<n. R (\<acute>Muts!i)<length \<acute>E 
                           \<and> T (\<acute>Muts!i)<length \<acute>M) \<guillemotright>"
 
-  Mul_Redirect_Edge  :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Redirect_Edge  :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
   "Mul_Redirect_Edge j n \<equiv>
   .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.
   \<langle>IF T(\<acute>Muts!j) \<in> Reach \<acute>E THEN  
   \<acute>E:= \<acute>E[R (\<acute>Muts!j):= (fst (\<acute>E!R(\<acute>Muts!j)), T (\<acute>Muts!j))] FI,, 
   \<acute>Muts:= \<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=False\<rparr>]\<rangle>"
 
-  Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
   "Mul_Color_Target j n \<equiv>
   .{\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)}. 
   \<langle>\<acute>M:=\<acute>M[T (\<acute>Muts!j):=Black],, \<acute>Muts:=\<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=True\<rparr>]\<rangle>"
 
-  Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Mutator j n \<equiv>
   .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.  
   WHILE True  
@@ -156,28 +155,25 @@
 
 subsection {* The Collector *}
 
-constdefs
-  Queue :: "mul_gar_coll_state \<Rightarrow> nat"
+definition Queue :: "mul_gar_coll_state \<Rightarrow> nat" where
  "Queue \<equiv> \<guillemotleft> length (filter (\<lambda>i. \<not> Z i \<and> \<acute>M!(T i) \<noteq> Black) \<acute>Muts) \<guillemotright>"
 
 consts  M_init :: nodes
 
-constdefs
-  Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool" where
   "Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
 
-  Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "Mul_Proper \<equiv> \<guillemotleft> \<lambda>n. Proper_Roots \<acute>M \<and> Proper_Edges (\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<and> n=length \<acute>Muts \<guillemotright>"
 
-  Safe :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Safe :: "mul_gar_coll_state \<Rightarrow> bool" where
   "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
 
 lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def
 
 subsubsection {* Blackening Roots *}
 
-constdefs
-  Mul_Blacken_Roots :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Blacken_Roots :: "nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Blacken_Roots n \<equiv>
   .{\<acute>Mul_Proper n}.
   \<acute>ind:=0;;
@@ -208,16 +204,14 @@
 
 subsubsection {* Propagating Black *} 
 
-constdefs
-  Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool" where
   "Mul_PBInv \<equiv>  \<guillemotleft>\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
                  \<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) \<and> \<acute>l\<le>\<acute>Queue\<guillemotright>"
 
-  Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool" where
   "Mul_Auxk \<equiv> \<guillemotleft>\<acute>l<\<acute>Queue \<or> \<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> \<acute>obc\<subset>Blacks \<acute>M\<guillemotright>"
 
-constdefs
-  Mul_Propagate_Black :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Propagate_Black :: "nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Propagate_Black n \<equiv>
  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
   \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)}. 
@@ -296,11 +290,10 @@
 
 subsubsection {* Counting Black Nodes *}
 
-constdefs
-  Mul_CountInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
- "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
+definition Mul_CountInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
+  "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
 
-  Mul_Count :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Count :: "nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Count n \<equiv> 
   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
     \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
@@ -396,11 +389,10 @@
   Append_to_free2: "i \<notin> Reach e 
            \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
 
-constdefs
-  Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
   "Mul_AppendInv \<equiv> \<guillemotleft> \<lambda>ind. (\<forall>i. ind\<le>i \<longrightarrow> i<length \<acute>M \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black)\<guillemotright>"
 
-  Mul_Append :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Append :: "nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Append n \<equiv> 
   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
   \<acute>ind:=0;;
@@ -438,8 +430,7 @@
 
 subsubsection {* Collector *}
 
-constdefs 
-  Mul_Collector :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+definition Mul_Collector :: "nat \<Rightarrow>  mul_gar_coll_state ann_com" where
   "Mul_Collector n \<equiv>
 .{\<acute>Mul_Proper n}.  
 WHILE True INV .{\<acute>Mul_Proper n}. 
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -27,12 +27,12 @@
 consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
 primrec "post (c, q) = q"
 
-constdefs  interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool"
+definition interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool" where
   "interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or>  
                     (\<forall>(r,a) \<in> atomics (the co'). \<parallel>= (q \<inter> r) a q \<and>
                     (co = None \<or> (\<forall>p \<in> assertions (the co). \<parallel>= (p \<inter> r) a p)))"
 
-constdefs interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" 
+definition interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" where 
   "interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow> 
                          interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
 
--- a/src/HOL/Hoare_Parallel/OG_Tactics.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -171,8 +171,7 @@
   "\<parallel>- (q \<inter> (r \<inter> b)) a q \<Longrightarrow> interfree_aux (None, q, Some (AnnAwait r b a))"
 by(simp add: interfree_aux_def oghoare_sound)
 
-constdefs 
-  interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \<Rightarrow> bool"
+definition interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \<Rightarrow> bool" where
   "interfree_swap == \<lambda>(x, xs). \<forall>y\<in>set xs. interfree_aux (com x, post x, com y)
   \<and> interfree_aux(com y, post y, com x)"
 
@@ -208,7 +207,7 @@
   \<Longrightarrow> interfree (map (\<lambda>k. (c k, Q k)) [a..<b])"
 by(force simp add: interfree_def less_diff_conv)
 
-constdefs map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45)
+definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45) where
   "[\<turnstile>] Ts == (\<forall>i<length Ts. \<exists>c q. Ts!i=(Some c, q) \<and> \<turnstile> c q)"
 
 lemma MapAnnEmpty: "[\<turnstile>] []"
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,14 +7,13 @@
   'a ann_com_op = "('a ann_com) option"
   'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
   
-consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
-primrec "com (c, q) = c"
+primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where
+  "com (c, q) = c"
 
-consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
-primrec "post (c, q) = q"
+primrec post :: "'a ann_triple_op \<Rightarrow> 'a assn" where
+  "post (c, q) = q"
 
-constdefs
-  All_None :: "'a ann_triple_op list \<Rightarrow> bool"
+definition All_None :: "'a ann_triple_op list \<Rightarrow> bool" where
   "All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
 
 subsection {* The Transition Relation *}
@@ -88,26 +87,24 @@
 
 subsection {* Definition of Semantics *}
 
-constdefs
-  ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set"
+definition ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set" where
   "ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
 
-  ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+definition ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
   "ann_SEM c S \<equiv> \<Union>ann_sem c ` S"  
 
-  sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set"
+definition sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set" where
   "sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
 
-  SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+definition SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
   "SEM c S \<equiv> \<Union>sem c ` S "
 
 abbreviation Omega :: "'a com"    ("\<Omega>" 63)
   where "\<Omega> \<equiv> While UNIV UNIV (Basic id)"
 
-consts fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com"
-primrec 
-   "fwhile b c 0 = \<Omega>"
-   "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"
+primrec fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com" where
+    "fwhile b c 0 = \<Omega>"
+  | "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"
 
 subsubsection {* Proofs *}
 
@@ -299,11 +296,10 @@
 
 section {* Validity of Correctness Formulas *}
 
-constdefs 
-  com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(3\<parallel>= _// _//_)" [90,55,90] 50)
+definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50) where
   "\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
 
-  ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool"   ("\<Turnstile> _ _" [60,90] 45)
+definition ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45) where
   "\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
 
 end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,8 +7,7 @@
 declare Un_subset_iff [simp del] le_sup_iff [simp del]
 declare Cons_eq_map_conv [iff]
 
-constdefs
-  stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"  
+definition stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where  
   "stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)" 
 
 inductive
@@ -39,16 +38,19 @@
              \<turnstile> P sat [pre', rely', guar', post'] \<rbrakk>
             \<Longrightarrow> \<turnstile> P sat [pre, rely, guar, post]"
 
-constdefs 
-  Pre :: "'a rgformula \<Rightarrow> 'a set"
+definition Pre :: "'a rgformula \<Rightarrow> 'a set" where
   "Pre x \<equiv> fst(snd x)"
-  Post :: "'a rgformula \<Rightarrow> 'a set"
+
+definition Post :: "'a rgformula \<Rightarrow> 'a set" where
   "Post x \<equiv> snd(snd(snd(snd x)))"
-  Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+
+definition Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set" where
   "Rely x \<equiv> fst(snd(snd x))"
-  Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+
+definition Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set" where
   "Guar x \<equiv> fst(snd(snd(snd x)))"
-  Com :: "'a rgformula \<Rightarrow> 'a com"
+
+definition Com :: "'a rgformula \<Rightarrow> 'a com" where
   "Com x \<equiv> fst x"
 
 subsection {* Proof System for Parallel Programs *}
@@ -1093,8 +1095,7 @@
 
 subsection {* Soundness of the System for Parallel Programs *}
 
-constdefs
-  ParallelCom :: "('a rgformula) list \<Rightarrow> 'a par_com"
+definition ParallelCom :: "('a rgformula) list \<Rightarrow> 'a par_com" where
   "ParallelCom Ps \<equiv> map (Some \<circ> fst) Ps" 
 
 lemma two: 
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -81,8 +81,7 @@
 | CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
 | CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
 
-constdefs
-  cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set"
+definition cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set" where
   "cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"  
 
 subsubsection {* Parallel computations *}
@@ -95,14 +94,12 @@
 | ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
 | ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
 
-constdefs
-  par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set"
+definition par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set" where
   "par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"  
 
 subsection{* Modular Definition of Computation *}
 
-constdefs 
-  lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf"
+definition lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf" where
   "lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"
 
 inductive_set cptn_mod :: "('a confs) set"
@@ -380,38 +377,36 @@
 
 types 'a rgformula = "'a com \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
 
-constdefs
-  assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set"
+definition assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set" where
   "assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
                c!i -e\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
 
-  comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set"
+definition comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set" where
   "comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow> 
                c!i -c\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
                (fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"
 
-  com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" 
-                 ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+definition com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" 
+                 ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45) where
   "\<Turnstile> P sat [pre, rely, guar, post] \<equiv> 
    \<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"
 
 subsection {* Validity for Parallel Programs. *}
 
-constdefs
-  All_None :: "(('a com) option) list \<Rightarrow> bool"
+definition All_None :: "(('a com) option) list \<Rightarrow> bool" where
   "All_None xs \<equiv> \<forall>c\<in>set xs. c=None"
 
-  par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set"
+definition par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set" where
   "par_assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
              c!i -pe\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
 
-  par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set"
+definition par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set" where
   "par_comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>   
         c!i -pc\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
          (All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"
 
-  par_com_validity :: "'a  par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set 
-\<Rightarrow> 'a set \<Rightarrow> bool"  ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+definition par_com_validity :: "'a  par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set 
+\<Rightarrow> 'a set \<Rightarrow> bool"  ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where
   "\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv> 
    \<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"
 
@@ -419,23 +414,22 @@
 
 subsubsection {* Definition of the conjoin operator *}
 
-constdefs
-  same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
   "same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"
  
-  same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
   "same_state c clist \<equiv> (\<forall>i <length clist. \<forall>j<length c. snd(c!j) = snd((clist!i)!j))"
 
-  same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
   "same_program c clist \<equiv> (\<forall>j<length c. fst(c!j) = map (\<lambda>x. fst(nth x j)) clist)"
 
-  compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
   "compat_label c clist \<equiv> (\<forall>j. Suc j<length c \<longrightarrow> 
          (c!j -pc\<rightarrow> c!Suc j \<and> (\<exists>i<length clist. (clist!i)!j -c\<rightarrow> (clist!i)! Suc j \<and> 
                        (\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or> 
          (c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"
 
-  conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"  ("_ \<propto> _" [65,65] 64)
+definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"  ("_ \<propto> _" [65,65] 64) where
   "c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"
 
 subsubsection {* Some previous lemmas *}
--- a/src/HOL/IOA/Solve.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/IOA/Solve.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -10,9 +10,7 @@
 imports IOA
 begin
 
-constdefs
-
-  is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool"
+definition is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool" where
   "is_weak_pmap f C A ==
    (!s:starts_of(C). f(s):starts_of(A)) &
    (!s t a. reachable C s &
--- a/src/HOL/Import/HOL/HOL4Base.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL/HOL4Base.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -4,22 +4,19 @@
 
 ;setup_theory bool
 
-constdefs
-  ARB :: "'a" 
+definition ARB :: "'a" where 
   "ARB == SOME x::'a::type. True"
 
 lemma ARB_DEF: "ARB = (SOME x::'a::type. True)"
   by (import bool ARB_DEF)
 
-constdefs
-  IN :: "'a => ('a => bool) => bool" 
+definition IN :: "'a => ('a => bool) => bool" where 
   "IN == %(x::'a::type) f::'a::type => bool. f x"
 
 lemma IN_DEF: "IN = (%(x::'a::type) f::'a::type => bool. f x)"
   by (import bool IN_DEF)
 
-constdefs
-  RES_FORALL :: "('a => bool) => ('a => bool) => bool" 
+definition RES_FORALL :: "('a => bool) => ('a => bool) => bool" where 
   "RES_FORALL ==
 %(p::'a::type => bool) m::'a::type => bool. ALL x::'a::type. IN x p --> m x"
 
@@ -28,8 +25,7 @@
     ALL x::'a::type. IN x p --> m x)"
   by (import bool RES_FORALL_DEF)
 
-constdefs
-  RES_EXISTS :: "('a => bool) => ('a => bool) => bool" 
+definition RES_EXISTS :: "('a => bool) => ('a => bool) => bool" where 
   "RES_EXISTS ==
 %(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x"
 
@@ -37,8 +33,7 @@
 (%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x)"
   by (import bool RES_EXISTS_DEF)
 
-constdefs
-  RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" 
+definition RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" where 
   "RES_EXISTS_UNIQUE ==
 %(p::'a::type => bool) m::'a::type => bool.
    RES_EXISTS p m &
@@ -52,8 +47,7 @@
      (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y)))"
   by (import bool RES_EXISTS_UNIQUE_DEF)
 
-constdefs
-  RES_SELECT :: "('a => bool) => ('a => bool) => 'a" 
+definition RES_SELECT :: "('a => bool) => ('a => bool) => 'a" where 
   "RES_SELECT ==
 %(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x"
 
@@ -264,15 +258,13 @@
 
 ;setup_theory combin
 
-constdefs
-  K :: "'a => 'b => 'a" 
+definition K :: "'a => 'b => 'a" where 
   "K == %(x::'a::type) y::'b::type. x"
 
 lemma K_DEF: "K = (%(x::'a::type) y::'b::type. x)"
   by (import combin K_DEF)
 
-constdefs
-  S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" 
+definition S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" where 
   "S ==
 %(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type)
    x::'a::type. f x (g x)"
@@ -282,8 +274,7 @@
     x::'a::type. f x (g x))"
   by (import combin S_DEF)
 
-constdefs
-  I :: "'a => 'a" 
+definition I :: "'a => 'a" where 
   "(op ==::('a::type => 'a::type) => ('a::type => 'a::type) => prop)
  (I::'a::type => 'a::type)
  ((S::('a::type => ('a::type => 'a::type) => 'a::type)
@@ -299,16 +290,14 @@
    (K::'a::type => 'a::type => 'a::type))"
   by (import combin I_DEF)
 
-constdefs
-  C :: "('a => 'b => 'c) => 'b => 'a => 'c" 
+definition C :: "('a => 'b => 'c) => 'b => 'a => 'c" where 
   "C == %(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x"
 
 lemma C_DEF: "C =
 (%(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x)"
   by (import combin C_DEF)
 
-constdefs
-  W :: "('a => 'a => 'b) => 'a => 'b" 
+definition W :: "('a => 'a => 'b) => 'a => 'b" where 
   "W == %(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x"
 
 lemma W_DEF: "W = (%(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x)"
@@ -582,8 +571,7 @@
 
 ;setup_theory relation
 
-constdefs
-  TC :: "('a => 'a => bool) => 'a => 'a => bool" 
+definition TC :: "('a => 'a => bool) => 'a => 'a => bool" where 
   "TC ==
 %(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
    ALL P::'a::type => 'a::type => bool.
@@ -601,8 +589,7 @@
        P a b)"
   by (import relation TC_DEF)
 
-constdefs
-  RTC :: "('a => 'a => bool) => 'a => 'a => bool" 
+definition RTC :: "('a => 'a => bool) => 'a => 'a => bool" where 
   "RTC ==
 %(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
    ALL P::'a::type => 'a::type => bool.
@@ -644,8 +631,7 @@
    (ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z)"
   by (import relation transitive_def)
 
-constdefs
-  pred_reflexive :: "('a => 'a => bool) => bool" 
+definition pred_reflexive :: "('a => 'a => bool) => bool" where 
   "pred_reflexive == %R::'a::type => 'a::type => bool. ALL x::'a::type. R x x"
 
 lemma reflexive_def: "ALL R::'a::type => 'a::type => bool.
@@ -788,8 +774,7 @@
    (ALL (x::'a::type) y::'a::type. RTC R x y --> RTC Q x y)"
   by (import relation RTC_MONOTONE)
 
-constdefs
-  WF :: "('a => 'a => bool) => bool" 
+definition WF :: "('a => 'a => bool) => bool" where 
   "WF ==
 %R::'a::type => 'a::type => bool.
    ALL B::'a::type => bool.
@@ -814,8 +799,7 @@
    WF x --> x xa xb --> xa ~= xb"
   by (import relation WF_NOT_REFL)
 
-constdefs
-  EMPTY_REL :: "'a => 'a => bool" 
+definition EMPTY_REL :: "'a => 'a => bool" where 
   "EMPTY_REL == %(x::'a::type) y::'a::type. False"
 
 lemma EMPTY_REL_DEF: "ALL (x::'a::type) y::'a::type. EMPTY_REL x y = False"
@@ -847,8 +831,7 @@
    WF R --> WF (relation.inv_image R f)"
   by (import relation WF_inv_image)
 
-constdefs
-  RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" 
+definition RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" where 
   "RESTRICT ==
 %(f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) (x::'a::type)
    y::'a::type. if R y x then f y else ARB"
@@ -891,8 +874,7 @@
    the_fun R M x = Eps (approx R M x)"
   by (import relation the_fun_def)
 
-constdefs
-  WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" 
+definition WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" where 
   "WFREC ==
 %(R::'a::type => 'a::type => bool)
    (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
@@ -1052,8 +1034,7 @@
    split xb x = split f' xa"
   by (import pair pair_case_cong)
 
-constdefs
-  LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
+definition LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where 
   "LEX ==
 %(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
    (s::'a::type, t::'b::type) (u::'a::type, v::'b::type).
@@ -1069,8 +1050,7 @@
    WF x & WF xa --> WF (LEX x xa)"
   by (import pair WF_LEX)
 
-constdefs
-  RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
+definition RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where 
   "RPROD ==
 %(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
    (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v"
@@ -1113,8 +1093,7 @@
 lemma NOT_LESS_EQ: "ALL (m::nat) n::nat. m = n --> ~ m < n"
   by (import prim_rec NOT_LESS_EQ)
 
-constdefs
-  SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" 
+definition SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" where 
   "(op ==::((nat => 'a::type)
          => 'a::type => ('a::type => 'a::type) => nat => bool)
         => ((nat => 'a::type)
@@ -1187,8 +1166,7 @@
    (ALL m::nat. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))"
   by (import prim_rec SIMP_REC_THM)
 
-constdefs
-  PRE :: "nat => nat" 
+definition PRE :: "nat => nat" where 
   "PRE == %m::nat. if m = 0 then 0 else SOME n::nat. m = Suc n"
 
 lemma PRE_DEF: "ALL m::nat. PRE m = (if m = 0 then 0 else SOME n::nat. m = Suc n)"
@@ -1197,8 +1175,7 @@
 lemma PRE: "PRE 0 = 0 & (ALL m::nat. PRE (Suc m) = m)"
   by (import prim_rec PRE)
 
-constdefs
-  PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" 
+definition PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" where 
   "PRIM_REC_FUN ==
 %(x::'a::type) f::'a::type => nat => 'a::type.
    SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)"
@@ -1214,8 +1191,7 @@
        PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)"
   by (import prim_rec PRIM_REC_EQN)
 
-constdefs
-  PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" 
+definition PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" where 
   "PRIM_REC ==
 %(x::'a::type) (f::'a::type => nat => 'a::type) m::nat.
    PRIM_REC_FUN x f m (PRE m)"
@@ -1286,8 +1262,7 @@
 
 ;setup_theory arithmetic
 
-constdefs
-  nat_elim__magic :: "nat => nat" 
+definition nat_elim__magic :: "nat => nat" where 
   "nat_elim__magic == %n::nat. n"
 
 lemma nat_elim__magic: "ALL n::nat. nat_elim__magic n = n"
@@ -1746,22 +1721,19 @@
 
 ;setup_theory hrat
 
-constdefs
-  trat_1 :: "nat * nat" 
+definition trat_1 :: "nat * nat" where 
   "trat_1 == (0, 0)"
 
 lemma trat_1: "trat_1 = (0, 0)"
   by (import hrat trat_1)
 
-constdefs
-  trat_inv :: "nat * nat => nat * nat" 
+definition trat_inv :: "nat * nat => nat * nat" where 
   "trat_inv == %(x::nat, y::nat). (y, x)"
 
 lemma trat_inv: "ALL (x::nat) y::nat. trat_inv (x, y) = (y, x)"
   by (import hrat trat_inv)
 
-constdefs
-  trat_add :: "nat * nat => nat * nat => nat * nat" 
+definition trat_add :: "nat * nat => nat * nat => nat * nat" where 
   "trat_add ==
 %(x::nat, y::nat) (x'::nat, y'::nat).
    (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
@@ -1771,8 +1743,7 @@
    (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
   by (import hrat trat_add)
 
-constdefs
-  trat_mul :: "nat * nat => nat * nat => nat * nat" 
+definition trat_mul :: "nat * nat => nat * nat => nat * nat" where 
   "trat_mul ==
 %(x::nat, y::nat) (x'::nat, y'::nat).
    (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
@@ -1788,8 +1759,7 @@
 (ALL n::nat. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)"
   by (import hrat trat_sucint)
 
-constdefs
-  trat_eq :: "nat * nat => nat * nat => bool" 
+definition trat_eq :: "nat * nat => nat * nat => bool" where 
   "trat_eq ==
 %(x::nat, y::nat) (x'::nat, y'::nat). Suc x * Suc y' = Suc x' * Suc y"
 
@@ -1901,23 +1871,20 @@
     (EX x::nat * nat. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))"
   by (import hrat hrat_tybij)
 
-constdefs
-  hrat_1 :: "hrat" 
+definition hrat_1 :: "hrat" where 
   "hrat_1 == mk_hrat (trat_eq trat_1)"
 
 lemma hrat_1: "hrat_1 = mk_hrat (trat_eq trat_1)"
   by (import hrat hrat_1)
 
-constdefs
-  hrat_inv :: "hrat => hrat" 
+definition hrat_inv :: "hrat => hrat" where 
   "hrat_inv == %T1::hrat. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
 
 lemma hrat_inv: "ALL T1::hrat.
    hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
   by (import hrat hrat_inv)
 
-constdefs
-  hrat_add :: "hrat => hrat => hrat" 
+definition hrat_add :: "hrat => hrat => hrat" where 
   "hrat_add ==
 %(T1::hrat) T2::hrat.
    mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
@@ -1927,8 +1894,7 @@
    mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
   by (import hrat hrat_add)
 
-constdefs
-  hrat_mul :: "hrat => hrat => hrat" 
+definition hrat_mul :: "hrat => hrat => hrat" where 
   "hrat_mul ==
 %(T1::hrat) T2::hrat.
    mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
@@ -1938,8 +1904,7 @@
    mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
   by (import hrat hrat_mul)
 
-constdefs
-  hrat_sucint :: "nat => hrat" 
+definition hrat_sucint :: "nat => hrat" where 
   "hrat_sucint == %T1::nat. mk_hrat (trat_eq (trat_sucint T1))"
 
 lemma hrat_sucint: "ALL T1::nat. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))"
@@ -1987,8 +1952,7 @@
 
 ;setup_theory hreal
 
-constdefs
-  hrat_lt :: "hrat => hrat => bool" 
+definition hrat_lt :: "hrat => hrat => bool" where 
   "hrat_lt == %(x::hrat) y::hrat. EX d::hrat. y = hrat_add x d"
 
 lemma hrat_lt: "ALL (x::hrat) y::hrat. hrat_lt x y = (EX d::hrat. y = hrat_add x d)"
@@ -2096,8 +2060,7 @@
    hrat_lt x y --> (EX xa::hrat. hrat_lt x xa & hrat_lt xa y)"
   by (import hreal HRAT_MEAN)
 
-constdefs
-  isacut :: "(hrat => bool) => bool" 
+definition isacut :: "(hrat => bool) => bool" where 
   "isacut ==
 %C::hrat => bool.
    Ex C &
@@ -2113,8 +2076,7 @@
     (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y)))"
   by (import hreal isacut)
 
-constdefs
-  cut_of_hrat :: "hrat => hrat => bool" 
+definition cut_of_hrat :: "hrat => hrat => bool" where 
   "cut_of_hrat == %(x::hrat) y::hrat. hrat_lt y x"
 
 lemma cut_of_hrat: "ALL x::hrat. cut_of_hrat x = (%y::hrat. hrat_lt y x)"
@@ -2173,15 +2135,13 @@
    (EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))"
   by (import hreal CUT_NEARTOP_MUL)
 
-constdefs
-  hreal_1 :: "hreal" 
+definition hreal_1 :: "hreal" where 
   "hreal_1 == hreal (cut_of_hrat hrat_1)"
 
 lemma hreal_1: "hreal_1 = hreal (cut_of_hrat hrat_1)"
   by (import hreal hreal_1)
 
-constdefs
-  hreal_add :: "hreal => hreal => hreal" 
+definition hreal_add :: "hreal => hreal => hreal" where 
   "hreal_add ==
 %(X::hreal) Y::hreal.
    hreal
@@ -2197,8 +2157,7 @@
            w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
   by (import hreal hreal_add)
 
-constdefs
-  hreal_mul :: "hreal => hreal => hreal" 
+definition hreal_mul :: "hreal => hreal => hreal" where 
   "hreal_mul ==
 %(X::hreal) Y::hreal.
    hreal
@@ -2214,8 +2173,7 @@
            w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
   by (import hreal hreal_mul)
 
-constdefs
-  hreal_inv :: "hreal => hreal" 
+definition hreal_inv :: "hreal => hreal" where 
   "hreal_inv ==
 %X::hreal.
    hreal
@@ -2233,8 +2191,7 @@
            (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
   by (import hreal hreal_inv)
 
-constdefs
-  hreal_sup :: "(hreal => bool) => hreal" 
+definition hreal_sup :: "(hreal => bool) => hreal" where 
   "hreal_sup ==
 %P::hreal => bool. hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
 
@@ -2242,8 +2199,7 @@
    hreal_sup P = hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
   by (import hreal hreal_sup)
 
-constdefs
-  hreal_lt :: "hreal => hreal => bool" 
+definition hreal_lt :: "hreal => hreal => bool" where 
   "hreal_lt ==
 %(X::hreal) Y::hreal.
    X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x)"
@@ -2301,8 +2257,7 @@
 lemma HREAL_NOZERO: "ALL (X::hreal) Y::hreal. hreal_add X Y ~= X"
   by (import hreal HREAL_NOZERO)
 
-constdefs
-  hreal_sub :: "hreal => hreal => hreal" 
+definition hreal_sub :: "hreal => hreal => hreal" where 
   "hreal_sub ==
 %(Y::hreal) X::hreal.
    hreal
@@ -2358,15 +2313,13 @@
 (ALL x::nat. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))"
   by (import numeral numeral_suc)
 
-constdefs
-  iZ :: "nat => nat" 
+definition iZ :: "nat => nat" where 
   "iZ == %x::nat. x"
 
 lemma iZ: "ALL x::nat. iZ x = x"
   by (import numeral iZ)
 
-constdefs
-  iiSUC :: "nat => nat" 
+definition iiSUC :: "nat => nat" where 
   "iiSUC == %n::nat. Suc (Suc n)"
 
 lemma iiSUC: "ALL n::nat. iiSUC n = Suc (Suc n)"
@@ -2699,8 +2652,7 @@
     iBIT_cases (NUMERAL_BIT2 n) zf bf1 bf2 = bf2 n)"
   by (import numeral iBIT_cases)
 
-constdefs
-  iDUB :: "nat => nat" 
+definition iDUB :: "nat => nat" where 
   "iDUB == %x::nat. x + x"
 
 lemma iDUB: "ALL x::nat. iDUB x = x + x"
@@ -2771,8 +2723,7 @@
    NUMERAL_BIT2 x * xa = iDUB (iZ (x * xa + xa))"
   by (import numeral numeral_mult)
 
-constdefs
-  iSQR :: "nat => nat" 
+definition iSQR :: "nat => nat" where 
   "iSQR == %x::nat. x * x"
 
 lemma iSQR: "ALL x::nat. iSQR x = x * x"
@@ -2809,8 +2760,7 @@
        ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)"
   by (import ind_type INJ_INVERSE2)
 
-constdefs
-  NUMPAIR :: "nat => nat => nat" 
+definition NUMPAIR :: "nat => nat => nat" where 
   "NUMPAIR == %(x::nat) y::nat. 2 ^ x * (2 * y + 1)"
 
 lemma NUMPAIR: "ALL (x::nat) y::nat. NUMPAIR x y = 2 ^ x * (2 * y + 1)"
@@ -2831,8 +2781,7 @@
 specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y"
   by (import ind_type NUMPAIR_DEST)
 
-constdefs
-  NUMSUM :: "bool => nat => nat" 
+definition NUMSUM :: "bool => nat => nat" where 
   "NUMSUM == %(b::bool) x::nat. if b then Suc (2 * x) else 2 * x"
 
 lemma NUMSUM: "ALL (b::bool) x::nat. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)"
@@ -2849,8 +2798,7 @@
 specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y"
   by (import ind_type NUMSUM_DEST)
 
-constdefs
-  INJN :: "nat => nat => 'a => bool" 
+definition INJN :: "nat => nat => 'a => bool" where 
   "INJN == %(m::nat) (n::nat) a::'a::type. n = m"
 
 lemma INJN: "ALL m::nat. INJN m = (%(n::nat) a::'a::type. n = m)"
@@ -2859,8 +2807,7 @@
 lemma INJN_INJ: "ALL (n1::nat) n2::nat. (INJN n1 = INJN n2) = (n1 = n2)"
   by (import ind_type INJN_INJ)
 
-constdefs
-  INJA :: "'a => nat => 'a => bool" 
+definition INJA :: "'a => nat => 'a => bool" where 
   "INJA == %(a::'a::type) (n::nat) b::'a::type. b = a"
 
 lemma INJA: "ALL a::'a::type. INJA a = (%(n::nat) b::'a::type. b = a)"
@@ -2869,8 +2816,7 @@
 lemma INJA_INJ: "ALL (a1::'a::type) a2::'a::type. (INJA a1 = INJA a2) = (a1 = a2)"
   by (import ind_type INJA_INJ)
 
-constdefs
-  INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" 
+definition INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" where 
   "INJF == %(f::nat => nat => 'a::type => bool) n::nat. f (NUMFST n) (NUMSND n)"
 
 lemma INJF: "ALL f::nat => nat => 'a::type => bool.
@@ -2881,8 +2827,7 @@
    (INJF f1 = INJF f2) = (f1 = f2)"
   by (import ind_type INJF_INJ)
 
-constdefs
-  INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" 
+definition INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" where 
   "INJP ==
 %(f1::nat => 'a::type => bool) (f2::nat => 'a::type => bool) (n::nat)
    a::'a::type. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a"
@@ -2898,8 +2843,7 @@
    (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
   by (import ind_type INJP_INJ)
 
-constdefs
-  ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" 
+definition ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" where 
   "ZCONSTR ==
 %(c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
    INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
@@ -2908,8 +2852,7 @@
    ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
   by (import ind_type ZCONSTR)
 
-constdefs
-  ZBOT :: "nat => 'a => bool" 
+definition ZBOT :: "nat => 'a => bool" where 
   "ZBOT == INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
 
 lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
@@ -2919,8 +2862,7 @@
    ZCONSTR x xa xb ~= ZBOT"
   by (import ind_type ZCONSTR_ZBOT)
 
-constdefs
-  ZRECSPACE :: "(nat => 'a => bool) => bool" 
+definition ZRECSPACE :: "(nat => 'a => bool) => bool" where 
   "ZRECSPACE ==
 %a0::nat => 'a::type => bool.
    ALL ZRECSPACE'::(nat => 'a::type => bool) => bool.
@@ -2993,15 +2935,13 @@
 (ALL r::nat => 'a::type => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))"
   by (import ind_type recspace_repfns)
 
-constdefs
-  BOTTOM :: "'a recspace" 
+definition BOTTOM :: "'a recspace" where 
   "BOTTOM == mk_rec ZBOT"
 
 lemma BOTTOM: "BOTTOM = mk_rec ZBOT"
   by (import ind_type BOTTOM)
 
-constdefs
-  CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" 
+definition CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" where 
   "CONSTR ==
 %(c::nat) (i::'a::type) r::nat => 'a::type recspace.
    mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))"
@@ -3049,15 +2989,13 @@
 (ALL (a::'a::type) (f::nat => 'a::type) n::nat. FCONS a f (Suc n) = f n)"
   by (import ind_type FCONS)
 
-constdefs
-  FNIL :: "nat => 'a" 
+definition FNIL :: "nat => 'a" where 
   "FNIL == %n::nat. SOME x::'a::type. True"
 
 lemma FNIL: "ALL n::nat. FNIL n = (SOME x::'a::type. True)"
   by (import ind_type FNIL)
 
-constdefs
-  ISO :: "('a => 'b) => ('b => 'a) => bool" 
+definition ISO :: "('a => 'b) => ('b => 'a) => bool" where 
   "ISO ==
 %(f::'a::type => 'b::type) g::'b::type => 'a::type.
    (ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y)"
@@ -3434,8 +3372,7 @@
    (EX x::'a::type. IN x s & (ALL y::'a::type. IN y s --> M x <= M y))"
   by (import pred_set SET_MINIMUM)
 
-constdefs
-  EMPTY :: "'a => bool" 
+definition EMPTY :: "'a => bool" where 
   "EMPTY == %x::'a::type. False"
 
 lemma EMPTY_DEF: "EMPTY = (%x::'a::type. False)"
@@ -3468,8 +3405,7 @@
 lemma EQ_UNIV: "(ALL x::'a::type. IN x (s::'a::type => bool)) = (s = pred_set.UNIV)"
   by (import pred_set EQ_UNIV)
 
-constdefs
-  SUBSET :: "('a => bool) => ('a => bool) => bool" 
+definition SUBSET :: "('a => bool) => ('a => bool) => bool" where 
   "SUBSET ==
 %(s::'a::type => bool) t::'a::type => bool.
    ALL x::'a::type. IN x s --> IN x t"
@@ -3501,8 +3437,7 @@
 lemma UNIV_SUBSET: "ALL x::'a::type => bool. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)"
   by (import pred_set UNIV_SUBSET)
 
-constdefs
-  PSUBSET :: "('a => bool) => ('a => bool) => bool" 
+definition PSUBSET :: "('a => bool) => ('a => bool) => bool" where 
   "PSUBSET == %(s::'a::type => bool) t::'a::type => bool. SUBSET s t & s ~= t"
 
 lemma PSUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
@@ -3640,8 +3575,7 @@
    pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)"
   by (import pred_set INTER_OVER_UNION)
 
-constdefs
-  DISJOINT :: "('a => bool) => ('a => bool) => bool" 
+definition DISJOINT :: "('a => bool) => ('a => bool) => bool" where 
   "DISJOINT ==
 %(s::'a::type => bool) t::'a::type => bool. pred_set.INTER s t = EMPTY"
 
@@ -3672,8 +3606,7 @@
    DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)"
   by (import pred_set DISJOINT_UNION_BOTH)
 
-constdefs
-  DIFF :: "('a => bool) => ('a => bool) => 'a => bool" 
+definition DIFF :: "('a => bool) => ('a => bool) => 'a => bool" where 
   "DIFF ==
 %(s::'a::type => bool) t::'a::type => bool.
    GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))"
@@ -3702,8 +3635,7 @@
 lemma DIFF_EQ_EMPTY: "ALL x::'a::type => bool. DIFF x x = EMPTY"
   by (import pred_set DIFF_EQ_EMPTY)
 
-constdefs
-  INSERT :: "'a => ('a => bool) => 'a => bool" 
+definition INSERT :: "'a => ('a => bool) => 'a => bool" where 
   "INSERT ==
 %(x::'a::type) s::'a::type => bool.
    GSPEC (%y::'a::type. (y, y = x | IN y s))"
@@ -3778,8 +3710,7 @@
    DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))"
   by (import pred_set INSERT_DIFF)
 
-constdefs
-  DELETE :: "('a => bool) => 'a => 'a => bool" 
+definition DELETE :: "('a => bool) => 'a => 'a => bool" where 
   "DELETE == %(s::'a::type => bool) x::'a::type. DIFF s (INSERT x EMPTY)"
 
 lemma DELETE_DEF: "ALL (s::'a::type => bool) x::'a::type. DELETE s x = DIFF s (INSERT x EMPTY)"
@@ -3852,8 +3783,7 @@
 specification (CHOICE) CHOICE_DEF: "ALL x::'a::type => bool. x ~= EMPTY --> IN (CHOICE x) x"
   by (import pred_set CHOICE_DEF)
 
-constdefs
-  REST :: "('a => bool) => 'a => bool" 
+definition REST :: "('a => bool) => 'a => bool" where 
   "REST == %s::'a::type => bool. DELETE s (CHOICE s)"
 
 lemma REST_DEF: "ALL s::'a::type => bool. REST s = DELETE s (CHOICE s)"
@@ -3871,8 +3801,7 @@
 lemma REST_PSUBSET: "ALL x::'a::type => bool. x ~= EMPTY --> PSUBSET (REST x) x"
   by (import pred_set REST_PSUBSET)
 
-constdefs
-  SING :: "('a => bool) => bool" 
+definition SING :: "('a => bool) => bool" where 
   "SING == %s::'a::type => bool. EX x::'a::type. s = INSERT x EMPTY"
 
 lemma SING_DEF: "ALL s::'a::type => bool. SING s = (EX x::'a::type. s = INSERT x EMPTY)"
@@ -3917,8 +3846,7 @@
 lemma SING_IFF_EMPTY_REST: "ALL x::'a::type => bool. SING x = (x ~= EMPTY & REST x = EMPTY)"
   by (import pred_set SING_IFF_EMPTY_REST)
 
-constdefs
-  IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" 
+definition IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" where 
   "IMAGE ==
 %(f::'a::type => 'b::type) s::'a::type => bool.
    GSPEC (%x::'a::type. (f x, IN x s))"
@@ -3971,8 +3899,7 @@
     (pred_set.INTER (IMAGE f s) (IMAGE f t))"
   by (import pred_set IMAGE_INTER)
 
-constdefs
-  INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+definition INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
   "INJ ==
 %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
    (ALL x::'a::type. IN x s --> IN (f x) t) &
@@ -3998,8 +3925,7 @@
    (ALL xa::'a::type => bool. INJ x xa EMPTY = (xa = EMPTY))"
   by (import pred_set INJ_EMPTY)
 
-constdefs
-  SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+definition SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
   "SURJ ==
 %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
    (ALL x::'a::type. IN x s --> IN (f x) t) &
@@ -4028,8 +3954,7 @@
    SURJ x xa xb = (IMAGE x xa = xb)"
   by (import pred_set IMAGE_SURJ)
 
-constdefs
-  BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+definition BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
   "BIJ ==
 %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
    INJ f s t & SURJ f s t"
@@ -4065,8 +3990,7 @@
    SURJ f s t --> (ALL x::'b::type. IN x t --> f (RINV f s x) = x)"
   by (import pred_set RINV_DEF)
 
-constdefs
-  FINITE :: "('a => bool) => bool" 
+definition FINITE :: "('a => bool) => bool" where 
   "FINITE ==
 %s::'a::type => bool.
    ALL P::('a::type => bool) => bool.
@@ -4219,8 +4143,7 @@
    (ALL x::'a::type => bool. FINITE x --> P x)"
   by (import pred_set FINITE_COMPLETE_INDUCTION)
 
-constdefs
-  INFINITE :: "('a => bool) => bool" 
+definition INFINITE :: "('a => bool) => bool" where 
   "INFINITE == %s::'a::type => bool. ~ FINITE s"
 
 lemma INFINITE_DEF: "ALL s::'a::type => bool. INFINITE s = (~ FINITE s)"
@@ -4320,8 +4243,7 @@
                                 (f n)))))))))"
   by (import pred_set FINITE_WEAK_ENUMERATE)
 
-constdefs
-  BIGUNION :: "(('a => bool) => bool) => 'a => bool" 
+definition BIGUNION :: "(('a => bool) => bool) => 'a => bool" where 
   "BIGUNION ==
 %P::('a::type => bool) => bool.
    GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))"
@@ -4367,8 +4289,7 @@
    FINITE (BIGUNION x)"
   by (import pred_set FINITE_BIGUNION)
 
-constdefs
-  BIGINTER :: "(('a => bool) => bool) => 'a => bool" 
+definition BIGINTER :: "(('a => bool) => bool) => 'a => bool" where 
   "BIGINTER ==
 %B::('a::type => bool) => bool.
    GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))"
@@ -4406,8 +4327,7 @@
    DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x"
   by (import pred_set DISJOINT_BIGINTER)
 
-constdefs
-  CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" 
+definition CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" where 
   "CROSS ==
 %(P::'a::type => bool) Q::'b::type => bool.
    GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))"
@@ -4460,8 +4380,7 @@
    FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)"
   by (import pred_set FINITE_CROSS_EQ)
 
-constdefs
-  COMPL :: "('a => bool) => 'a => bool" 
+definition COMPL :: "('a => bool) => 'a => bool" where 
   "COMPL == DIFF pred_set.UNIV"
 
 lemma COMPL_DEF: "ALL P::'a::type => bool. COMPL P = DIFF pred_set.UNIV P"
@@ -4513,8 +4432,7 @@
 lemma CARD_COUNT: "ALL n::nat. CARD (count n) = n"
   by (import pred_set CARD_COUNT)
 
-constdefs
-  ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" 
+definition ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" where 
   "ITSET_tupled ==
 %f::'a::type => 'b::type => 'b::type.
    WFREC
@@ -4546,8 +4464,7 @@
         else ARB)"
   by (import pred_set ITSET_tupled_primitive_def)
 
-constdefs
-  ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" 
+definition ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" where 
   "ITSET ==
 %(f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) x1::'b::type.
    ITSET_tupled f (x, x1)"
@@ -4578,8 +4495,7 @@
 
 ;setup_theory operator
 
-constdefs
-  ASSOC :: "('a => 'a => 'a) => bool" 
+definition ASSOC :: "('a => 'a => 'a) => bool" where 
   "ASSOC ==
 %f::'a::type => 'a::type => 'a::type.
    ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z"
@@ -4589,8 +4505,7 @@
    (ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z)"
   by (import operator ASSOC_DEF)
 
-constdefs
-  COMM :: "('a => 'a => 'b) => bool" 
+definition COMM :: "('a => 'a => 'b) => bool" where 
   "COMM ==
 %f::'a::type => 'a::type => 'b::type.
    ALL (x::'a::type) y::'a::type. f x y = f y x"
@@ -4599,8 +4514,7 @@
    COMM f = (ALL (x::'a::type) y::'a::type. f x y = f y x)"
   by (import operator COMM_DEF)
 
-constdefs
-  FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" 
+definition FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" where 
   "FCOMM ==
 %(f::'a::type => 'b::type => 'a::type) g::'c::type => 'a::type => 'a::type.
    ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z"
@@ -4611,8 +4525,7 @@
    (ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z)"
   by (import operator FCOMM_DEF)
 
-constdefs
-  RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" 
+definition RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" where 
   "RIGHT_ID ==
 %(f::'a::type => 'b::type => 'a::type) e::'b::type.
    ALL x::'a::type. f x e = x"
@@ -4621,8 +4534,7 @@
    RIGHT_ID f e = (ALL x::'a::type. f x e = x)"
   by (import operator RIGHT_ID_DEF)
 
-constdefs
-  LEFT_ID :: "('a => 'b => 'b) => 'a => bool" 
+definition LEFT_ID :: "('a => 'b => 'b) => 'a => bool" where 
   "LEFT_ID ==
 %(f::'a::type => 'b::type => 'b::type) e::'a::type.
    ALL x::'b::type. f e x = x"
@@ -4631,8 +4543,7 @@
    LEFT_ID f e = (ALL x::'b::type. f e x = x)"
   by (import operator LEFT_ID_DEF)
 
-constdefs
-  MONOID :: "('a => 'a => 'a) => 'a => bool" 
+definition MONOID :: "('a => 'a => 'a) => 'a => bool" where 
   "MONOID ==
 %(f::'a::type => 'a::type => 'a::type) e::'a::type.
    ASSOC f & RIGHT_ID f e & LEFT_ID f e"
@@ -4690,15 +4601,13 @@
 lemma IS_EL_DEF: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l"
   by (import rich_list IS_EL_DEF)
 
-constdefs
-  AND_EL :: "bool list => bool" 
+definition AND_EL :: "bool list => bool" where 
   "AND_EL == list_all I"
 
 lemma AND_EL_DEF: "AND_EL = list_all I"
   by (import rich_list AND_EL_DEF)
 
-constdefs
-  OR_EL :: "bool list => bool" 
+definition OR_EL :: "bool list => bool" where 
   "OR_EL == list_exists I"
 
 lemma OR_EL_DEF: "OR_EL = list_exists I"
@@ -4816,16 +4725,14 @@
     (if P x then ([], x # l) else (x # fst (SPLITP P l), snd (SPLITP P l))))"
   by (import rich_list SPLITP)
 
-constdefs
-  PREFIX :: "('a => bool) => 'a list => 'a list" 
+definition PREFIX :: "('a => bool) => 'a list => 'a list" where 
   "PREFIX == %(P::'a::type => bool) l::'a::type list. fst (SPLITP (Not o P) l)"
 
 lemma PREFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list.
    PREFIX P l = fst (SPLITP (Not o P) l)"
   by (import rich_list PREFIX_DEF)
 
-constdefs
-  SUFFIX :: "('a => bool) => 'a list => 'a list" 
+definition SUFFIX :: "('a => bool) => 'a list => 'a list" where 
   "SUFFIX ==
 %P::'a::type => bool.
    foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else [])
@@ -4837,15 +4744,13 @@
     [] l"
   by (import rich_list SUFFIX_DEF)
 
-constdefs
-  UNZIP_FST :: "('a * 'b) list => 'a list" 
+definition UNZIP_FST :: "('a * 'b) list => 'a list" where 
   "UNZIP_FST == %l::('a::type * 'b::type) list. fst (unzip l)"
 
 lemma UNZIP_FST_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_FST l = fst (unzip l)"
   by (import rich_list UNZIP_FST_DEF)
 
-constdefs
-  UNZIP_SND :: "('a * 'b) list => 'b list" 
+definition UNZIP_SND :: "('a * 'b) list => 'b list" where 
   "UNZIP_SND == %l::('a::type * 'b::type) list. snd (unzip l)"
 
 lemma UNZIP_SND_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_SND l = snd (unzip l)"
@@ -5916,8 +5821,7 @@
 
 ;setup_theory state_transformer
 
-constdefs
-  UNIT :: "'b => 'a => 'b * 'a" 
+definition UNIT :: "'b => 'a => 'b * 'a" where 
   "(op ==::('b::type => 'a::type => 'b::type * 'a::type)
         => ('b::type => 'a::type => 'b::type * 'a::type) => prop)
  (UNIT::'b::type => 'a::type => 'b::type * 'a::type)
@@ -5926,8 +5830,7 @@
 lemma UNIT_DEF: "ALL x::'b::type. UNIT x = Pair x"
   by (import state_transformer UNIT_DEF)
 
-constdefs
-  BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" 
+definition BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" where 
   "(op ==::(('a::type => 'b::type * 'a::type)
          => ('b::type => 'a::type => 'c::type * 'a::type)
             => 'a::type => 'c::type * 'a::type)
@@ -5967,8 +5870,7 @@
              g)))"
   by (import state_transformer BIND_DEF)
 
-constdefs
-  MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" 
+definition MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" where 
   "MMAP ==
 %(f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type.
    BIND m (UNIT o f)"
@@ -5977,8 +5879,7 @@
    MMAP f m = BIND m (UNIT o f)"
   by (import state_transformer MMAP_DEF)
 
-constdefs
-  JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" 
+definition JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" where 
   "JOIN ==
 %z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. BIND z I"
 
--- a/src/HOL/Import/HOL/HOL4Prob.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL/HOL4Prob.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -373,8 +373,7 @@
    alg_twin x y = (EX l::bool list. x = SNOC True l & y = SNOC False l)"
   by (import prob_canon alg_twin_def)
 
-constdefs
-  alg_order_tupled :: "bool list * bool list => bool" 
+definition alg_order_tupled :: "bool list * bool list => bool" where 
   "(op ==::(bool list * bool list => bool)
         => (bool list * bool list => bool) => prop)
  (alg_order_tupled::bool list * bool list => bool)
@@ -1917,8 +1916,7 @@
    P 0 & (ALL v::nat. P (Suc v div 2) --> P (Suc v)) --> All P"
   by (import prob_uniform unif_bound_ind)
 
-constdefs
-  unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" 
+definition unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" where 
   "unif_tupled ==
 WFREC
  (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool.
@@ -1963,8 +1961,7 @@
    (ALL v::nat. All (P v))"
   by (import prob_uniform unif_ind)
 
-constdefs
-  uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" 
+definition uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" where 
   "uniform_tupled ==
 WFREC
  (SOME R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool.
--- a/src/HOL/Import/HOL/HOL4Real.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL/HOL4Real.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -39,29 +39,25 @@
    hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z"
   by (import realax HREAL_LT_LADD)
 
-constdefs
-  treal_0 :: "hreal * hreal" 
+definition treal_0 :: "hreal * hreal" where 
   "treal_0 == (hreal_1, hreal_1)"
 
 lemma treal_0: "treal_0 = (hreal_1, hreal_1)"
   by (import realax treal_0)
 
-constdefs
-  treal_1 :: "hreal * hreal" 
+definition treal_1 :: "hreal * hreal" where 
   "treal_1 == (hreal_add hreal_1 hreal_1, hreal_1)"
 
 lemma treal_1: "treal_1 = (hreal_add hreal_1 hreal_1, hreal_1)"
   by (import realax treal_1)
 
-constdefs
-  treal_neg :: "hreal * hreal => hreal * hreal" 
+definition treal_neg :: "hreal * hreal => hreal * hreal" where 
   "treal_neg == %(x::hreal, y::hreal). (y, x)"
 
 lemma treal_neg: "ALL (x::hreal) y::hreal. treal_neg (x, y) = (y, x)"
   by (import realax treal_neg)
 
-constdefs
-  treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
   "treal_add ==
 %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
    (hreal_add x1 x2, hreal_add y1 y2)"
@@ -70,8 +66,7 @@
    treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)"
   by (import realax treal_add)
 
-constdefs
-  treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
   "treal_mul ==
 %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
    (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2),
@@ -83,8 +78,7 @@
     hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))"
   by (import realax treal_mul)
 
-constdefs
-  treal_lt :: "hreal * hreal => hreal * hreal => bool" 
+definition treal_lt :: "hreal * hreal => hreal * hreal => bool" where 
   "treal_lt ==
 %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
    hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
@@ -93,8 +87,7 @@
    treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
   by (import realax treal_lt)
 
-constdefs
-  treal_inv :: "hreal * hreal => hreal * hreal" 
+definition treal_inv :: "hreal * hreal => hreal * hreal" where 
   "treal_inv ==
 %(x::hreal, y::hreal).
    if x = y then treal_0
@@ -110,8 +103,7 @@
          else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1))"
   by (import realax treal_inv)
 
-constdefs
-  treal_eq :: "hreal * hreal => hreal * hreal => bool" 
+definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where 
   "treal_eq ==
 %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
    hreal_add x1 y2 = hreal_add x2 y1"
@@ -194,15 +186,13 @@
    treal_lt treal_0 (treal_mul x y)"
   by (import realax TREAL_LT_MUL)
 
-constdefs
-  treal_of_hreal :: "hreal => hreal * hreal" 
+definition treal_of_hreal :: "hreal => hreal * hreal" where 
   "treal_of_hreal == %x::hreal. (hreal_add x hreal_1, hreal_1)"
 
 lemma treal_of_hreal: "ALL x::hreal. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)"
   by (import realax treal_of_hreal)
 
-constdefs
-  hreal_of_treal :: "hreal * hreal => hreal" 
+definition hreal_of_treal :: "hreal * hreal => hreal" where 
   "hreal_of_treal == %(x::hreal, y::hreal). SOME d::hreal. x = hreal_add y d"
 
 lemma hreal_of_treal: "ALL (x::hreal) y::hreal.
@@ -579,8 +569,7 @@
    (EX x::real. ALL y::real. (EX x::real. P x & y < x) = (y < x))"
   by (import real REAL_SUP_EXISTS)
 
-constdefs
-  sup :: "(real => bool) => real" 
+definition sup :: "(real => bool) => real" where 
   "sup ==
 %P::real => bool.
    SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)"
@@ -781,8 +770,7 @@
 
 ;setup_theory topology
 
-constdefs
-  re_Union :: "(('a => bool) => bool) => 'a => bool" 
+definition re_Union :: "(('a => bool) => bool) => 'a => bool" where 
   "re_Union ==
 %(P::('a::type => bool) => bool) x::'a::type.
    EX s::'a::type => bool. P s & s x"
@@ -791,8 +779,7 @@
    re_Union P = (%x::'a::type. EX s::'a::type => bool. P s & s x)"
   by (import topology re_Union)
 
-constdefs
-  re_union :: "('a => bool) => ('a => bool) => 'a => bool" 
+definition re_union :: "('a => bool) => ('a => bool) => 'a => bool" where 
   "re_union ==
 %(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x | Q x"
 
@@ -800,8 +787,7 @@
    re_union P Q = (%x::'a::type. P x | Q x)"
   by (import topology re_union)
 
-constdefs
-  re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" 
+definition re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" where 
   "re_intersect ==
 %(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x & Q x"
 
@@ -809,22 +795,19 @@
    re_intersect P Q = (%x::'a::type. P x & Q x)"
   by (import topology re_intersect)
 
-constdefs
-  re_null :: "'a => bool" 
+definition re_null :: "'a => bool" where 
   "re_null == %x::'a::type. False"
 
 lemma re_null: "re_null = (%x::'a::type. False)"
   by (import topology re_null)
 
-constdefs
-  re_universe :: "'a => bool" 
+definition re_universe :: "'a => bool" where 
   "re_universe == %x::'a::type. True"
 
 lemma re_universe: "re_universe = (%x::'a::type. True)"
   by (import topology re_universe)
 
-constdefs
-  re_subset :: "('a => bool) => ('a => bool) => bool" 
+definition re_subset :: "('a => bool) => ('a => bool) => bool" where 
   "re_subset ==
 %(P::'a::type => bool) Q::'a::type => bool. ALL x::'a::type. P x --> Q x"
 
@@ -832,8 +815,7 @@
    re_subset P Q = (ALL x::'a::type. P x --> Q x)"
   by (import topology re_subset)
 
-constdefs
-  re_compl :: "('a => bool) => 'a => bool" 
+definition re_compl :: "('a => bool) => 'a => bool" where 
   "re_compl == %(P::'a::type => bool) x::'a::type. ~ P x"
 
 lemma re_compl: "ALL P::'a::type => bool. re_compl P = (%x::'a::type. ~ P x)"
@@ -853,8 +835,7 @@
    re_subset P Q & re_subset Q R --> re_subset P R"
   by (import topology SUBSET_TRANS)
 
-constdefs
-  istopology :: "(('a => bool) => bool) => bool" 
+definition istopology :: "(('a => bool) => bool) => bool" where 
   "istopology ==
 %L::('a::type => bool) => bool.
    L re_null &
@@ -900,8 +881,7 @@
    re_subset xa (open x) --> open x (re_Union xa)"
   by (import topology TOPOLOGY_UNION)
 
-constdefs
-  neigh :: "'a topology => ('a => bool) * 'a => bool" 
+definition neigh :: "'a topology => ('a => bool) * 'a => bool" where 
   "neigh ==
 %(top::'a::type topology) (N::'a::type => bool, x::'a::type).
    EX P::'a::type => bool. open top P & re_subset P N & P x"
@@ -932,16 +912,14 @@
        S' x --> (EX N::'a::type => bool. neigh top (N, x) & re_subset N S'))"
   by (import topology OPEN_NEIGH)
 
-constdefs
-  closed :: "'a topology => ('a => bool) => bool" 
+definition closed :: "'a topology => ('a => bool) => bool" where 
   "closed == %(L::'a::type topology) S'::'a::type => bool. open L (re_compl S')"
 
 lemma closed: "ALL (L::'a::type topology) S'::'a::type => bool.
    closed L S' = open L (re_compl S')"
   by (import topology closed)
 
-constdefs
-  limpt :: "'a topology => 'a => ('a => bool) => bool" 
+definition limpt :: "'a topology => 'a => ('a => bool) => bool" where 
   "limpt ==
 %(top::'a::type topology) (x::'a::type) S'::'a::type => bool.
    ALL N::'a::type => bool.
@@ -957,8 +935,7 @@
    closed top S' = (ALL x::'a::type. limpt top x S' --> S' x)"
   by (import topology CLOSED_LIMPT)
 
-constdefs
-  ismet :: "('a * 'a => real) => bool" 
+definition ismet :: "('a * 'a => real) => bool" where 
   "ismet ==
 %m::'a::type * 'a::type => real.
    (ALL (x::'a::type) y::'a::type. (m (x, y) = 0) = (x = y)) &
@@ -1012,8 +989,7 @@
    x ~= y --> 0 < dist m (x, y)"
   by (import topology METRIC_NZ)
 
-constdefs
-  mtop :: "'a metric => 'a topology" 
+definition mtop :: "'a metric => 'a topology" where 
   "mtop ==
 %m::'a::type metric.
    topology
@@ -1042,8 +1018,7 @@
        S' xa --> (EX e>0. ALL y::'a::type. dist x (xa, y) < e --> S' y))"
   by (import topology MTOP_OPEN)
 
-constdefs
-  B :: "'a metric => 'a * real => 'a => bool" 
+definition B :: "'a metric => 'a * real => 'a => bool" where 
   "B ==
 %(m::'a::type metric) (x::'a::type, e::real) y::'a::type. dist m (x, y) < e"
 
@@ -1067,8 +1042,7 @@
 lemma ISMET_R1: "ismet (%(x::real, y::real). abs (y - x))"
   by (import topology ISMET_R1)
 
-constdefs
-  mr1 :: "real metric" 
+definition mr1 :: "real metric" where 
   "mr1 == metric (%(x::real, y::real). abs (y - x))"
 
 lemma mr1: "mr1 = metric (%(x::real, y::real). abs (y - x))"
@@ -1105,8 +1079,7 @@
 
 ;setup_theory nets
 
-constdefs
-  dorder :: "('a => 'a => bool) => bool" 
+definition dorder :: "('a => 'a => bool) => bool" where 
   "dorder ==
 %g::'a::type => 'a::type => bool.
    ALL (x::'a::type) y::'a::type.
@@ -1120,8 +1093,7 @@
        (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y)))"
   by (import nets dorder)
 
-constdefs
-  tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" 
+definition tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" where 
   "tends ==
 %(s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology,
    g::'b::type => 'b::type => bool).
@@ -1137,8 +1109,7 @@
        (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m))))"
   by (import nets tends)
 
-constdefs
-  bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" 
+definition bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" where 
   "bounded ==
 %(m::'a::type metric, g::'b::type => 'b::type => bool)
    f::'b::type => 'a::type.
@@ -1152,8 +1123,7 @@
        g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k))"
   by (import nets bounded)
 
-constdefs
-  tendsto :: "'a metric * 'a => 'a => 'a => bool" 
+definition tendsto :: "'a metric * 'a => 'a => 'a => bool" where 
   "tendsto ==
 %(m::'a::type metric, x::'a::type) (y::'a::type) z::'a::type.
    0 < dist m (x, y) & dist m (x, y) <= dist m (x, z)"
@@ -1366,15 +1336,13 @@
    hol4--> x x1 & hol4--> x x2 --> x1 = x2"
   by (import seq SEQ_UNIQ)
 
-constdefs
-  convergent :: "(nat => real) => bool" 
+definition convergent :: "(nat => real) => bool" where 
   "convergent == %f::nat => real. Ex (hol4--> f)"
 
 lemma convergent: "ALL f::nat => real. convergent f = Ex (hol4--> f)"
   by (import seq convergent)
 
-constdefs
-  cauchy :: "(nat => real) => bool" 
+definition cauchy :: "(nat => real) => bool" where 
   "cauchy ==
 %f::nat => real.
    ALL e>0.
@@ -1388,8 +1356,7 @@
           ALL (m::nat) n::nat. N <= m & N <= n --> abs (f m - f n) < e)"
   by (import seq cauchy)
 
-constdefs
-  lim :: "(nat => real) => real" 
+definition lim :: "(nat => real) => real" where 
   "lim == %f::nat => real. Eps (hol4--> f)"
 
 lemma lim: "ALL f::nat => real. lim f = Eps (hol4--> f)"
@@ -1398,8 +1365,7 @@
 lemma SEQ_LIM: "ALL f::nat => real. convergent f = hol4--> f (lim f)"
   by (import seq SEQ_LIM)
 
-constdefs
-  subseq :: "(nat => nat) => bool" 
+definition subseq :: "(nat => nat) => bool" where 
   "subseq == %f::nat => nat. ALL (m::nat) n::nat. m < n --> f m < f n"
 
 lemma subseq: "ALL f::nat => nat. subseq f = (ALL (m::nat) n::nat. m < n --> f m < f n)"
@@ -1541,23 +1507,20 @@
    (ALL (a::real) b::real. a <= b --> P (a, b))"
   by (import seq BOLZANO_LEMMA)
 
-constdefs
-  sums :: "(nat => real) => real => bool" 
+definition sums :: "(nat => real) => real => bool" where 
   "sums == %f::nat => real. hol4--> (%n::nat. real.sum (0, n) f)"
 
 lemma sums: "ALL (f::nat => real) s::real.
    sums f s = hol4--> (%n::nat. real.sum (0, n) f) s"
   by (import seq sums)
 
-constdefs
-  summable :: "(nat => real) => bool" 
+definition summable :: "(nat => real) => bool" where 
   "summable == %f::nat => real. Ex (sums f)"
 
 lemma summable: "ALL f::nat => real. summable f = Ex (sums f)"
   by (import seq summable)
 
-constdefs
-  suminf :: "(nat => real) => real" 
+definition suminf :: "(nat => real) => real" where 
   "suminf == %f::nat => real. Eps (sums f)"
 
 lemma suminf: "ALL f::nat => real. suminf f = Eps (sums f)"
@@ -1692,8 +1655,7 @@
 
 ;setup_theory lim
 
-constdefs
-  tends_real_real :: "(real => real) => real => real => bool" 
+definition tends_real_real :: "(real => real) => real => real => bool" where 
   "tends_real_real ==
 %(f::real => real) (l::real) x0::real.
    tends f l (mtop mr1, tendsto (mr1, x0))"
@@ -1763,8 +1725,7 @@
    tends_real_real f l x0"
   by (import lim LIM_TRANSFORM)
 
-constdefs
-  diffl :: "(real => real) => real => real => bool" 
+definition diffl :: "(real => real) => real => real => bool" where 
   "diffl ==
 %(f::real => real) (l::real) x::real.
    tends_real_real (%h::real. (f (x + h) - f x) / h) l 0"
@@ -1773,8 +1734,7 @@
    diffl f l x = tends_real_real (%h::real. (f (x + h) - f x) / h) l 0"
   by (import lim diffl)
 
-constdefs
-  contl :: "(real => real) => real => bool" 
+definition contl :: "(real => real) => real => bool" where 
   "contl ==
 %(f::real => real) x::real. tends_real_real (%h::real. f (x + h)) (f x) 0"
 
@@ -1782,8 +1742,7 @@
    contl f x = tends_real_real (%h::real. f (x + h)) (f x) 0"
   by (import lim contl)
 
-constdefs
-  differentiable :: "(real => real) => real => bool" 
+definition differentiable :: "(real => real) => real => bool" where 
   "differentiable == %(f::real => real) x::real. EX l::real. diffl f l x"
 
 lemma differentiable: "ALL (f::real => real) x::real.
@@ -2127,8 +2086,7 @@
    summable (%n::nat. f n * z ^ n)"
   by (import powser POWSER_INSIDE)
 
-constdefs
-  diffs :: "(nat => real) => nat => real" 
+definition diffs :: "(nat => real) => nat => real" where 
   "diffs == %(c::nat => real) n::nat. real (Suc n) * c (Suc n)"
 
 lemma diffs: "ALL c::nat => real. diffs c = (%n::nat. real (Suc n) * c (Suc n))"
@@ -2204,15 +2162,13 @@
 
 ;setup_theory transc
 
-constdefs
-  exp :: "real => real" 
+definition exp :: "real => real" where 
   "exp == %x::real. suminf (%n::nat. inverse (real (FACT n)) * x ^ n)"
 
 lemma exp: "ALL x::real. exp x = suminf (%n::nat. inverse (real (FACT n)) * x ^ n)"
   by (import transc exp)
 
-constdefs
-  cos :: "real => real" 
+definition cos :: "real => real" where 
   "cos ==
 %x::real.
    suminf
@@ -2226,8 +2182,7 @@
         (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
   by (import transc cos)
 
-constdefs
-  sin :: "real => real" 
+definition sin :: "real => real" where 
   "sin ==
 %x::real.
    suminf
@@ -2364,8 +2319,7 @@
 lemma EXP_TOTAL: "ALL y>0. EX x::real. exp x = y"
   by (import transc EXP_TOTAL)
 
-constdefs
-  ln :: "real => real" 
+definition ln :: "real => real" where 
   "ln == %x::real. SOME u::real. exp u = x"
 
 lemma ln: "ALL x::real. ln x = (SOME u::real. exp u = x)"
@@ -2410,16 +2364,14 @@
 lemma LN_POS: "ALL x>=1. 0 <= ln x"
   by (import transc LN_POS)
 
-constdefs
-  root :: "nat => real => real" 
+definition root :: "nat => real => real" where 
   "root == %(n::nat) x::real. SOME u::real. (0 < x --> 0 < u) & u ^ n = x"
 
 lemma root: "ALL (n::nat) x::real.
    root n x = (SOME u::real. (0 < x --> 0 < u) & u ^ n = x)"
   by (import transc root)
 
-constdefs
-  sqrt :: "real => real" 
+definition sqrt :: "real => real" where 
   "sqrt == root 2"
 
 lemma sqrt: "ALL x::real. sqrt x = root 2 x"
@@ -2584,8 +2536,7 @@
 lemma COS_ISZERO: "EX! x::real. 0 <= x & x <= 2 & cos x = 0"
   by (import transc COS_ISZERO)
 
-constdefs
-  pi :: "real" 
+definition pi :: "real" where 
   "pi == 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)"
 
 lemma pi: "pi = 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)"
@@ -2689,8 +2640,7 @@
     (EX n::nat. EVEN n & x = - (real n * (pi / 2))))"
   by (import transc SIN_ZERO)
 
-constdefs
-  tan :: "real => real" 
+definition tan :: "real => real" where 
   "tan == %x::real. sin x / cos x"
 
 lemma tan: "ALL x::real. tan x = sin x / cos x"
@@ -2736,23 +2686,20 @@
 lemma TAN_TOTAL: "ALL y::real. EX! x::real. - (pi / 2) < x & x < pi / 2 & tan x = y"
   by (import transc TAN_TOTAL)
 
-constdefs
-  asn :: "real => real" 
+definition asn :: "real => real" where 
   "asn == %y::real. SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y"
 
 lemma asn: "ALL y::real.
    asn y = (SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y)"
   by (import transc asn)
 
-constdefs
-  acs :: "real => real" 
+definition acs :: "real => real" where 
   "acs == %y::real. SOME x::real. 0 <= x & x <= pi & cos x = y"
 
 lemma acs: "ALL y::real. acs y = (SOME x::real. 0 <= x & x <= pi & cos x = y)"
   by (import transc acs)
 
-constdefs
-  atn :: "real => real" 
+definition atn :: "real => real" where 
   "atn == %y::real. SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y"
 
 lemma atn: "ALL y::real. atn y = (SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y)"
@@ -2845,8 +2792,7 @@
 lemma DIFF_ATN: "ALL x::real. diffl atn (inverse (1 + x ^ 2)) x"
   by (import transc DIFF_ATN)
 
-constdefs
-  division :: "real * real => (nat => real) => bool" 
+definition division :: "real * real => (nat => real) => bool" where 
   "(op ==::(real * real => (nat => real) => bool)
         => (real * real => (nat => real) => bool) => prop)
  (division::real * real => (nat => real) => bool)
@@ -2898,8 +2844,7 @@
                                   b)))))))))"
   by (import transc division)
 
-constdefs
-  dsize :: "(nat => real) => nat" 
+definition dsize :: "(nat => real) => nat" where 
   "(op ==::((nat => real) => nat) => ((nat => real) => nat) => prop)
  (dsize::(nat => real) => nat)
  (%D::nat => real.
@@ -2937,8 +2882,7 @@
                     ((op =::real => real => bool) (D n) (D N)))))))"
   by (import transc dsize)
 
-constdefs
-  tdiv :: "real * real => (nat => real) * (nat => real) => bool" 
+definition tdiv :: "real * real => (nat => real) * (nat => real) => bool" where 
   "tdiv ==
 %(a::real, b::real) (D::nat => real, p::nat => real).
    division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n))"
@@ -2948,16 +2892,14 @@
    (division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n)))"
   by (import transc tdiv)
 
-constdefs
-  gauge :: "(real => bool) => (real => real) => bool" 
+definition gauge :: "(real => bool) => (real => real) => bool" where 
   "gauge == %(E::real => bool) g::real => real. ALL x::real. E x --> 0 < g x"
 
 lemma gauge: "ALL (E::real => bool) g::real => real.
    gauge E g = (ALL x::real. E x --> 0 < g x)"
   by (import transc gauge)
 
-constdefs
-  fine :: "(real => real) => (nat => real) * (nat => real) => bool" 
+definition fine :: "(real => real) => (nat => real) * (nat => real) => bool" where 
   "(op ==::((real => real) => (nat => real) * (nat => real) => bool)
         => ((real => real) => (nat => real) * (nat => real) => bool)
            => prop)
@@ -3000,8 +2942,7 @@
                          (g (p n))))))))"
   by (import transc fine)
 
-constdefs
-  rsum :: "(nat => real) * (nat => real) => (real => real) => real" 
+definition rsum :: "(nat => real) * (nat => real) => (real => real) => real" where 
   "rsum ==
 %(D::nat => real, p::nat => real) f::real => real.
    real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))"
@@ -3011,8 +2952,7 @@
    real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))"
   by (import transc rsum)
 
-constdefs
-  Dint :: "real * real => (real => real) => real => bool" 
+definition Dint :: "real * real => (real => real) => real => bool" where 
   "Dint ==
 %(a::real, b::real) (f::real => real) k::real.
    ALL e>0.
@@ -3313,8 +3253,7 @@
     poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)"
   by (import poly poly_diff_aux_def)
 
-constdefs
-  diff :: "real list => real list" 
+definition diff :: "real list => real list" where 
   "diff == %l::real list. if l = [] then [] else poly_diff_aux 1 (tl l)"
 
 lemma poly_diff_def: "ALL l::real list. diff l = (if l = [] then [] else poly_diff_aux 1 (tl l))"
@@ -3622,8 +3561,7 @@
    poly p = poly q --> poly (diff p) = poly (diff q)"
   by (import poly POLY_DIFF_WELLDEF)
 
-constdefs
-  poly_divides :: "real list => real list => bool" 
+definition poly_divides :: "real list => real list => bool" where 
   "poly_divides ==
 %(p1::real list) p2::real list.
    EX q::real list. poly p2 = poly (poly_mul p1 q)"
@@ -3681,8 +3619,7 @@
        ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)"
   by (import poly POLY_ORDER)
 
-constdefs
-  poly_order :: "real => real list => nat" 
+definition poly_order :: "real => real list => nat" where 
   "poly_order ==
 %(a::real) p::real list.
    SOME n::nat.
@@ -3754,8 +3691,7 @@
    (ALL a::real. poly_order a q = (if poly_order a p = 0 then 0 else 1))"
   by (import poly POLY_SQUAREFREE_DECOMP_ORDER)
 
-constdefs
-  rsquarefree :: "real list => bool" 
+definition rsquarefree :: "real list => bool" where 
   "rsquarefree ==
 %p::real list.
    poly p ~= poly [] &
@@ -3798,8 +3734,7 @@
 lemma POLY_NORMALIZE: "ALL t::real list. poly (normalize t) = poly t"
   by (import poly POLY_NORMALIZE)
 
-constdefs
-  degree :: "real list => nat" 
+definition degree :: "real list => nat" where 
   "degree == %p::real list. PRE (length (normalize p))"
 
 lemma degree: "ALL p::real list. degree p = PRE (length (normalize p))"
--- a/src/HOL/Import/HOL/HOL4Vec.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL/HOL4Vec.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -164,8 +164,7 @@
 lemma word_base0_def: "word_base0 = (%a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM)))"
   by (import word_base word_base0_def)
 
-constdefs
-  WORD :: "'a list => 'a word" 
+definition WORD :: "'a list => 'a word" where 
   "WORD == word_base0"
 
 lemma WORD: "WORD = word_base0"
@@ -680,8 +679,7 @@
 
 ;setup_theory word_num
 
-constdefs
-  LVAL :: "('a => nat) => nat => 'a list => nat" 
+definition LVAL :: "('a => nat) => nat => 'a list => nat" where 
   "LVAL ==
 %(f::'a::type => nat) b::nat. foldl (%(e::nat) x::'a::type. b * e + f x) 0"
 
@@ -756,8 +754,7 @@
     SNOC (frep (m mod b)) (NLIST n frep b (m div b)))"
   by (import word_num NLIST_DEF)
 
-constdefs
-  NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" 
+definition NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" where 
   "NWORD ==
 %(n::nat) (frep::nat => 'a::type) (b::nat) m::nat. WORD (NLIST n frep b m)"
 
@@ -1076,8 +1073,7 @@
    EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)"
   by (import word_bitop EXISTSABIT_WCAT)
 
-constdefs
-  SHR :: "bool => 'a => 'a word => 'a word * 'a" 
+definition SHR :: "bool => 'a => 'a word => 'a word * 'a" where 
   "SHR ==
 %(f::bool) (b::'a::type) w::'a::type word.
    (WCAT
@@ -1093,8 +1089,7 @@
     bit 0 w)"
   by (import word_bitop SHR_DEF)
 
-constdefs
-  SHL :: "bool => 'a word => 'a => 'a * 'a word" 
+definition SHL :: "bool => 'a word => 'a => 'a * 'a word" where 
   "SHL ==
 %(f::bool) (w::'a::type word) b::'a::type.
    (bit (PRE (WORDLEN w)) w,
@@ -1196,8 +1191,7 @@
 
 ;setup_theory bword_num
 
-constdefs
-  BV :: "bool => nat" 
+definition BV :: "bool => nat" where 
   "BV == %b::bool. if b then Suc 0 else 0"
 
 lemma BV_DEF: "ALL b::bool. BV b = (if b then Suc 0 else 0)"
@@ -1248,15 +1242,13 @@
              BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))"
   by (import bword_num BNVAL_WCAT)
 
-constdefs
-  VB :: "nat => bool" 
+definition VB :: "nat => bool" where 
   "VB == %n::nat. n mod 2 ~= 0"
 
 lemma VB_DEF: "ALL n::nat. VB n = (n mod 2 ~= 0)"
   by (import bword_num VB_DEF)
 
-constdefs
-  NBWORD :: "nat => nat => bool word" 
+definition NBWORD :: "nat => nat => bool word" where 
   "NBWORD == %(n::nat) m::nat. WORD (NLIST n VB 2 m)"
 
 lemma NBWORD_DEF: "ALL (n::nat) m::nat. NBWORD n m = WORD (NLIST n VB 2 m)"
--- a/src/HOL/Import/HOL/HOL4Word32.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL/HOL4Word32.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -68,8 +68,7 @@
    BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
   by (import bits BITS_def)
 
-constdefs
-  bit :: "nat => nat => bool" 
+definition bit :: "nat => nat => bool" where 
   "bit == %(b::nat) n::nat. BITS b b n = 1"
 
 lemma BIT_def: "ALL (b::nat) n::nat. bit b n = (BITS b b n = 1)"
@@ -840,15 +839,13 @@
 lemma w_T_def: "w_T = mk_word32 (EQUIV COMP0)"
   by (import word32 w_T_def)
 
-constdefs
-  word_suc :: "word32 => word32" 
+definition word_suc :: "word32 => word32" where 
   "word_suc == %T1::word32. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
 
 lemma word_suc: "ALL T1::word32. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
   by (import word32 word_suc)
 
-constdefs
-  word_add :: "word32 => word32 => word32" 
+definition word_add :: "word32 => word32 => word32" where 
   "word_add ==
 %(T1::word32) T2::word32.
    mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
@@ -858,8 +855,7 @@
    mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
   by (import word32 word_add)
 
-constdefs
-  word_mul :: "word32 => word32 => word32" 
+definition word_mul :: "word32 => word32 => word32" where 
   "word_mul ==
 %(T1::word32) T2::word32.
    mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
@@ -869,8 +865,7 @@
    mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
   by (import word32 word_mul)
 
-constdefs
-  word_1comp :: "word32 => word32" 
+definition word_1comp :: "word32 => word32" where 
   "word_1comp ==
 %T1::word32. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
 
@@ -878,8 +873,7 @@
    word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
   by (import word32 word_1comp)
 
-constdefs
-  word_2comp :: "word32 => word32" 
+definition word_2comp :: "word32 => word32" where 
   "word_2comp ==
 %T1::word32. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
 
@@ -887,24 +881,21 @@
    word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
   by (import word32 word_2comp)
 
-constdefs
-  word_lsr1 :: "word32 => word32" 
+definition word_lsr1 :: "word32 => word32" where 
   "word_lsr1 == %T1::word32. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
 
 lemma word_lsr1: "ALL T1::word32.
    word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
   by (import word32 word_lsr1)
 
-constdefs
-  word_asr1 :: "word32 => word32" 
+definition word_asr1 :: "word32 => word32" where 
   "word_asr1 == %T1::word32. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
 
 lemma word_asr1: "ALL T1::word32.
    word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
   by (import word32 word_asr1)
 
-constdefs
-  word_ror1 :: "word32 => word32" 
+definition word_ror1 :: "word32 => word32" where 
   "word_ror1 == %T1::word32. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
 
 lemma word_ror1: "ALL T1::word32.
@@ -940,8 +931,7 @@
 lemma MSB_def: "ALL T1::word32. MSB T1 = MSBn (Eps (dest_word32 T1))"
   by (import word32 MSB_def)
 
-constdefs
-  bitwise_or :: "word32 => word32 => word32" 
+definition bitwise_or :: "word32 => word32 => word32" where 
   "bitwise_or ==
 %(T1::word32) T2::word32.
    mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
@@ -951,8 +941,7 @@
    mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
   by (import word32 bitwise_or)
 
-constdefs
-  bitwise_eor :: "word32 => word32 => word32" 
+definition bitwise_eor :: "word32 => word32 => word32" where 
   "bitwise_eor ==
 %(T1::word32) T2::word32.
    mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
@@ -962,8 +951,7 @@
    mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
   by (import word32 bitwise_eor)
 
-constdefs
-  bitwise_and :: "word32 => word32 => word32" 
+definition bitwise_and :: "word32 => word32 => word32" where 
   "bitwise_and ==
 %(T1::word32) T2::word32.
    mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
@@ -1154,36 +1142,31 @@
 lemma ADD_TWO_COMP2: "ALL x::word32. word_add (word_2comp x) x = w_0"
   by (import word32 ADD_TWO_COMP2)
 
-constdefs
-  word_sub :: "word32 => word32 => word32" 
+definition word_sub :: "word32 => word32 => word32" where 
   "word_sub == %(a::word32) b::word32. word_add a (word_2comp b)"
 
 lemma word_sub: "ALL (a::word32) b::word32. word_sub a b = word_add a (word_2comp b)"
   by (import word32 word_sub)
 
-constdefs
-  word_lsl :: "word32 => nat => word32" 
+definition word_lsl :: "word32 => nat => word32" where 
   "word_lsl == %(a::word32) n::nat. word_mul a (n2w (2 ^ n))"
 
 lemma word_lsl: "ALL (a::word32) n::nat. word_lsl a n = word_mul a (n2w (2 ^ n))"
   by (import word32 word_lsl)
 
-constdefs
-  word_lsr :: "word32 => nat => word32" 
+definition word_lsr :: "word32 => nat => word32" where 
   "word_lsr == %(a::word32) n::nat. (word_lsr1 ^^ n) a"
 
 lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^^ n) a"
   by (import word32 word_lsr)
 
-constdefs
-  word_asr :: "word32 => nat => word32" 
+definition word_asr :: "word32 => nat => word32" where 
   "word_asr == %(a::word32) n::nat. (word_asr1 ^^ n) a"
 
 lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^^ n) a"
   by (import word32 word_asr)
 
-constdefs
-  word_ror :: "word32 => nat => word32" 
+definition word_ror :: "word32 => nat => word32" where 
   "word_ror == %(a::word32) n::nat. (word_ror1 ^^ n) a"
 
 lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^^ n) a"
--- a/src/HOL/Import/HOL4Compat.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOL4Compat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -15,8 +15,7 @@
 lemma COND_DEF:"(If b t f) = (@x. ((b = True) --> (x = t)) & ((b = False) --> (x = f)))"
   by auto
 
-constdefs
-  LET :: "['a \<Rightarrow> 'b,'a] \<Rightarrow> 'b"
+definition LET :: "['a \<Rightarrow> 'b,'a] \<Rightarrow> 'b" where
   "LET f s == f s"
 
 lemma [hol4rew]: "LET f s = Let s f"
@@ -119,10 +118,10 @@
 lemma LESS_OR_EQ: "m <= (n::nat) = (m < n | m = n)"
   by auto
 
-constdefs
-  nat_gt :: "nat => nat => bool"
+definition nat_gt :: "nat => nat => bool" where
   "nat_gt == %m n. n < m"
-  nat_ge :: "nat => nat => bool"
+
+definition nat_ge :: "nat => nat => bool" where
   "nat_ge == %m n. nat_gt m n | m = n"
 
 lemma [hol4rew]: "nat_gt m n = (n < m)"
@@ -200,8 +199,7 @@
   qed
 qed;
 
-constdefs
-  FUNPOW :: "('a => 'a) => nat => 'a => 'a"
+definition FUNPOW :: "('a => 'a) => nat => 'a => 'a" where
   "FUNPOW f n == f ^^ n"
 
 lemma FUNPOW: "(ALL f x. (f ^^ 0) x = x) &
@@ -229,14 +227,16 @@
 lemma DIVISION: "(0::nat) < n --> (!k. (k = k div n * n + k mod n) & k mod n < n)"
   by simp
 
-constdefs
-  ALT_ZERO :: nat
+definition ALT_ZERO :: nat where 
   "ALT_ZERO == 0"
-  NUMERAL_BIT1 :: "nat \<Rightarrow> nat"
+
+definition NUMERAL_BIT1 :: "nat \<Rightarrow> nat" where 
   "NUMERAL_BIT1 n == n + (n + Suc 0)"
-  NUMERAL_BIT2 :: "nat \<Rightarrow> nat"
+
+definition NUMERAL_BIT2 :: "nat \<Rightarrow> nat" where 
   "NUMERAL_BIT2 n == n + (n + Suc (Suc 0))"
-  NUMERAL :: "nat \<Rightarrow> nat"
+
+definition NUMERAL :: "nat \<Rightarrow> nat" where 
   "NUMERAL x == x"
 
 lemma [hol4rew]: "NUMERAL ALT_ZERO = 0"
@@ -329,8 +329,7 @@
 lemma NULL_DEF: "(null [] = True) & (!h t. null (h # t) = False)"
   by simp
 
-constdefs
-  sum :: "nat list \<Rightarrow> nat"
+definition sum :: "nat list \<Rightarrow> nat" where
   "sum l == foldr (op +) l 0"
 
 lemma SUM: "(sum [] = 0) & (!h t. sum (h#t) = h + sum t)"
@@ -359,8 +358,7 @@
   (ALL n x. replicate (Suc n) x = x # replicate n x)"
   by simp
 
-constdefs
-  FOLDR :: "[['a,'b]\<Rightarrow>'b,'b,'a list] \<Rightarrow> 'b"
+definition FOLDR :: "[['a,'b]\<Rightarrow>'b,'b,'a list] \<Rightarrow> 'b" where
   "FOLDR f e l == foldr f l e"
 
 lemma [hol4rew]: "FOLDR f e l = foldr f l e"
@@ -418,8 +416,7 @@
 lemma list_CASES: "(l = []) | (? t h. l = h#t)"
   by (induct l,auto)
 
-constdefs
-  ZIP :: "'a list * 'b list \<Rightarrow> ('a * 'b) list"
+definition ZIP :: "'a list * 'b list \<Rightarrow> ('a * 'b) list" where
   "ZIP == %(a,b). zip a b"
 
 lemma ZIP: "(zip [] [] = []) &
@@ -514,8 +511,7 @@
 lemma pow: "(!x::real. x ^ 0 = 1) & (!x::real. ALL n. x ^ (Suc n) = x * x ^ n)"
   by simp
 
-constdefs
-  real_gt :: "real => real => bool" 
+definition real_gt :: "real => real => bool" where 
   "real_gt == %x y. y < x"
 
 lemma [hol4rew]: "real_gt x y = (y < x)"
@@ -524,8 +520,7 @@
 lemma real_gt: "ALL x (y::real). (y < x) = (y < x)"
   by simp
 
-constdefs
-  real_ge :: "real => real => bool"
+definition real_ge :: "real => real => bool" where
   "real_ge x y == y <= x"
 
 lemma [hol4rew]: "real_ge x y = (y <= x)"
--- a/src/HOL/Import/HOLLight/HOLLight.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOLLight/HOLLight.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -95,8 +95,7 @@
 lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t"
   by (import hollight EXCLUDED_MIDDLE)
 
-constdefs
-  COND :: "bool => 'A => 'A => 'A" 
+definition COND :: "bool => 'A => 'A => 'A" where 
   "COND ==
 %(t::bool) (t1::'A::type) t2::'A::type.
    SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2)"
@@ -173,15 +172,13 @@
 (b & P x True | ~ b & P y False)"
   by (import hollight th_cond)
 
-constdefs
-  LET_END :: "'A => 'A" 
+definition LET_END :: "'A => 'A" where 
   "LET_END == %t::'A::type. t"
 
 lemma DEF_LET_END: "LET_END = (%t::'A::type. t)"
   by (import hollight DEF_LET_END)
 
-constdefs
-  GABS :: "('A => bool) => 'A" 
+definition GABS :: "('A => bool) => 'A" where 
   "(op ==::(('A::type => bool) => 'A::type)
         => (('A::type => bool) => 'A::type) => prop)
  (GABS::('A::type => bool) => 'A::type)
@@ -193,8 +190,7 @@
  (Eps::('A::type => bool) => 'A::type)"
   by (import hollight DEF_GABS)
 
-constdefs
-  GEQ :: "'A => 'A => bool" 
+definition GEQ :: "'A => 'A => bool" where 
   "(op ==::('A::type => 'A::type => bool)
         => ('A::type => 'A::type => bool) => prop)
  (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)"
@@ -208,8 +204,7 @@
    x = Pair_Rep a b"
   by (import hollight PAIR_EXISTS_THM)
 
-constdefs
-  CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" 
+definition CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" where 
   "CURRY ==
 %(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type.
    u (ua, ub)"
@@ -219,8 +214,7 @@
     u (ua, ub))"
   by (import hollight DEF_CURRY)
 
-constdefs
-  UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" 
+definition UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where 
   "UNCURRY ==
 %(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type.
    u (fst ua) (snd ua)"
@@ -230,8 +224,7 @@
     u (fst ua) (snd ua))"
   by (import hollight DEF_UNCURRY)
 
-constdefs
-  PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" 
+definition PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where 
   "PASSOC ==
 %(u::('A::type * 'B::type) * 'C::type => 'D::type)
    ua::'A::type * 'B::type * 'C::type.
@@ -291,8 +284,7 @@
    (m * n = NUMERAL_BIT1 0) = (m = NUMERAL_BIT1 0 & n = NUMERAL_BIT1 0)"
   by (import hollight MULT_EQ_1)
 
-constdefs
-  EXP :: "nat => nat => nat" 
+definition EXP :: "nat => nat => nat" where 
   "EXP ==
 SOME EXP::nat => nat => nat.
    (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) &
@@ -549,8 +541,7 @@
    (EX m::nat. P m & (ALL x::nat. P x --> <= x m))"
   by (import hollight num_MAX)
 
-constdefs
-  EVEN :: "nat => bool" 
+definition EVEN :: "nat => bool" where 
   "EVEN ==
 SOME EVEN::nat => bool.
    EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))"
@@ -560,8 +551,7 @@
     EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))"
   by (import hollight DEF_EVEN)
 
-constdefs
-  ODD :: "nat => bool" 
+definition ODD :: "nat => bool" where 
   "ODD ==
 SOME ODD::nat => bool. ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))"
 
@@ -641,8 +631,7 @@
 lemma SUC_SUB1: "ALL x::nat. Suc x - NUMERAL_BIT1 0 = x"
   by (import hollight SUC_SUB1)
 
-constdefs
-  FACT :: "nat => nat" 
+definition FACT :: "nat => nat" where 
   "FACT ==
 SOME FACT::nat => nat.
    FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)"
@@ -669,8 +658,7 @@
       COND (n = 0) (x = 0 & xa = 0) (m = x * n + xa & < xa n)"
   by (import hollight DIVMOD_EXIST_0)
 
-constdefs
-  DIV :: "nat => nat => nat" 
+definition DIV :: "nat => nat => nat" where 
   "DIV ==
 SOME q::nat => nat => nat.
    EX r::nat => nat => nat.
@@ -686,8 +674,7 @@
            (m = q m n * n + r m n & < (r m n) n))"
   by (import hollight DEF_DIV)
 
-constdefs
-  MOD :: "nat => nat => nat" 
+definition MOD :: "nat => nat => nat" where 
   "MOD ==
 SOME r::nat => nat => nat.
    ALL (m::nat) n::nat.
@@ -877,8 +864,7 @@
  n ~= 0 & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))"
   by (import hollight DIVMOD_ELIM_THM)
 
-constdefs
-  eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" 
+definition eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" where 
   "eqeq ==
 %(u::'q_9910::type) (ua::'q_9909::type)
    ub::'q_9910::type => 'q_9909::type => bool. ub u ua"
@@ -888,8 +874,7 @@
     ub::'q_9910::type => 'q_9909::type => bool. ub u ua)"
   by (import hollight DEF__equal__equal_)
 
-constdefs
-  mod_nat :: "nat => nat => nat => bool" 
+definition mod_nat :: "nat => nat => nat => bool" where 
   "mod_nat ==
 %(u::nat) (ua::nat) ub::nat. EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2"
 
@@ -898,8 +883,7 @@
     EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2)"
   by (import hollight DEF_mod_nat)
 
-constdefs
-  minimal :: "(nat => bool) => nat" 
+definition minimal :: "(nat => bool) => nat" where 
   "minimal == %u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m)"
 
 lemma DEF_minimal: "minimal =
@@ -910,8 +894,7 @@
    Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))"
   by (import hollight MINIMAL)
 
-constdefs
-  WF :: "('A => 'A => bool) => bool" 
+definition WF :: "('A => 'A => bool) => bool" where 
   "WF ==
 %u::'A::type => 'A::type => bool.
    ALL P::'A::type => bool.
@@ -1605,8 +1588,7 @@
        ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)"
   by (import hollight INJ_INVERSE2)
 
-constdefs
-  NUMPAIR :: "nat => nat => nat" 
+definition NUMPAIR :: "nat => nat => nat" where 
   "NUMPAIR ==
 %(u::nat) ua::nat.
    EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u *
@@ -1626,8 +1608,7 @@
    (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
   by (import hollight NUMPAIR_INJ)
 
-constdefs
-  NUMFST :: "nat => nat" 
+definition NUMFST :: "nat => nat" where 
   "NUMFST ==
 SOME X::nat => nat.
    EX Y::nat => nat.
@@ -1639,8 +1620,7 @@
        ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
   by (import hollight DEF_NUMFST)
 
-constdefs
-  NUMSND :: "nat => nat" 
+definition NUMSND :: "nat => nat" where 
   "NUMSND ==
 SOME Y::nat => nat.
    ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
@@ -1650,8 +1630,7 @@
     ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
   by (import hollight DEF_NUMSND)
 
-constdefs
-  NUMSUM :: "bool => nat => nat" 
+definition NUMSUM :: "bool => nat => nat" where 
   "NUMSUM ==
 %(u::bool) ua::nat.
    COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))
@@ -1667,8 +1646,7 @@
    (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
   by (import hollight NUMSUM_INJ)
 
-constdefs
-  NUMLEFT :: "nat => bool" 
+definition NUMLEFT :: "nat => bool" where 
   "NUMLEFT ==
 SOME X::nat => bool.
    EX Y::nat => nat.
@@ -1680,8 +1658,7 @@
        ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
   by (import hollight DEF_NUMLEFT)
 
-constdefs
-  NUMRIGHT :: "nat => nat" 
+definition NUMRIGHT :: "nat => nat" where 
   "NUMRIGHT ==
 SOME Y::nat => nat.
    ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y"
@@ -1691,8 +1668,7 @@
     ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
   by (import hollight DEF_NUMRIGHT)
 
-constdefs
-  INJN :: "nat => nat => 'A => bool" 
+definition INJN :: "nat => nat => 'A => bool" where 
   "INJN == %(u::nat) (n::nat) a::'A::type. n = u"
 
 lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. n = u)"
@@ -1710,8 +1686,7 @@
            ((op =::nat => nat => bool) n1 n2)))"
   by (import hollight INJN_INJ)
 
-constdefs
-  INJA :: "'A => nat => 'A => bool" 
+definition INJA :: "'A => nat => 'A => bool" where 
   "INJA == %(u::'A::type) (n::nat) b::'A::type. b = u"
 
 lemma DEF_INJA: "INJA = (%(u::'A::type) (n::nat) b::'A::type. b = u)"
@@ -1720,8 +1695,7 @@
 lemma INJA_INJ: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)"
   by (import hollight INJA_INJ)
 
-constdefs
-  INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" 
+definition INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" where 
   "INJF == %(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n)"
 
 lemma DEF_INJF: "INJF =
@@ -1732,8 +1706,7 @@
    (INJF f1 = INJF f2) = (f1 = f2)"
   by (import hollight INJF_INJ)
 
-constdefs
-  INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" 
+definition INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where 
   "INJP ==
 %(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat)
    a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a)"
@@ -1748,8 +1721,7 @@
    (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
   by (import hollight INJP_INJ)
 
-constdefs
-  ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" 
+definition ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where 
   "ZCONSTR ==
 %(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool.
    INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))"
@@ -1759,8 +1731,7 @@
     INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))"
   by (import hollight DEF_ZCONSTR)
 
-constdefs
-  ZBOT :: "nat => 'A => bool" 
+definition ZBOT :: "nat => 'A => bool" where 
   "ZBOT == INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)"
 
 lemma DEF_ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)"
@@ -1770,8 +1741,7 @@
    ZCONSTR x xa xb ~= ZBOT"
   by (import hollight ZCONSTR_ZBOT)
 
-constdefs
-  ZRECSPACE :: "(nat => 'A => bool) => bool" 
+definition ZRECSPACE :: "(nat => 'A => bool) => bool" where 
   "ZRECSPACE ==
 %a::nat => 'A::type => bool.
    ALL ZRECSPACE'::(nat => 'A::type => bool) => bool.
@@ -1809,8 +1779,7 @@
   [where a="a :: 'A recspace" and r=r ,
    OF type_definition_recspace]
 
-constdefs
-  BOTTOM :: "'A recspace" 
+definition BOTTOM :: "'A recspace" where 
   "(op ==::'A::type recspace => 'A::type recspace => prop)
  (BOTTOM::'A::type recspace)
  ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace)
@@ -1822,8 +1791,7 @@
    (ZBOT::nat => 'A::type => bool))"
   by (import hollight DEF_BOTTOM)
 
-constdefs
-  CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" 
+definition CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" where 
   "(op ==::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
         => (nat
             => 'A::type => (nat => 'A::type recspace) => 'A::type recspace)
@@ -1900,8 +1868,7 @@
          f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))"
   by (import hollight CONSTR_REC)
 
-constdefs
-  FCONS :: "'A => (nat => 'A) => nat => 'A" 
+definition FCONS :: "'A => (nat => 'A) => nat => 'A" where 
   "FCONS ==
 SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type.
    (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) &
@@ -1917,8 +1884,7 @@
 lemma FCONS_UNDO: "ALL f::nat => 'A::type. f = FCONS (f 0) (f o Suc)"
   by (import hollight FCONS_UNDO)
 
-constdefs
-  FNIL :: "nat => 'A" 
+definition FNIL :: "nat => 'A" where 
   "FNIL == %u::nat. SOME x::'A::type. True"
 
 lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)"
@@ -1995,8 +1961,7 @@
   [where a="a :: 'A hollight.option" and r=r ,
    OF type_definition_option]
 
-constdefs
-  NONE :: "'A hollight.option" 
+definition NONE :: "'A hollight.option" where 
   "(op ==::'A::type hollight.option => 'A::type hollight.option => prop)
  (NONE::'A::type hollight.option)
  ((_mk_option::'A::type recspace => 'A::type hollight.option)
@@ -2093,8 +2058,7 @@
   [where a="a :: 'A hollight.list" and r=r ,
    OF type_definition_list]
 
-constdefs
-  NIL :: "'A hollight.list" 
+definition NIL :: "'A hollight.list" where 
   "(op ==::'A::type hollight.list => 'A::type hollight.list => prop)
  (NIL::'A::type hollight.list)
  ((_mk_list::'A::type recspace => 'A::type hollight.list)
@@ -2114,8 +2078,7 @@
      (%n::nat. BOTTOM::'A::type recspace)))"
   by (import hollight DEF_NIL)
 
-constdefs
-  CONS :: "'A => 'A hollight.list => 'A hollight.list" 
+definition CONS :: "'A => 'A hollight.list => 'A hollight.list" where 
   "(op ==::('A::type => 'A::type hollight.list => 'A::type hollight.list)
         => ('A::type => 'A::type hollight.list => 'A::type hollight.list)
            => prop)
@@ -2160,8 +2123,7 @@
    EX x::bool => 'A::type. x False = a & x True = b"
   by (import hollight bool_RECURSION)
 
-constdefs
-  ISO :: "('A => 'B) => ('B => 'A) => bool" 
+definition ISO :: "('A => 'B) => ('B => 'A) => bool" where 
   "ISO ==
 %(u::'A::type => 'B::type) ua::'B::type => 'A::type.
    (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y)"
@@ -2244,15 +2206,13 @@
      (%n::nat. BOTTOM::bool recspace)))"
   by (import hollight DEF__10303)
 
-constdefs
-  two_1 :: "N_2" 
+definition two_1 :: "N_2" where 
   "two_1 == _10302"
 
 lemma DEF_two_1: "two_1 = _10302"
   by (import hollight DEF_two_1)
 
-constdefs
-  two_2 :: "N_2" 
+definition two_2 :: "N_2" where 
   "two_2 == _10303"
 
 lemma DEF_two_2: "two_2 = _10303"
@@ -2337,22 +2297,19 @@
      (%n::nat. BOTTOM::bool recspace)))"
   by (import hollight DEF__10328)
 
-constdefs
-  three_1 :: "N_3" 
+definition three_1 :: "N_3" where 
   "three_1 == _10326"
 
 lemma DEF_three_1: "three_1 = _10326"
   by (import hollight DEF_three_1)
 
-constdefs
-  three_2 :: "N_3" 
+definition three_2 :: "N_3" where 
   "three_2 == _10327"
 
 lemma DEF_three_2: "three_2 = _10327"
   by (import hollight DEF_three_2)
 
-constdefs
-  three_3 :: "N_3" 
+definition three_3 :: "N_3" where 
   "three_3 == _10328"
 
 lemma DEF_three_3: "three_3 = _10328"
@@ -2365,8 +2322,7 @@
    All P"
   by (import hollight list_INDUCT)
 
-constdefs
-  HD :: "'A hollight.list => 'A" 
+definition HD :: "'A hollight.list => 'A" where 
   "HD ==
 SOME HD::'A::type hollight.list => 'A::type.
    ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h"
@@ -2376,8 +2332,7 @@
     ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h)"
   by (import hollight DEF_HD)
 
-constdefs
-  TL :: "'A hollight.list => 'A hollight.list" 
+definition TL :: "'A hollight.list => 'A hollight.list" where 
   "TL ==
 SOME TL::'A::type hollight.list => 'A::type hollight.list.
    ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t"
@@ -2387,8 +2342,7 @@
     ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t)"
   by (import hollight DEF_TL)
 
-constdefs
-  APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" 
+definition APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" where 
   "APPEND ==
 SOME APPEND::'A::type hollight.list
              => 'A::type hollight.list => 'A::type hollight.list.
@@ -2405,8 +2359,7 @@
         APPEND (CONS h t) l = CONS h (APPEND t l)))"
   by (import hollight DEF_APPEND)
 
-constdefs
-  REVERSE :: "'A hollight.list => 'A hollight.list" 
+definition REVERSE :: "'A hollight.list => 'A hollight.list" where 
   "REVERSE ==
 SOME REVERSE::'A::type hollight.list => 'A::type hollight.list.
    REVERSE NIL = NIL &
@@ -2420,8 +2373,7 @@
         REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL)))"
   by (import hollight DEF_REVERSE)
 
-constdefs
-  LENGTH :: "'A hollight.list => nat" 
+definition LENGTH :: "'A hollight.list => nat" where 
   "LENGTH ==
 SOME LENGTH::'A::type hollight.list => nat.
    LENGTH NIL = 0 &
@@ -2435,8 +2387,7 @@
         LENGTH (CONS h t) = Suc (LENGTH t)))"
   by (import hollight DEF_LENGTH)
 
-constdefs
-  MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" 
+definition MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" where 
   "MAP ==
 SOME MAP::('A::type => 'B::type)
           => 'A::type hollight.list => 'B::type hollight.list.
@@ -2452,8 +2403,7 @@
         MAP f (CONS h t) = CONS (f h) (MAP f t)))"
   by (import hollight DEF_MAP)
 
-constdefs
-  LAST :: "'A hollight.list => 'A" 
+definition LAST :: "'A hollight.list => 'A" where 
   "LAST ==
 SOME LAST::'A::type hollight.list => 'A::type.
    ALL (h::'A::type) t::'A::type hollight.list.
@@ -2465,8 +2415,7 @@
        LAST (CONS h t) = COND (t = NIL) h (LAST t))"
   by (import hollight DEF_LAST)
 
-constdefs
-  REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" 
+definition REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" where 
   "REPLICATE ==
 SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list.
    (ALL x::'q_16860::type. REPLICATE 0 x = NIL) &
@@ -2480,8 +2429,7 @@
         REPLICATE (Suc n) x = CONS x (REPLICATE n x)))"
   by (import hollight DEF_REPLICATE)
 
-constdefs
-  NULL :: "'q_16875 hollight.list => bool" 
+definition NULL :: "'q_16875 hollight.list => bool" where 
   "NULL ==
 SOME NULL::'q_16875::type hollight.list => bool.
    NULL NIL = True &
@@ -2495,8 +2443,7 @@
         NULL (CONS h t) = False))"
   by (import hollight DEF_NULL)
 
-constdefs
-  ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" 
+definition ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" where 
   "ALL_list ==
 SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool.
    (ALL P::'q_16895::type => bool. u P NIL = True) &
@@ -2527,9 +2474,8 @@
         t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t)))"
   by (import hollight DEF_EX)
 
-constdefs
-  ITLIST :: "('q_16939 => 'q_16938 => 'q_16938)
-=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" 
+definition ITLIST :: "('q_16939 => 'q_16938 => 'q_16938)
+=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" where 
   "ITLIST ==
 SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type)
              => 'q_16939::type hollight.list
@@ -2553,8 +2499,7 @@
         ITLIST f (CONS h t) b = f h (ITLIST f t b)))"
   by (import hollight DEF_ITLIST)
 
-constdefs
-  MEM :: "'q_16964 => 'q_16964 hollight.list => bool" 
+definition MEM :: "'q_16964 => 'q_16964 hollight.list => bool" where 
   "MEM ==
 SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool.
    (ALL x::'q_16964::type. MEM x NIL = False) &
@@ -2570,9 +2515,8 @@
         MEM x (CONS h t) = (x = h | MEM x t)))"
   by (import hollight DEF_MEM)
 
-constdefs
-  ALL2 :: "('q_16997 => 'q_17004 => bool)
-=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" 
+definition ALL2 :: "('q_16997 => 'q_17004 => bool)
+=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" where 
   "ALL2 ==
 SOME ALL2::('q_16997::type => 'q_17004::type => bool)
            => 'q_16997::type hollight.list
@@ -2604,10 +2548,9 @@
 ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)"
   by (import hollight ALL2)
 
-constdefs
-  MAP2 :: "('q_17089 => 'q_17096 => 'q_17086)
+definition MAP2 :: "('q_17089 => 'q_17096 => 'q_17086)
 => 'q_17089 hollight.list
-   => 'q_17096 hollight.list => 'q_17086 hollight.list" 
+   => 'q_17096 hollight.list => 'q_17086 hollight.list" where 
   "MAP2 ==
 SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type)
            => 'q_17089::type hollight.list
@@ -2639,8 +2582,7 @@
 CONS (f h1 h2) (MAP2 f t1 t2)"
   by (import hollight MAP2)
 
-constdefs
-  EL :: "nat => 'q_17157 hollight.list => 'q_17157" 
+definition EL :: "nat => 'q_17157 hollight.list => 'q_17157" where 
   "EL ==
 SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type.
    (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) &
@@ -2654,8 +2596,7 @@
         EL (Suc n) l = EL n (TL l)))"
   by (import hollight DEF_EL)
 
-constdefs
-  FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" 
+definition FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" where 
   "FILTER ==
 SOME FILTER::('q_17182::type => bool)
              => 'q_17182::type hollight.list
@@ -2676,8 +2617,7 @@
         COND (P h) (CONS h (FILTER P t)) (FILTER P t)))"
   by (import hollight DEF_FILTER)
 
-constdefs
-  ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" 
+definition ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" where 
   "ASSOC ==
 SOME ASSOC::'q_17211::type
             => ('q_17211::type * 'q_17205::type) hollight.list
@@ -2695,9 +2635,8 @@
        ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))"
   by (import hollight DEF_ASSOC)
 
-constdefs
-  ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233)
-=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" 
+definition ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233)
+=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" where 
   "ITLIST2 ==
 SOME ITLIST2::('q_17235::type
                => 'q_17243::type => 'q_17233::type => 'q_17233::type)
@@ -3041,8 +2980,7 @@
 ALL2 Q l l'"
   by (import hollight MONO_ALL2)
 
-constdefs
-  dist :: "nat * nat => nat" 
+definition dist :: "nat * nat => nat" where 
   "dist == %u::nat * nat. fst u - snd u + (snd u - fst u)"
 
 lemma DEF_dist: "dist = (%u::nat * nat. fst u - snd u + (snd u - fst u))"
@@ -3133,8 +3071,7 @@
    (EX (x::nat) N::nat. ALL i::nat. <= N i --> <= (P i) (Q i + x))"
   by (import hollight BOUNDS_IGNORE)
 
-constdefs
-  is_nadd :: "(nat => nat) => bool" 
+definition is_nadd :: "(nat => nat) => bool" where 
   "is_nadd ==
 %u::nat => nat.
    EX B::nat.
@@ -3211,8 +3148,7 @@
           (A * n + B)"
   by (import hollight NADD_ALTMUL)
 
-constdefs
-  nadd_eq :: "nadd => nadd => bool" 
+definition nadd_eq :: "nadd => nadd => bool" where 
   "nadd_eq ==
 %(u::nadd) ua::nadd.
    EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B"
@@ -3231,8 +3167,7 @@
 lemma NADD_EQ_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_eq x y & nadd_eq y z --> nadd_eq x z"
   by (import hollight NADD_EQ_TRANS)
 
-constdefs
-  nadd_of_num :: "nat => nadd" 
+definition nadd_of_num :: "nat => nadd" where 
   "nadd_of_num == %u::nat. mk_nadd (op * u)"
 
 lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. mk_nadd (op * u))"
@@ -3247,8 +3182,7 @@
 lemma NADD_OF_NUM_EQ: "ALL (m::nat) n::nat. nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)"
   by (import hollight NADD_OF_NUM_EQ)
 
-constdefs
-  nadd_le :: "nadd => nadd => bool" 
+definition nadd_le :: "nadd => nadd => bool" where 
   "nadd_le ==
 %(u::nadd) ua::nadd.
    EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B)"
@@ -3289,8 +3223,7 @@
 lemma NADD_OF_NUM_LE: "ALL (m::nat) n::nat. nadd_le (nadd_of_num m) (nadd_of_num n) = <= m n"
   by (import hollight NADD_OF_NUM_LE)
 
-constdefs
-  nadd_add :: "nadd => nadd => nadd" 
+definition nadd_add :: "nadd => nadd => nadd" where 
   "nadd_add ==
 %(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n)"
 
@@ -3332,8 +3265,7 @@
     (nadd_of_num (x + xa))"
   by (import hollight NADD_OF_NUM_ADD)
 
-constdefs
-  nadd_mul :: "nadd => nadd => nadd" 
+definition nadd_mul :: "nadd => nadd => nadd" where 
   "nadd_mul ==
 %(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n))"
 
@@ -3438,8 +3370,7 @@
    (EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))"
   by (import hollight NADD_LBOUND)
 
-constdefs
-  nadd_rinv :: "nadd => nat => nat" 
+definition nadd_rinv :: "nadd => nat => nat" where 
   "nadd_rinv == %(u::nadd) n::nat. DIV (n * n) (dest_nadd u n)"
 
 lemma DEF_nadd_rinv: "nadd_rinv = (%(u::nadd) n::nat. DIV (n * n) (dest_nadd u n))"
@@ -3528,8 +3459,7 @@
           <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))"
   by (import hollight NADD_MUL_LINV_LEMMA8)
 
-constdefs
-  nadd_inv :: "nadd => nadd" 
+definition nadd_inv :: "nadd => nadd" where 
   "nadd_inv ==
 %u::nadd.
    COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) (mk_nadd (nadd_rinv u))"
@@ -3570,15 +3500,13 @@
   [where a="a :: hreal" and r=r ,
    OF type_definition_hreal]
 
-constdefs
-  hreal_of_num :: "nat => hreal" 
+definition hreal_of_num :: "nat => hreal" where 
   "hreal_of_num == %m::nat. mk_hreal (nadd_eq (nadd_of_num m))"
 
 lemma DEF_hreal_of_num: "hreal_of_num = (%m::nat. mk_hreal (nadd_eq (nadd_of_num m)))"
   by (import hollight DEF_hreal_of_num)
 
-constdefs
-  hreal_add :: "hreal => hreal => hreal" 
+definition hreal_add :: "hreal => hreal => hreal" where 
   "hreal_add ==
 %(x::hreal) y::hreal.
    mk_hreal
@@ -3594,8 +3522,7 @@
             nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya))"
   by (import hollight DEF_hreal_add)
 
-constdefs
-  hreal_mul :: "hreal => hreal => hreal" 
+definition hreal_mul :: "hreal => hreal => hreal" where 
   "hreal_mul ==
 %(x::hreal) y::hreal.
    mk_hreal
@@ -3611,8 +3538,7 @@
             nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya))"
   by (import hollight DEF_hreal_mul)
 
-constdefs
-  hreal_le :: "hreal => hreal => bool" 
+definition hreal_le :: "hreal => hreal => bool" where 
   "hreal_le ==
 %(x::hreal) y::hreal.
    SOME u::bool.
@@ -3626,8 +3552,7 @@
           nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)"
   by (import hollight DEF_hreal_le)
 
-constdefs
-  hreal_inv :: "hreal => hreal" 
+definition hreal_inv :: "hreal => hreal" where 
   "hreal_inv ==
 %x::hreal.
    mk_hreal
@@ -3685,22 +3610,19 @@
    hreal_le a b --> hreal_le (hreal_mul a c) (hreal_mul b c)"
   by (import hollight HREAL_LE_MUL_RCANCEL_IMP)
 
-constdefs
-  treal_of_num :: "nat => hreal * hreal" 
+definition treal_of_num :: "nat => hreal * hreal" where 
   "treal_of_num == %u::nat. (hreal_of_num u, hreal_of_num 0)"
 
 lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num 0))"
   by (import hollight DEF_treal_of_num)
 
-constdefs
-  treal_neg :: "hreal * hreal => hreal * hreal" 
+definition treal_neg :: "hreal * hreal => hreal * hreal" where 
   "treal_neg == %u::hreal * hreal. (snd u, fst u)"
 
 lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))"
   by (import hollight DEF_treal_neg)
 
-constdefs
-  treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
   "treal_add ==
 %(u::hreal * hreal) ua::hreal * hreal.
    (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))"
@@ -3710,8 +3632,7 @@
     (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))"
   by (import hollight DEF_treal_add)
 
-constdefs
-  treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
   "treal_mul ==
 %(u::hreal * hreal) ua::hreal * hreal.
    (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
@@ -3723,8 +3644,7 @@
      hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))"
   by (import hollight DEF_treal_mul)
 
-constdefs
-  treal_le :: "hreal * hreal => hreal * hreal => bool" 
+definition treal_le :: "hreal * hreal => hreal * hreal => bool" where 
   "treal_le ==
 %(u::hreal * hreal) ua::hreal * hreal.
    hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))"
@@ -3734,8 +3654,7 @@
     hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))"
   by (import hollight DEF_treal_le)
 
-constdefs
-  treal_inv :: "hreal * hreal => hreal * hreal" 
+definition treal_inv :: "hreal * hreal => hreal * hreal" where 
   "treal_inv ==
 %u::hreal * hreal.
    COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0)
@@ -3755,8 +3674,7 @@
         hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))"
   by (import hollight DEF_treal_inv)
 
-constdefs
-  treal_eq :: "hreal * hreal => hreal * hreal => bool" 
+definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where 
   "treal_eq ==
 %(u::hreal * hreal) ua::hreal * hreal.
    hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)"
@@ -3916,15 +3834,13 @@
   [where a="a :: hollight.real" and r=r ,
    OF type_definition_real]
 
-constdefs
-  real_of_num :: "nat => hollight.real" 
+definition real_of_num :: "nat => hollight.real" where 
   "real_of_num == %m::nat. mk_real (treal_eq (treal_of_num m))"
 
 lemma DEF_real_of_num: "real_of_num = (%m::nat. mk_real (treal_eq (treal_of_num m)))"
   by (import hollight DEF_real_of_num)
 
-constdefs
-  real_neg :: "hollight.real => hollight.real" 
+definition real_neg :: "hollight.real => hollight.real" where 
   "real_neg ==
 %x1::hollight.real.
    mk_real
@@ -3940,8 +3856,7 @@
             treal_eq (treal_neg x1a) u & dest_real x1 x1a))"
   by (import hollight DEF_real_neg)
 
-constdefs
-  real_add :: "hollight.real => hollight.real => hollight.real" 
+definition real_add :: "hollight.real => hollight.real => hollight.real" where 
   "real_add ==
 %(x1::hollight.real) y1::hollight.real.
    mk_real
@@ -3959,8 +3874,7 @@
             dest_real x1 x1a & dest_real y1 y1a))"
   by (import hollight DEF_real_add)
 
-constdefs
-  real_mul :: "hollight.real => hollight.real => hollight.real" 
+definition real_mul :: "hollight.real => hollight.real => hollight.real" where 
   "real_mul ==
 %(x1::hollight.real) y1::hollight.real.
    mk_real
@@ -3978,8 +3892,7 @@
             dest_real x1 x1a & dest_real y1 y1a))"
   by (import hollight DEF_real_mul)
 
-constdefs
-  real_le :: "hollight.real => hollight.real => bool" 
+definition real_le :: "hollight.real => hollight.real => bool" where 
   "real_le ==
 %(x1::hollight.real) y1::hollight.real.
    SOME u::bool.
@@ -3993,8 +3906,7 @@
           treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)"
   by (import hollight DEF_real_le)
 
-constdefs
-  real_inv :: "hollight.real => hollight.real" 
+definition real_inv :: "hollight.real => hollight.real" where 
   "real_inv ==
 %x::hollight.real.
    mk_real
@@ -4008,15 +3920,13 @@
          EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa))"
   by (import hollight DEF_real_inv)
 
-constdefs
-  real_sub :: "hollight.real => hollight.real => hollight.real" 
+definition real_sub :: "hollight.real => hollight.real => hollight.real" where 
   "real_sub == %(u::hollight.real) ua::hollight.real. real_add u (real_neg ua)"
 
 lemma DEF_real_sub: "real_sub = (%(u::hollight.real) ua::hollight.real. real_add u (real_neg ua))"
   by (import hollight DEF_real_sub)
 
-constdefs
-  real_lt :: "hollight.real => hollight.real => bool" 
+definition real_lt :: "hollight.real => hollight.real => bool" where 
   "real_lt == %(u::hollight.real) ua::hollight.real. ~ real_le ua u"
 
 lemma DEF_real_lt: "real_lt = (%(u::hollight.real) ua::hollight.real. ~ real_le ua u)"
@@ -4040,8 +3950,7 @@
 lemma DEF_real_gt: "hollight.real_gt = (%(u::hollight.real) ua::hollight.real. real_lt ua u)"
   by (import hollight DEF_real_gt)
 
-constdefs
-  real_abs :: "hollight.real => hollight.real" 
+definition real_abs :: "hollight.real => hollight.real" where 
   "real_abs ==
 %u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u)"
 
@@ -4049,8 +3958,7 @@
 (%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u))"
   by (import hollight DEF_real_abs)
 
-constdefs
-  real_pow :: "hollight.real => nat => hollight.real" 
+definition real_pow :: "hollight.real => nat => hollight.real" where 
   "real_pow ==
 SOME real_pow::hollight.real => nat => hollight.real.
    (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) &
@@ -4064,22 +3972,19 @@
         real_pow x (Suc n) = real_mul x (real_pow x n)))"
   by (import hollight DEF_real_pow)
 
-constdefs
-  real_div :: "hollight.real => hollight.real => hollight.real" 
+definition real_div :: "hollight.real => hollight.real => hollight.real" where 
   "real_div == %(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua)"
 
 lemma DEF_real_div: "real_div = (%(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua))"
   by (import hollight DEF_real_div)
 
-constdefs
-  real_max :: "hollight.real => hollight.real => hollight.real" 
+definition real_max :: "hollight.real => hollight.real => hollight.real" where 
   "real_max == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u"
 
 lemma DEF_real_max: "real_max = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u)"
   by (import hollight DEF_real_max)
 
-constdefs
-  real_min :: "hollight.real => hollight.real => hollight.real" 
+definition real_min :: "hollight.real => hollight.real => hollight.real" where 
   "real_min == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua"
 
 lemma DEF_real_min: "real_min = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua)"
@@ -5212,8 +5117,7 @@
 (ALL x::hollight.real. All (P x))"
   by (import hollight REAL_WLOG_LT)
 
-constdefs
-  mod_real :: "hollight.real => hollight.real => hollight.real => bool" 
+definition mod_real :: "hollight.real => hollight.real => hollight.real => bool" where 
   "mod_real ==
 %(u::hollight.real) (ua::hollight.real) ub::hollight.real.
    EX q::hollight.real. real_sub ua ub = real_mul q u"
@@ -5223,8 +5127,7 @@
     EX q::hollight.real. real_sub ua ub = real_mul q u)"
   by (import hollight DEF_mod_real)
 
-constdefs
-  DECIMAL :: "nat => nat => hollight.real" 
+definition DECIMAL :: "nat => nat => hollight.real" where 
   "DECIMAL == %(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua)"
 
 lemma DEF_DECIMAL: "DECIMAL = (%(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua))"
@@ -5267,8 +5170,7 @@
 (real_mul x1 y2 = real_mul x2 y1)"
   by (import hollight RAT_LEMMA5)
 
-constdefs
-  is_int :: "hollight.real => bool" 
+definition is_int :: "hollight.real => bool" where 
   "is_int ==
 %u::hollight.real.
    EX n::nat. u = real_of_num n | u = real_neg (real_of_num n)"
@@ -5297,8 +5199,7 @@
       dest_int x = real_of_num n | dest_int x = real_neg (real_of_num n)"
   by (import hollight dest_int_rep)
 
-constdefs
-  int_le :: "hollight.int => hollight.int => bool" 
+definition int_le :: "hollight.int => hollight.int => bool" where 
   "int_le ==
 %(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua)"
 
@@ -5306,8 +5207,7 @@
 (%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua))"
   by (import hollight DEF_int_le)
 
-constdefs
-  int_lt :: "hollight.int => hollight.int => bool" 
+definition int_lt :: "hollight.int => hollight.int => bool" where 
   "int_lt ==
 %(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua)"
 
@@ -5315,8 +5215,7 @@
 (%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua))"
   by (import hollight DEF_int_lt)
 
-constdefs
-  int_ge :: "hollight.int => hollight.int => bool" 
+definition int_ge :: "hollight.int => hollight.int => bool" where 
   "int_ge ==
 %(u::hollight.int) ua::hollight.int.
    hollight.real_ge (dest_int u) (dest_int ua)"
@@ -5326,8 +5225,7 @@
     hollight.real_ge (dest_int u) (dest_int ua))"
   by (import hollight DEF_int_ge)
 
-constdefs
-  int_gt :: "hollight.int => hollight.int => bool" 
+definition int_gt :: "hollight.int => hollight.int => bool" where 
   "int_gt ==
 %(u::hollight.int) ua::hollight.int.
    hollight.real_gt (dest_int u) (dest_int ua)"
@@ -5337,8 +5235,7 @@
     hollight.real_gt (dest_int u) (dest_int ua))"
   by (import hollight DEF_int_gt)
 
-constdefs
-  int_of_num :: "nat => hollight.int" 
+definition int_of_num :: "nat => hollight.int" where 
   "int_of_num == %u::nat. mk_int (real_of_num u)"
 
 lemma DEF_int_of_num: "int_of_num = (%u::nat. mk_int (real_of_num u))"
@@ -5347,8 +5244,7 @@
 lemma int_of_num_th: "ALL x::nat. dest_int (int_of_num x) = real_of_num x"
   by (import hollight int_of_num_th)
 
-constdefs
-  int_neg :: "hollight.int => hollight.int" 
+definition int_neg :: "hollight.int => hollight.int" where 
   "int_neg == %u::hollight.int. mk_int (real_neg (dest_int u))"
 
 lemma DEF_int_neg: "int_neg = (%u::hollight.int. mk_int (real_neg (dest_int u)))"
@@ -5357,8 +5253,7 @@
 lemma int_neg_th: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)"
   by (import hollight int_neg_th)
 
-constdefs
-  int_add :: "hollight.int => hollight.int => hollight.int" 
+definition int_add :: "hollight.int => hollight.int => hollight.int" where 
   "int_add ==
 %(u::hollight.int) ua::hollight.int.
    mk_int (real_add (dest_int u) (dest_int ua))"
@@ -5372,8 +5267,7 @@
    dest_int (int_add x xa) = real_add (dest_int x) (dest_int xa)"
   by (import hollight int_add_th)
 
-constdefs
-  int_sub :: "hollight.int => hollight.int => hollight.int" 
+definition int_sub :: "hollight.int => hollight.int => hollight.int" where 
   "int_sub ==
 %(u::hollight.int) ua::hollight.int.
    mk_int (real_sub (dest_int u) (dest_int ua))"
@@ -5387,8 +5281,7 @@
    dest_int (int_sub x xa) = real_sub (dest_int x) (dest_int xa)"
   by (import hollight int_sub_th)
 
-constdefs
-  int_mul :: "hollight.int => hollight.int => hollight.int" 
+definition int_mul :: "hollight.int => hollight.int => hollight.int" where 
   "int_mul ==
 %(u::hollight.int) ua::hollight.int.
    mk_int (real_mul (dest_int u) (dest_int ua))"
@@ -5402,8 +5295,7 @@
    dest_int (int_mul x y) = real_mul (dest_int x) (dest_int y)"
   by (import hollight int_mul_th)
 
-constdefs
-  int_abs :: "hollight.int => hollight.int" 
+definition int_abs :: "hollight.int => hollight.int" where 
   "int_abs == %u::hollight.int. mk_int (real_abs (dest_int u))"
 
 lemma DEF_int_abs: "int_abs = (%u::hollight.int. mk_int (real_abs (dest_int u)))"
@@ -5412,8 +5304,7 @@
 lemma int_abs_th: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)"
   by (import hollight int_abs_th)
 
-constdefs
-  int_max :: "hollight.int => hollight.int => hollight.int" 
+definition int_max :: "hollight.int => hollight.int => hollight.int" where 
   "int_max ==
 %(u::hollight.int) ua::hollight.int.
    mk_int (real_max (dest_int u) (dest_int ua))"
@@ -5427,8 +5318,7 @@
    dest_int (int_max x y) = real_max (dest_int x) (dest_int y)"
   by (import hollight int_max_th)
 
-constdefs
-  int_min :: "hollight.int => hollight.int => hollight.int" 
+definition int_min :: "hollight.int => hollight.int => hollight.int" where 
   "int_min ==
 %(u::hollight.int) ua::hollight.int.
    mk_int (real_min (dest_int u) (dest_int ua))"
@@ -5442,8 +5332,7 @@
    dest_int (int_min x y) = real_min (dest_int x) (dest_int y)"
   by (import hollight int_min_th)
 
-constdefs
-  int_pow :: "hollight.int => nat => hollight.int" 
+definition int_pow :: "hollight.int => nat => hollight.int" where 
   "int_pow == %(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua)"
 
 lemma DEF_int_pow: "int_pow = (%(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua))"
@@ -5496,8 +5385,7 @@
    d ~= int_of_num 0 --> (EX c::hollight.int. int_lt x (int_mul c d))"
   by (import hollight INT_ARCH)
 
-constdefs
-  mod_int :: "hollight.int => hollight.int => hollight.int => bool" 
+definition mod_int :: "hollight.int => hollight.int => hollight.int => bool" where 
   "mod_int ==
 %(u::hollight.int) (ua::hollight.int) ub::hollight.int.
    EX q::hollight.int. int_sub ua ub = int_mul q u"
@@ -5507,8 +5395,7 @@
     EX q::hollight.int. int_sub ua ub = int_mul q u)"
   by (import hollight DEF_mod_int)
 
-constdefs
-  IN :: "'A => ('A => bool) => bool" 
+definition IN :: "'A => ('A => bool) => bool" where 
   "IN == %(u::'A::type) ua::'A::type => bool. ua u"
 
 lemma DEF_IN: "IN = (%(u::'A::type) ua::'A::type => bool. ua u)"
@@ -5518,15 +5405,13 @@
    (x = xa) = (ALL xb::'A::type. IN xb x = IN xb xa)"
   by (import hollight EXTENSION)
 
-constdefs
-  GSPEC :: "('A => bool) => 'A => bool" 
+definition GSPEC :: "('A => bool) => 'A => bool" where 
   "GSPEC == %u::'A::type => bool. u"
 
 lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)"
   by (import hollight DEF_GSPEC)
 
-constdefs
-  SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" 
+definition SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" where 
   "SETSPEC == %(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub"
 
 lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub)"
@@ -5548,15 +5433,13 @@
 (ALL (p::'q_37194::type => bool) x::'q_37194::type. IN x p = p x)"
   by (import hollight IN_ELIM_THM)
 
-constdefs
-  EMPTY :: "'A => bool" 
+definition EMPTY :: "'A => bool" where 
   "EMPTY == %x::'A::type. False"
 
 lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)"
   by (import hollight DEF_EMPTY)
 
-constdefs
-  INSERT :: "'A => ('A => bool) => 'A => bool" 
+definition INSERT :: "'A => ('A => bool) => 'A => bool" where 
   "INSERT == %(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u"
 
 lemma DEF_INSERT: "INSERT =
@@ -5585,8 +5468,7 @@
     GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x))"
   by (import hollight DEF_UNION)
 
-constdefs
-  UNIONS :: "(('A => bool) => bool) => 'A => bool" 
+definition UNIONS :: "(('A => bool) => bool) => 'A => bool" where 
   "UNIONS ==
 %u::('A::type => bool) => bool.
    GSPEC
@@ -5615,8 +5497,7 @@
     GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x))"
   by (import hollight DEF_INTER)
 
-constdefs
-  INTERS :: "(('A => bool) => bool) => 'A => bool" 
+definition INTERS :: "(('A => bool) => bool) => 'A => bool" where 
   "INTERS ==
 %u::('A::type => bool) => bool.
    GSPEC
@@ -5632,8 +5513,7 @@
             SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x))"
   by (import hollight DEF_INTERS)
 
-constdefs
-  DIFF :: "('A => bool) => ('A => bool) => 'A => bool" 
+definition DIFF :: "('A => bool) => ('A => bool) => 'A => bool" where 
   "DIFF ==
 %(u::'A::type => bool) ua::'A::type => bool.
    GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x)"
@@ -5648,8 +5528,7 @@
 GSPEC (%u::'A::type. EX y::'A::type. SETSPEC u (IN y s | y = x) y)"
   by (import hollight INSERT)
 
-constdefs
-  DELETE :: "('A => bool) => 'A => 'A => bool" 
+definition DELETE :: "('A => bool) => 'A => 'A => bool" where 
   "DELETE ==
 %(u::'A::type => bool) ua::'A::type.
    GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y)"
@@ -5659,8 +5538,7 @@
     GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y))"
   by (import hollight DEF_DELETE)
 
-constdefs
-  SUBSET :: "('A => bool) => ('A => bool) => bool" 
+definition SUBSET :: "('A => bool) => ('A => bool) => bool" where 
   "SUBSET ==
 %(u::'A::type => bool) ua::'A::type => bool.
    ALL x::'A::type. IN x u --> IN x ua"
@@ -5670,8 +5548,7 @@
     ALL x::'A::type. IN x u --> IN x ua)"
   by (import hollight DEF_SUBSET)
 
-constdefs
-  PSUBSET :: "('A => bool) => ('A => bool) => bool" 
+definition PSUBSET :: "('A => bool) => ('A => bool) => bool" where 
   "PSUBSET ==
 %(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua"
 
@@ -5679,8 +5556,7 @@
 (%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua)"
   by (import hollight DEF_PSUBSET)
 
-constdefs
-  DISJOINT :: "('A => bool) => ('A => bool) => bool" 
+definition DISJOINT :: "('A => bool) => ('A => bool) => bool" where 
   "DISJOINT ==
 %(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY"
 
@@ -5688,15 +5564,13 @@
 (%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY)"
   by (import hollight DEF_DISJOINT)
 
-constdefs
-  SING :: "('A => bool) => bool" 
+definition SING :: "('A => bool) => bool" where 
   "SING == %u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY"
 
 lemma DEF_SING: "SING = (%u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY)"
   by (import hollight DEF_SING)
 
-constdefs
-  FINITE :: "('A => bool) => bool" 
+definition FINITE :: "('A => bool) => bool" where 
   "FINITE ==
 %a::'A::type => bool.
    ALL FINITE'::('A::type => bool) => bool.
@@ -5718,15 +5592,13 @@
        FINITE' a)"
   by (import hollight DEF_FINITE)
 
-constdefs
-  INFINITE :: "('A => bool) => bool" 
+definition INFINITE :: "('A => bool) => bool" where 
   "INFINITE == %u::'A::type => bool. ~ FINITE u"
 
 lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)"
   by (import hollight DEF_INFINITE)
 
-constdefs
-  IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" 
+definition IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" where 
   "IMAGE ==
 %(u::'A::type => 'B::type) ua::'A::type => bool.
    GSPEC
@@ -5740,8 +5612,7 @@
          EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y))"
   by (import hollight DEF_IMAGE)
 
-constdefs
-  INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" 
+definition INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
   "INJ ==
 %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
    (ALL x::'A::type. IN x ua --> IN (u x) ub) &
@@ -5754,8 +5625,7 @@
         IN x ua & IN y ua & u x = u y --> x = y))"
   by (import hollight DEF_INJ)
 
-constdefs
-  SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" 
+definition SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
   "SURJ ==
 %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
    (ALL x::'A::type. IN x ua --> IN (u x) ub) &
@@ -5767,8 +5637,7 @@
     (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x)))"
   by (import hollight DEF_SURJ)
 
-constdefs
-  BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" 
+definition BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where 
   "BIJ ==
 %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool.
    INJ u ua ub & SURJ u ua ub"
@@ -5778,22 +5647,19 @@
     INJ u ua ub & SURJ u ua ub)"
   by (import hollight DEF_BIJ)
 
-constdefs
-  CHOICE :: "('A => bool) => 'A" 
+definition CHOICE :: "('A => bool) => 'A" where 
   "CHOICE == %u::'A::type => bool. SOME x::'A::type. IN x u"
 
 lemma DEF_CHOICE: "CHOICE = (%u::'A::type => bool. SOME x::'A::type. IN x u)"
   by (import hollight DEF_CHOICE)
 
-constdefs
-  REST :: "('A => bool) => 'A => bool" 
+definition REST :: "('A => bool) => 'A => bool" where 
   "REST == %u::'A::type => bool. DELETE u (CHOICE u)"
 
 lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))"
   by (import hollight DEF_REST)
 
-constdefs
-  CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" 
+definition CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" where 
   "CARD_GE ==
 %(u::'q_37693::type => bool) ua::'q_37690::type => bool.
    EX f::'q_37693::type => 'q_37690::type.
@@ -5807,8 +5673,7 @@
           IN y ua --> (EX x::'q_37693::type. IN x u & y = f x))"
   by (import hollight DEF_CARD_GE)
 
-constdefs
-  CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" 
+definition CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" where 
   "CARD_LE ==
 %(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u"
 
@@ -5816,8 +5681,7 @@
 (%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u)"
   by (import hollight DEF_CARD_LE)
 
-constdefs
-  CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" 
+definition CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" where 
   "CARD_EQ ==
 %(u::'q_37712::type => bool) ua::'q_37713::type => bool.
    CARD_LE u ua & CARD_LE ua u"
@@ -5827,8 +5691,7 @@
     CARD_LE u ua & CARD_LE ua u)"
   by (import hollight DEF_CARD_EQ)
 
-constdefs
-  CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" 
+definition CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" where 
   "CARD_GT ==
 %(u::'q_37727::type => bool) ua::'q_37728::type => bool.
    CARD_GE u ua & ~ CARD_GE ua u"
@@ -5838,8 +5701,7 @@
     CARD_GE u ua & ~ CARD_GE ua u)"
   by (import hollight DEF_CARD_GT)
 
-constdefs
-  CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" 
+definition CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" where 
   "CARD_LT ==
 %(u::'q_37743::type => bool) ua::'q_37744::type => bool.
    CARD_LE u ua & ~ CARD_LE ua u"
@@ -5849,8 +5711,7 @@
     CARD_LE u ua & ~ CARD_LE ua u)"
   by (import hollight DEF_CARD_LT)
 
-constdefs
-  COUNTABLE :: "('q_37757 => bool) => bool" 
+definition COUNTABLE :: "('q_37757 => bool) => bool" where 
   "(op ==::(('q_37757::type => bool) => bool)
         => (('q_37757::type => bool) => bool) => prop)
  (COUNTABLE::('q_37757::type => bool) => bool)
@@ -6470,9 +6331,8 @@
    FINITE s --> FINITE (DIFF s t)"
   by (import hollight FINITE_DIFF)
 
-constdefs
-  FINREC :: "('q_41824 => 'q_41823 => 'q_41823)
-=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" 
+definition FINREC :: "('q_41824 => 'q_41823 => 'q_41823)
+=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" where 
   "FINREC ==
 SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type)
              => 'q_41823::type
@@ -6558,9 +6418,8 @@
            FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (f x (g s))))"
   by (import hollight SET_RECURSION_LEMMA)
 
-constdefs
-  ITSET :: "('q_42525 => 'q_42524 => 'q_42524)
-=> ('q_42525 => bool) => 'q_42524 => 'q_42524" 
+definition ITSET :: "('q_42525 => 'q_42524 => 'q_42524)
+=> ('q_42525 => bool) => 'q_42524 => 'q_42524" where 
   "ITSET ==
 %(u::'q_42525::type => 'q_42524::type => 'q_42524::type)
    (ua::'q_42525::type => bool) ub::'q_42524::type.
@@ -6630,8 +6489,7 @@
           EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))"
   by (import hollight FINITE_RESTRICT)
 
-constdefs
-  CARD :: "('q_42918 => bool) => nat" 
+definition CARD :: "('q_42918 => bool) => nat" where 
   "CARD == %u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0"
 
 lemma DEF_CARD: "CARD = (%u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0)"
@@ -6674,8 +6532,7 @@
    CARD s + CARD t = CARD u"
   by (import hollight CARD_UNION_EQ)
 
-constdefs
-  HAS_SIZE :: "('q_43199 => bool) => nat => bool" 
+definition HAS_SIZE :: "('q_43199 => bool) => nat => bool" where 
   "HAS_SIZE == %(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua"
 
 lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua)"
@@ -6944,8 +6801,7 @@
        (ALL xa::'A::type. IN xa x --> (EX! m::nat. < m n & f m = xa)))"
   by (import hollight HAS_SIZE_INDEX)
 
-constdefs
-  set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" 
+definition set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" where 
   "set_of_list ==
 SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool.
    set_of_list NIL = EMPTY &
@@ -6959,8 +6815,7 @@
         set_of_list (CONS h t) = INSERT h (set_of_list t)))"
   by (import hollight DEF_set_of_list)
 
-constdefs
-  list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" 
+definition list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" where 
   "list_of_set ==
 %u::'q_45986::type => bool.
    SOME l::'q_45986::type hollight.list.
@@ -6999,8 +6854,7 @@
    hollight.UNION (set_of_list x) (set_of_list xa)"
   by (import hollight SET_OF_LIST_APPEND)
 
-constdefs
-  pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" 
+definition pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" where 
   "pairwise ==
 %(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool.
    ALL (x::'q_46198::type) y::'q_46198::type.
@@ -7012,8 +6866,7 @@
        IN x ua & IN y ua & x ~= y --> u x y)"
   by (import hollight DEF_pairwise)
 
-constdefs
-  PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" 
+definition PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" where 
   "PAIRWISE ==
 SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool)
                => 'q_46220::type hollight.list => bool.
@@ -7075,8 +6928,7 @@
        (EMPTY::'A::type => bool))))"
   by (import hollight FINITE_IMAGE_IMAGE)
 
-constdefs
-  dimindex :: "('A => bool) => nat" 
+definition dimindex :: "('A => bool) => nat" where 
   "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop)
  (dimindex::('A::type => bool) => nat)
  (%u::'A::type => bool.
@@ -7125,8 +6977,7 @@
    dimindex s = dimindex t"
   by (import hollight DIMINDEX_FINITE_IMAGE)
 
-constdefs
-  finite_index :: "nat => 'A" 
+definition finite_index :: "nat => 'A" where 
   "(op ==::(nat => 'A::type) => (nat => 'A::type) => prop)
  (finite_index::nat => 'A::type)
  ((Eps::((nat => 'A::type) => bool) => nat => 'A::type)
@@ -7287,8 +7138,7 @@
                       xa))))))"
   by (import hollight CART_EQ)
 
-constdefs
-  lambda :: "(nat => 'A) => ('A, 'B) cart" 
+definition lambda :: "(nat => 'A) => ('A, 'B) cart" where 
   "(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart)
         => ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop)
  (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
@@ -7388,8 +7238,7 @@
   [where a="a :: ('A, 'B) finite_sum" and r=r ,
    OF type_definition_finite_sum]
 
-constdefs
-  pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" 
+definition pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where 
   "(op ==::(('A::type, 'M::type) cart
          => ('A::type, 'N::type) cart
             => ('A::type, ('M::type, 'N::type) finite_sum) cart)
@@ -7439,8 +7288,7 @@
                  (hollight.UNIV::'M::type => bool))))))"
   by (import hollight DEF_pastecart)
 
-constdefs
-  fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" 
+definition fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where 
   "fstcart ==
 %u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u)"
 
@@ -7448,8 +7296,7 @@
 (%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u))"
   by (import hollight DEF_fstcart)
 
-constdefs
-  sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" 
+definition sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where 
   "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart
          => ('A::type, 'N::type) cart)
         => (('A::type, ('M::type, 'N::type) finite_sum) cart
@@ -7616,8 +7463,7 @@
    (EX xb::'q_48070::type => 'q_48091::type. x = xb o xa)"
   by (import hollight FUNCTION_FACTORS_LEFT)
 
-constdefs
-  dotdot :: "nat => nat => nat => bool" 
+definition dotdot :: "nat => nat => nat => bool" where 
   "dotdot ==
 %(u::nat) ua::nat.
    GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x)"
@@ -7718,8 +7564,7 @@
    SUBSET (dotdot m n) (dotdot p q) = (< n m | <= p m & <= n q)"
   by (import hollight SUBSET_NUMSEG)
 
-constdefs
-  neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" 
+definition neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" where 
   "neutral ==
 %u::'q_48985::type => 'q_48985::type => 'q_48985::type.
    SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y"
@@ -7729,8 +7574,7 @@
     SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y)"
   by (import hollight DEF_neutral)
 
-constdefs
-  monoidal :: "('A => 'A => 'A) => bool" 
+definition monoidal :: "('A => 'A => 'A) => bool" where 
   "monoidal ==
 %u::'A::type => 'A::type => 'A::type.
    (ALL (x::'A::type) y::'A::type. u x y = u y x) &
@@ -7746,8 +7590,7 @@
     (ALL x::'A::type. u (neutral u) x = x))"
   by (import hollight DEF_monoidal)
 
-constdefs
-  support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" 
+definition support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where 
   "support ==
 %(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type)
    ub::'A::type => bool.
@@ -7763,9 +7606,8 @@
          EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x))"
   by (import hollight DEF_support)
 
-constdefs
-  iterate :: "('q_49090 => 'q_49090 => 'q_49090)
-=> ('A => bool) => ('A => 'q_49090) => 'q_49090" 
+definition iterate :: "('q_49090 => 'q_49090 => 'q_49090)
+=> ('A => bool) => ('A => 'q_49090) => 'q_49090" where 
   "iterate ==
 %(u::'q_49090::type => 'q_49090::type => 'q_49090::type)
    (ua::'A::type => bool) ub::'A::type => 'q_49090::type.
@@ -8017,8 +7859,7 @@
        iterate u_4247 s f = iterate u_4247 t g)"
   by (import hollight ITERATE_EQ_GENERAL)
 
-constdefs
-  nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" 
+definition nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" where 
   "(op ==::(('q_51017::type => bool) => ('q_51017::type => nat) => nat)
         => (('q_51017::type => bool) => ('q_51017::type => nat) => nat)
            => prop)
@@ -8965,9 +8806,8 @@
    hollight.sum x xb = hollight.sum xa xc"
   by (import hollight SUM_EQ_GENERAL)
 
-constdefs
-  CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list
-=> 'q_57931 => 'q_57930 => 'q_57890" 
+definition CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list
+=> 'q_57931 => 'q_57930 => 'q_57890" where 
   "CASEWISE ==
 SOME CASEWISE::(('q_57926::type => 'q_57930::type) *
                 ('q_57931::type
@@ -9084,11 +8924,10 @@
     x"
   by (import hollight CASEWISE_WORKS)
 
-constdefs
-  admissible :: "('q_58228 => 'q_58221 => bool)
+definition admissible :: "('q_58228 => 'q_58221 => bool)
 => (('q_58228 => 'q_58224) => 'q_58234 => bool)
    => ('q_58234 => 'q_58221)
-      => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" 
+      => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" where 
   "admissible ==
 %(u::'q_58228::type => 'q_58221::type => bool)
    (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool)
@@ -9114,10 +8953,9 @@
        uc f a = uc g a)"
   by (import hollight DEF_admissible)
 
-constdefs
-  tailadmissible :: "('A => 'A => bool)
+definition tailadmissible :: "('A => 'A => bool)
 => (('A => 'B) => 'P => bool)
-   => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" 
+   => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where 
   "tailadmissible ==
 %(u::'A::type => 'A::type => bool)
    (ua::('A::type => 'B::type) => 'P::type => bool)
@@ -9151,11 +8989,10 @@
            ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a)))"
   by (import hollight DEF_tailadmissible)
 
-constdefs
-  superadmissible :: "('q_58378 => 'q_58378 => bool)
+definition superadmissible :: "('q_58378 => 'q_58378 => bool)
 => (('q_58378 => 'q_58380) => 'q_58386 => bool)
    => ('q_58386 => 'q_58378)
-      => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" 
+      => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" where 
   "superadmissible ==
 %(u::'q_58378::type => 'q_58378::type => bool)
    (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool)
--- a/src/HOL/Import/HOLLightCompat.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/HOLLightCompat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -30,8 +30,7 @@
     by simp
 qed
 
-constdefs
-   Pred :: "nat \<Rightarrow> nat"
+definition Pred :: "nat \<Rightarrow> nat" where
    "Pred n \<equiv> n - (Suc 0)"
 
 lemma Pred_altdef: "Pred = (SOME PRE. PRE 0 = 0 & (ALL n. PRE (Suc n) = n))"
@@ -84,8 +83,7 @@
   apply auto
   done
 
-constdefs
-  NUMERAL_BIT0 :: "nat \<Rightarrow> nat"
+definition NUMERAL_BIT0 :: "nat \<Rightarrow> nat" where
   "NUMERAL_BIT0 n \<equiv> n + n"
 
 lemma NUMERAL_BIT1_altdef: "NUMERAL_BIT1 n = Suc (n + n)"
--- a/src/HOL/Import/proof_kernel.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/proof_kernel.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -186,7 +186,7 @@
 
 fun mk_syn thy c =
   if Syntax.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
-  else Syntax.literal c
+  else Delimfix (Syntax.escape c)
 
 fun quotename c =
   if Syntax.is_identifier c andalso not (OuterKeyword.is_keyword c) then c else quote c
--- a/src/HOL/Import/shuffler.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Import/shuffler.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -351,7 +351,7 @@
 
 fun equals_fun thy assume t =
     case t of
-        Const("op ==",_) $ u $ v => if TermOrd.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
+        Const("op ==",_) $ u $ v => if Term_Ord.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
       | _ => NONE
 
 fun eta_expand thy assumes origt =
--- a/src/HOL/Induct/Tree.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Induct/Tree.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -13,20 +13,20 @@
     Atom 'a
   | Branch "nat => 'a tree"
 
-consts
+primrec
   map_tree :: "('a => 'b) => 'a tree => 'b tree"
-primrec
+where
   "map_tree f (Atom a) = Atom (f a)"
-  "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"
+| "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"
 
 lemma tree_map_compose: "map_tree g (map_tree f t) = map_tree (g \<circ> f) t"
   by (induct t) simp_all
 
-consts
+primrec
   exists_tree :: "('a => bool) => 'a tree => bool"
-primrec
+where
   "exists_tree P (Atom a) = P a"
-  "exists_tree P (Branch ts) = (\<exists>x. exists_tree P (ts x))"
+| "exists_tree P (Branch ts) = (\<exists>x. exists_tree P (ts x))"
 
 lemma exists_map:
   "(!!x. P x ==> Q (f x)) ==>
@@ -39,23 +39,23 @@
 datatype brouwer = Zero | Succ "brouwer" | Lim "nat => brouwer"
 
 text{*Addition of ordinals*}
-consts
+primrec
   add :: "[brouwer,brouwer] => brouwer"
-primrec
+where
   "add i Zero = i"
-  "add i (Succ j) = Succ (add i j)"
-  "add i (Lim f) = Lim (%n. add i (f n))"
+| "add i (Succ j) = Succ (add i j)"
+| "add i (Lim f) = Lim (%n. add i (f n))"
 
 lemma add_assoc: "add (add i j) k = add i (add j k)"
   by (induct k) auto
 
 text{*Multiplication of ordinals*}
-consts
+primrec
   mult :: "[brouwer,brouwer] => brouwer"
-primrec
+where
   "mult i Zero = Zero"
-  "mult i (Succ j) = add (mult i j) i"
-  "mult i (Lim f) = Lim (%n. mult i (f n))"
+| "mult i (Succ j) = add (mult i j) i"
+| "mult i (Lim f) = Lim (%n. mult i (f n))"
 
 lemma add_mult_distrib: "mult i (add j k) = add (mult i j) (mult i k)"
   by (induct k) (auto simp add: add_assoc)
@@ -68,7 +68,7 @@
 
 subsection{*A WF Ordering for The Brouwer ordinals (Michael Compton)*}
 
-text{*To define recdef style functions we need an ordering on the Brouwer
+text{*To use the function package we need an ordering on the Brouwer
   ordinals.  Start with a predecessor relation and form its transitive 
   closure. *} 
 
@@ -83,7 +83,7 @@
 lemma wf_brouwer_pred: "wf brouwer_pred"
   by(unfold wf_def brouwer_pred_def, clarify, induct_tac x, blast+)
 
-lemma wf_brouwer_order: "wf brouwer_order"
+lemma wf_brouwer_order[simp]: "wf brouwer_order"
   by(unfold brouwer_order_def, rule wf_trancl[OF wf_brouwer_pred])
 
 lemma [simp]: "(j, Succ j) : brouwer_order"
@@ -92,14 +92,16 @@
 lemma [simp]: "(f n, Lim f) : brouwer_order"
   by(auto simp add: brouwer_order_def brouwer_pred_def)
 
-text{*Example of a recdef*}
-consts
+text{*Example of a general function*}
+
+function
   add2 :: "(brouwer*brouwer) => brouwer"
-recdef add2 "inv_image brouwer_order (\<lambda> (x,y). y)"
+where
   "add2 (i, Zero) = i"
-  "add2 (i, (Succ j)) = Succ (add2 (i, j))"
-  "add2 (i, (Lim f)) = Lim (\<lambda> n. add2 (i, (f n)))"
-  (hints recdef_wf: wf_brouwer_order)
+| "add2 (i, (Succ j)) = Succ (add2 (i, j))"
+| "add2 (i, (Lim f)) = Lim (\<lambda> n. add2 (i, (f n)))"
+by pat_completeness auto
+termination by (relation "inv_image brouwer_order snd") auto
 
 lemma add2_assoc: "add2 (add2 (i, j), k) = add2 (i, add2 (j, k))"
   by (induct k) auto
--- a/src/HOL/IsaMakefile	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/IsaMakefile	Tue Mar 02 04:35:44 2010 -0800
@@ -352,7 +352,7 @@
   PReal.thy \
   Parity.thy \
   RComplete.thy \
-  Rational.thy \
+  Rat.thy \
   Real.thy \
   RealDef.thy \
   RealPow.thy \
--- a/src/HOL/Isar_Examples/Expr_Compiler.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Isar_Examples/Expr_Compiler.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -85,8 +85,7 @@
     | Apply f => exec instrs (f (hd stack) (hd (tl stack))
                    # (tl (tl stack))) env)"
 
-constdefs
-  execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val"
+definition execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val" where
   "execute instrs env == hd (exec instrs [] env)"
 
 
--- a/src/HOL/Isar_Examples/Hoare.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Isar_Examples/Hoare.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -55,14 +55,10 @@
     (if s : b then Sem c1 s s' else Sem c2 s s')"
   "Sem (While b x c) s s' = (EX n. iter n b (Sem c) s s')"
 
-constdefs
-  Valid :: "'a bexp => 'a com => 'a bexp => bool"
-    ("(3|- _/ (2_)/ _)" [100, 55, 100] 50)
+definition Valid :: "'a bexp => 'a com => 'a bexp => bool" ("(3|- _/ (2_)/ _)" [100, 55, 100] 50) where
   "|- P c Q == ALL s s'. Sem c s s' --> s : P --> s' : Q"
 
-syntax (xsymbols)
-  Valid :: "'a bexp => 'a com => 'a bexp => bool"
-    ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
+notation (xsymbols) Valid ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
 
 lemma ValidI [intro?]:
     "(!!s s'. Sem c s s' ==> s : P ==> s' : Q) ==> |- P c Q"
--- a/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -60,8 +60,7 @@
 
 subsection {* Basic properties of ``below'' *}
 
-constdefs
-  below :: "nat => nat set"
+definition below :: "nat => nat set" where
   "below n == {i. i < n}"
 
 lemma below_less_iff [iff]: "(i: below k) = (i < k)"
@@ -84,8 +83,7 @@
 
 subsection {* Basic properties of ``evnodd'' *}
 
-constdefs
-  evnodd :: "(nat * nat) set => nat => (nat * nat) set"
+definition evnodd :: "(nat * nat) set => nat => (nat * nat) set" where
   "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}"
 
 lemma evnodd_iff:
@@ -247,8 +245,7 @@
 
 subsection {* Main theorem *}
 
-constdefs
-  mutilated_board :: "nat => nat => (nat * nat) set"
+definition mutilated_board :: "nat => nat => (nat * nat) set" where
   "mutilated_board m n ==
     below (2 * (m + 1)) <*> below (2 * (n + 1))
       - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}"
--- a/src/HOL/Lambda/ParRed.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Lambda/ParRed.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -85,14 +85,14 @@
 
 subsection {* Complete developments *}
 
-consts
+fun
   "cd" :: "dB => dB"
-recdef "cd" "measure size"
+where
   "cd (Var n) = Var n"
-  "cd (Var n \<degree> t) = Var n \<degree> cd t"
-  "cd ((s1 \<degree> s2) \<degree> t) = cd (s1 \<degree> s2) \<degree> cd t"
-  "cd (Abs u \<degree> t) = (cd u)[cd t/0]"
-  "cd (Abs s) = Abs (cd s)"
+| "cd (Var n \<degree> t) = Var n \<degree> cd t"
+| "cd ((s1 \<degree> s2) \<degree> t) = cd (s1 \<degree> s2) \<degree> cd t"
+| "cd (Abs u \<degree> t) = (cd u)[cd t/0]"
+| "cd (Abs s) = Abs (cd s)"
 
 lemma par_beta_cd: "s => t \<Longrightarrow> t => cd s"
   apply (induct s arbitrary: t rule: cd.induct)
--- a/src/HOL/Library/Binomial.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Binomial.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Binomial.thy
+(*  Title:      HOL/Library/Binomial.thy
     Author:     Lawrence C Paulson, Amine Chaieb
     Copyright   1997  University of Cambridge
 *)
@@ -6,7 +6,7 @@
 header {* Binomial Coefficients *}
 
 theory Binomial
-imports Fact SetInterval Presburger Main Rational
+imports Complex_Main
 begin
 
 text {* This development is based on the work of Andy Gordon and
--- a/src/HOL/Library/Countable.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Countable.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -5,12 +5,7 @@
 header {* Encoding (almost) everything into natural numbers *}
 
 theory Countable
-imports
-  "~~/src/HOL/List"
-  "~~/src/HOL/Hilbert_Choice"
-  "~~/src/HOL/Nat_Int_Bij"
-  "~~/src/HOL/Rational"
-  Main
+imports Main Rat Nat_Int_Bij
 begin
 
 subsection {* The class of countable types *}
@@ -213,8 +208,8 @@
 proof
   fix r::rat
   show "\<exists>n. r = nat_to_rat_surj n"
-  proof(cases r)
-    fix i j assume [simp]: "r = Fract i j" and "j \<noteq> 0"
+  proof (cases r)
+    fix i j assume [simp]: "r = Fract i j" and "j > 0"
     have "r = (let m = inv nat_to_int_bij i; n = inv nat_to_int_bij j
                in nat_to_rat_surj(nat2_to_nat (m,n)))"
       using nat2_to_nat_inj surj_f_inv_f[OF surj_nat_to_int_bij]
--- a/src/HOL/Library/Fraction_Field.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Fraction_Field.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,12 +1,12 @@
-(*  Title:      Fraction_Field.thy
+(*  Title:      HOL/Library/Fraction_Field.thy
     Author:     Amine Chaieb, University of Cambridge
 *)
 
 header{* A formalization of the fraction field of any integral domain 
-         A generalization of Rational.thy from int to any integral domain *}
+         A generalization of Rat.thy from int to any integral domain *}
 
 theory Fraction_Field
-imports Main (* Equiv_Relations Plain *)
+imports Main
 begin
 
 subsection {* General fractions construction *}
--- a/src/HOL/Library/Multiset.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Multiset.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1502,13 +1502,13 @@
 by (cases M) auto
 
 syntax
-  comprehension1_mset :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"
+  "_comprehension1_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"
       ("({#_/. _ :# _#})")
 translations
   "{#e. x:#M#}" == "CONST image_mset (%x. e) M"
 
 syntax
-  comprehension2_mset :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
+  "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
       ("({#_/ | _ :# _./ _#})")
 translations
   "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
@@ -1631,9 +1631,9 @@
 
 setup {*
 let
-  fun msetT T = Type ("Multiset.multiset", [T]);
+  fun msetT T = Type (@{type_name multiset}, [T]);
 
-  fun mk_mset T [] = Const (@{const_name Mempty}, msetT T)
+  fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T)
     | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x
     | mk_mset T (x :: xs) =
           Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $
@@ -1649,7 +1649,7 @@
       rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
 
   val regroup_munion_conv =
-      Function_Lib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
+      Function_Lib.regroup_conv @{const_abbrev Mempty} @{const_name plus}
         (map (fn t => t RS eq_reflection) (@{thms add_ac} @ @{thms empty_idemp}))
 
   fun unfold_pwleq_tac i =
--- a/src/HOL/Library/Numeral_Type.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Numeral_Type.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -167,7 +167,7 @@
 begin
 
 lemma size0: "0 < n"
-by (cut_tac size1, simp)
+using size1 by simp
 
 lemmas definitions =
   zero_def one_def add_def mult_def minus_def diff_def
@@ -384,23 +384,18 @@
   "_NumeralType1" :: type ("1")
 
 translations
-  "_NumeralType1" == (type) "num1"
-  "_NumeralType0" == (type) "num0"
+  (type) "1" == (type) "num1"
+  (type) "0" == (type) "num0"
 
 parse_translation {*
 let
-(* FIXME @{type_syntax} *)
-val num1_const = Syntax.const "Numeral_Type.num1";
-val num0_const = Syntax.const "Numeral_Type.num0";
-val B0_const = Syntax.const "Numeral_Type.bit0";
-val B1_const = Syntax.const "Numeral_Type.bit1";
-
 fun mk_bintype n =
   let
-    fun mk_bit n = if n = 0 then B0_const else B1_const;
+    fun mk_bit 0 = Syntax.const @{type_syntax bit0}
+      | mk_bit 1 = Syntax.const @{type_syntax bit1};
     fun bin_of n =
-      if n = 1 then num1_const
-      else if n = 0 then num0_const
+      if n = 1 then Syntax.const @{type_syntax num1}
+      else if n = 0 then Syntax.const @{type_syntax num0}
       else if n = ~1 then raise TERM ("negative type numeral", [])
       else
         let val (q, r) = Integer.div_mod n 2;
@@ -419,25 +414,22 @@
 fun int_of [] = 0
   | int_of (b :: bs) = b + 2 * int_of bs;
 
-(* FIXME @{type_syntax} *)
-fun bin_of (Const ("num0", _)) = []
-  | bin_of (Const ("num1", _)) = [1]
-  | bin_of (Const ("bit0", _) $ bs) = 0 :: bin_of bs
-  | bin_of (Const ("bit1", _) $ bs) = 1 :: bin_of bs
-  | bin_of t = raise TERM("bin_of", [t]);
+fun bin_of (Const (@{type_syntax num0}, _)) = []
+  | bin_of (Const (@{type_syntax num1}, _)) = [1]
+  | bin_of (Const (@{type_syntax bit0}, _) $ bs) = 0 :: bin_of bs
+  | bin_of (Const (@{type_syntax bit1}, _) $ bs) = 1 :: bin_of bs
+  | bin_of t = raise TERM ("bin_of", [t]);
 
 fun bit_tr' b [t] =
-  let
-    val rev_digs = b :: bin_of t handle TERM _ => raise Match
-    val i = int_of rev_digs;
-    val num = string_of_int (abs i);
-  in
-    Syntax.const "_NumeralType" $ Syntax.free num
-  end
+      let
+        val rev_digs = b :: bin_of t handle TERM _ => raise Match
+        val i = int_of rev_digs;
+        val num = string_of_int (abs i);
+      in
+        Syntax.const @{syntax_const "_NumeralType"} $ Syntax.free num
+      end
   | bit_tr' b _ = raise Match;
-
-(* FIXME @{type_syntax} *)
-in [("bit0", bit_tr' 0), ("bit1", bit_tr' 1)] end;
+in [(@{type_syntax bit0}, bit_tr' 0), (@{type_syntax bit1}, bit_tr' 1)] end;
 *}
 
 subsection {* Examples *}
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -1189,7 +1189,7 @@
 local
   open Conv
   val concl = Thm.dest_arg o cprop_of
-  fun simple_cterm_ord t u = TermOrd.fast_term_ord (term_of t, term_of u) = LESS
+  fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
 in
   (* FIXME: Replace tryfind by get_first !! *)
 fun real_nonlinear_prover proof_method ctxt =
@@ -1279,7 +1279,7 @@
 
 local
  open Conv
-  fun simple_cterm_ord t u = TermOrd.fast_term_ord (term_of t, term_of u) = LESS
+  fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
  val concl = Thm.dest_arg o cprop_of
  val shuffle1 =
    fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
--- a/src/HOL/Library/Word.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/Word.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -311,11 +311,11 @@
 lemma norm_unsigned_idem [simp]: "norm_unsigned (norm_unsigned w) = norm_unsigned w"
   by (rule bit_list_induct [of _ w],simp_all)
 
-consts
+fun
   nat_to_bv_helper :: "nat => bit list => bit list"
-recdef nat_to_bv_helper "measure (\<lambda>n. n)"
-  "nat_to_bv_helper n = (%bs. (if n = 0 then bs
-                               else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \<zero> else \<one>)#bs)))"
+where
+  "nat_to_bv_helper n bs = (if n = 0 then bs
+                               else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \<zero> else \<one>)#bs))"
 
 definition
   nat_to_bv :: "nat => bit list" where
--- a/src/HOL/Library/normarith.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/normarith.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -216,7 +216,7 @@
       let 
        val l = headvector lt 
        val r = headvector rt
-      in (case TermOrd.fast_term_ord (l,r) of
+      in (case Term_Ord.fast_term_ord (l,r) of
          LESS => (apply_pthb then_conv arg_conv vector_add_conv 
                   then_conv apply_pthd) ct
         | GREATER => (apply_pthc then_conv arg_conv vector_add_conv 
@@ -378,7 +378,7 @@
 local
  val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
  fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
- fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
+ fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
   (* FIXME: Lookup in the context every time!!! Fix this !!!*)
  fun splitequation ctxt th acc =
   let 
--- a/src/HOL/Library/positivstellensatz.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Library/positivstellensatz.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -61,9 +61,9 @@
 structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
 structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
 structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
-structure Termfunc = FuncFun(type key = term val ord = TermOrd.fast_term_ord);
+structure Termfunc = FuncFun(type key = term val ord = Term_Ord.fast_term_ord);
 
-val cterm_ord = TermOrd.fast_term_ord o pairself term_of
+val cterm_ord = Term_Ord.fast_term_ord o pairself term_of
 
 structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);
 
@@ -745,7 +745,7 @@
 
 fun gen_prover_real_arith ctxt prover = 
  let
-  fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS
+  fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS
   val {add,mul,neg,pow,sub,main} = 
      Normalizer.semiring_normalizers_ord_wrapper ctxt
       (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
--- a/src/HOL/Matrix/ComputeNumeral.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Matrix/ComputeNumeral.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -20,7 +20,7 @@
 lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
 
 (* lezero for bit strings *)
-constdefs
+definition
   "lezero x == (x \<le> 0)"
 lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
 lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
@@ -60,7 +60,7 @@
 
 lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul 
 
-constdefs 
+definition
   "nat_norm_number_of (x::nat) == x"
 
 lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
--- a/src/HOL/Matrix/Matrix.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Matrix/Matrix.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -24,10 +24,10 @@
 apply (rule Abs_matrix_induct)
 by (simp add: Abs_matrix_inverse matrix_def)
 
-constdefs
-  nrows :: "('a::zero) matrix \<Rightarrow> nat"
+definition nrows :: "('a::zero) matrix \<Rightarrow> nat" where
   "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))"
-  ncols :: "('a::zero) matrix \<Rightarrow> nat"
+
+definition ncols :: "('a::zero) matrix \<Rightarrow> nat" where
   "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))"
 
 lemma nrows:
@@ -50,10 +50,10 @@
   thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect)
 qed
 
-constdefs
-  transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix"
+definition transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix" where
   "transpose_infmatrix A j i == A i j"
-  transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix"
+
+definition transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix" where
   "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix"
 
 declare transpose_infmatrix_def[simp]
@@ -256,14 +256,16 @@
   ultimately show "finite ?u" by (rule finite_subset)
 qed
 
-constdefs
-  apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix"
+definition apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix" where
   "apply_infmatrix f == % A. (% j i. f (A j i))"
-  apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix"
+
+definition apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix" where
   "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))"
-  combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix"
+
+definition combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix" where
   "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))"
-  combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix"
+
+definition combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix" where
   "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))"
 
 lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)"
@@ -272,10 +274,10 @@
 lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)"
 by (simp add: combine_infmatrix_def)
 
-constdefs
-commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
+definition commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool" where
 "commutative f == ! x y. f x y = f y x"
-associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
 "associative f == ! x y z. f (f x y) z = f x (f y z)"
 
 text{*
@@ -356,12 +358,13 @@
 lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q"
   by (simp add: ncols_le)
 
-constdefs
-  zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool"
+definition zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool" where
   "zero_r_neutral f == ! a. f a 0 = a"
-  zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
+
+definition zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
   "zero_l_neutral f == ! a. f 0 a = a"
-  zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool"
+
+definition zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool" where
   "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)"
 
 consts foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
@@ -648,10 +651,10 @@
   then show ?concl by simp
 qed
 
-constdefs
-  mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
+definition mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix" where
   "mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)"
-  mult_matrix :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
+
+definition mult_matrix :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix" where
   "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B"
 
 lemma mult_matrix_n:
@@ -673,12 +676,13 @@
   finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp
 qed
 
-constdefs
-  r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
+definition r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
   "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)"
-  l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
   "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"
-  distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
   "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"
 
 lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> a <= max x y" by (arith)
@@ -835,20 +839,22 @@
   apply (simp add: apply_matrix_def apply_infmatrix_def)
   by (simp add: zero_matrix_def)
 
-constdefs
-  singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix"
+definition singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix" where
   "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)"
-  move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix"
+
+definition move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix" where
   "move_matrix A y x == Abs_matrix(% j i. if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))"
-  take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
   "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)"
-  take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
   "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)"
 
-constdefs
-  column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
+definition column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
   "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1"
-  row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
   "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1"
 
 lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)"
@@ -1042,10 +1048,10 @@
   with contraprems show "False" by simp
 qed
 
-constdefs
-  foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
+definition foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
   "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m"
-  foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
+
+definition foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
   "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m"
 
 lemma foldmatrix_transpose:
@@ -1691,12 +1697,13 @@
 lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)"
 by (simp add: minus_matrix_def transpose_apply_matrix)
 
-constdefs 
-  right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+definition right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
   "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \<and> nrows X \<le> ncols A" 
-  left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+
+definition left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
   "left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \<and> ncols X \<le> nrows A" 
-  inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+
+definition inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
   "inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)"
 
 lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"
@@ -1781,8 +1788,7 @@
 lemma move_matrix_mult: "move_matrix ((A::('a::semiring_0) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)"
 by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult)
 
-constdefs
-  scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix"
+definition scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix" where
   "scalar_mult a m == apply_matrix (op * a) m"
 
 lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" 
--- a/src/HOL/Matrix/SparseMatrix.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Matrix/SparseMatrix.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -552,8 +552,7 @@
   else if j < i then (le_spvec [] b & le_spmat ((i,a)#as) bs)
   else (le_spvec a b & le_spmat as bs))"
 
-constdefs
-  disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+definition disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
   "disj_matrices A B == (! j i. (Rep_matrix A j i \<noteq> 0) \<longrightarrow> (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \<noteq> 0) \<longrightarrow> (Rep_matrix A j i = 0))"  
 
 declare [[simp_depth_limit = 6]]
@@ -802,8 +801,7 @@
   apply (simp_all add: sorted_spvec_abs_spvec)
   done
 
-constdefs
-  diff_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
+definition diff_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat" where
   "diff_spmat A B == add_spmat A (minus_spmat B)"
 
 lemma sorted_spmat_diff_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> sorted_spmat (diff_spmat A B)"
@@ -815,8 +813,7 @@
 lemma sparse_row_diff_spmat: "sparse_row_matrix (diff_spmat A B ) = (sparse_row_matrix A) - (sparse_row_matrix B)"
   by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus)
 
-constdefs
-  sorted_sparse_matrix :: "'a spmat \<Rightarrow> bool"
+definition sorted_sparse_matrix :: "'a spmat \<Rightarrow> bool" where
   "sorted_sparse_matrix A == (sorted_spvec A) & (sorted_spmat A)"
 
 lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \<Longrightarrow> sorted_spvec A"
@@ -1014,8 +1011,7 @@
   apply (simp_all add: sorted_nprt_spvec)
   done
 
-constdefs
-  mult_est_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
+definition mult_est_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat" where
   "mult_est_spmat r1 r2 s1 s2 == 
   add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) 
   (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))"  
--- a/src/HOL/Metis_Examples/BigO.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Metis_Examples/BigO.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1099,9 +1099,7 @@
 
 subsection {* Less than or equal to *}
 
-constdefs 
-  lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)"
-      (infixl "<o" 70)
+definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
   "f <o g == (%x. max (f x - g x) 0)"
 
 lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==>
--- a/src/HOL/Metis_Examples/Message.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Metis_Examples/Message.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -26,8 +26,7 @@
 text{*The inverse of a symmetric key is itself; that of a public key
       is the private key and vice versa*}
 
-constdefs
-  symKeys :: "key set"
+definition symKeys :: "key set" where
   "symKeys == {K. invKey K = K}"
 
 datatype  --{*We allow any number of friendly agents*}
@@ -55,12 +54,11 @@
   "{|x, y|}"      == "CONST MPair x y"
 
 
-constdefs
-  HPair :: "[msg,msg] => msg"                       ("(4Hash[_] /_)" [0, 1000])
+definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
     --{*Message Y paired with a MAC computed with the help of X*}
     "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
 
-  keysFor :: "msg set => key set"
+definition keysFor :: "msg set => key set" where
     --{*Keys useful to decrypt elements of a message set*}
   "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
 
--- a/src/HOL/Metis_Examples/Tarski.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Metis_Examples/Tarski.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -20,59 +20,56 @@
   pset  :: "'a set"
   order :: "('a * 'a) set"
 
-constdefs
-  monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
+definition monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" where
   "monotone f A r == \<forall>x\<in>A. \<forall>y\<in>A. (x, y): r --> ((f x), (f y)) : r"
 
-  least :: "['a => bool, 'a potype] => 'a"
+definition least :: "['a => bool, 'a potype] => 'a" where
   "least P po == @ x. x: pset po & P x &
                        (\<forall>y \<in> pset po. P y --> (x,y): order po)"
 
-  greatest :: "['a => bool, 'a potype] => 'a"
+definition greatest :: "['a => bool, 'a potype] => 'a" where
   "greatest P po == @ x. x: pset po & P x &
                           (\<forall>y \<in> pset po. P y --> (y,x): order po)"
 
-  lub  :: "['a set, 'a potype] => 'a"
+definition lub  :: "['a set, 'a potype] => 'a" where
   "lub S po == least (%x. \<forall>y\<in>S. (y,x): order po) po"
 
-  glb  :: "['a set, 'a potype] => 'a"
+definition glb  :: "['a set, 'a potype] => 'a" where
   "glb S po == greatest (%x. \<forall>y\<in>S. (x,y): order po) po"
 
-  isLub :: "['a set, 'a potype, 'a] => bool"
+definition isLub :: "['a set, 'a potype, 'a] => bool" where
   "isLub S po == %L. (L: pset po & (\<forall>y\<in>S. (y,L): order po) &
                    (\<forall>z\<in>pset po. (\<forall>y\<in>S. (y,z): order po) --> (L,z): order po))"
 
-  isGlb :: "['a set, 'a potype, 'a] => bool"
+definition isGlb :: "['a set, 'a potype, 'a] => bool" where
   "isGlb S po == %G. (G: pset po & (\<forall>y\<in>S. (G,y): order po) &
                  (\<forall>z \<in> pset po. (\<forall>y\<in>S. (z,y): order po) --> (z,G): order po))"
 
-  "fix"    :: "[('a => 'a), 'a set] => 'a set"
+definition "fix"    :: "[('a => 'a), 'a set] => 'a set" where
   "fix f A  == {x. x: A & f x = x}"
 
-  interval :: "[('a*'a) set,'a, 'a ] => 'a set"
+definition interval :: "[('a*'a) set,'a, 'a ] => 'a set" where
   "interval r a b == {x. (a,x): r & (x,b): r}"
 
-constdefs
-  Bot :: "'a potype => 'a"
+definition Bot :: "'a potype => 'a" where
   "Bot po == least (%x. True) po"
 
-  Top :: "'a potype => 'a"
+definition Top :: "'a potype => 'a" where
   "Top po == greatest (%x. True) po"
 
-  PartialOrder :: "('a potype) set"
+definition PartialOrder :: "('a potype) set" where
   "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) &
                        trans (order P)}"
 
-  CompleteLattice :: "('a potype) set"
+definition CompleteLattice :: "('a potype) set" where
   "CompleteLattice == {cl. cl: PartialOrder &
                         (\<forall>S. S \<subseteq> pset cl --> (\<exists>L. isLub S cl L)) &
                         (\<forall>S. S \<subseteq> pset cl --> (\<exists>G. isGlb S cl G))}"
 
-  induced :: "['a set, ('a * 'a) set] => ('a *'a)set"
+definition induced :: "['a set, ('a * 'a) set] => ('a *'a)set" where
   "induced A r == {(a,b). a : A & b: A & (a,b): r}"
 
-constdefs
-  sublattice :: "('a potype * 'a set)set"
+definition sublattice :: "('a potype * 'a set)set" where
   "sublattice ==
       SIGMA cl: CompleteLattice.
           {S. S \<subseteq> pset cl &
@@ -82,8 +79,7 @@
   sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50)
   where "S <<= cl \<equiv> S : sublattice `` {cl}"
 
-constdefs
-  dual :: "'a potype => 'a potype"
+definition dual :: "'a potype => 'a potype" where
   "dual po == (| pset = pset po, order = converse (order po) |)"
 
 locale PO =
--- a/src/HOL/MicroJava/BV/Altern.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/Altern.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,19 +8,18 @@
 imports BVSpec
 begin
 
-constdefs
-  check_type :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state \<Rightarrow> bool"
+definition check_type :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state \<Rightarrow> bool" where
   "check_type G mxs mxr s \<equiv> s \<in> states G mxs mxr"
 
-  wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count,
-                exception_table,p_count] \<Rightarrow> bool"
+definition wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count,
+                exception_table,p_count] \<Rightarrow> bool" where
   "wt_instr_altern i G rT phi mxs mxr max_pc et pc \<equiv>
   app i G mxs rT pc et (phi!pc) \<and>
   check_type G mxs mxr (OK (phi!pc)) \<and>
   (\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')"
 
-  wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
-                 exception_table,method_type] \<Rightarrow> bool"
+definition wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
+                 exception_table,method_type] \<Rightarrow> bool" where
   "wt_method_altern G C pTs rT mxs mxl ins et phi \<equiv>
   let max_pc = length ins in
   0 < max_pc \<and> 
--- a/src/HOL/MicroJava/BV/BVExample.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/BVExample.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -167,8 +167,7 @@
 
   @{prop [display] "P n"} 
 *}
-constdefs 
-  intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')")
+definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')") where
   "x \<in> [a, b) \<equiv> a \<le> x \<and> x < b"
 
 lemma pc_0: "x < n \<Longrightarrow> (x \<in> [0, n) \<Longrightarrow> P x) \<Longrightarrow> P x"
@@ -238,8 +237,7 @@
 lemmas eff_simps [simp] = eff_def norm_eff_def xcpt_eff_def
 declare appInvoke [simp del]
 
-constdefs
-  phi_append :: method_type ("\<phi>\<^sub>a")
+definition phi_append :: method_type ("\<phi>\<^sub>a") where
   "\<phi>\<^sub>a \<equiv> map (\<lambda>(x,y). Some (x, map OK y)) [ 
    (                                    [], [Class list_name, Class list_name]),
    (                     [Class list_name], [Class list_name, Class list_name]),
@@ -301,8 +299,7 @@
 abbreviation Ctest :: ty
   where "Ctest == Class test_name"
 
-constdefs
-  phi_makelist :: method_type ("\<phi>\<^sub>m")
+definition phi_makelist :: method_type ("\<phi>\<^sub>m") where
   "\<phi>\<^sub>m \<equiv> map (\<lambda>(x,y). Some (x, y)) [ 
     (                                   [], [OK Ctest, Err     , Err     ]),
     (                              [Clist], [OK Ctest, Err     , Err     ]),
@@ -376,8 +373,7 @@
   done
 
 text {* The whole program is welltyped: *}
-constdefs 
-  Phi :: prog_type ("\<Phi>")
+definition Phi :: prog_type ("\<Phi>") where
   "\<Phi> C sg \<equiv> if C = test_name \<and> sg = (makelist_name, []) then \<phi>\<^sub>m else          
              if C = list_name \<and> sg = (append_name, [Class list_name]) then \<phi>\<^sub>a else []"
 
--- a/src/HOL/MicroJava/BV/Correct.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/Correct.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,47 +9,40 @@
 imports BVSpec "../JVM/JVMExec"
 begin
 
-constdefs
-  approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool"
+definition approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool" where
   "approx_val G h v any == case any of Err \<Rightarrow> True | OK T \<Rightarrow> G,h\<turnstile>v::\<preceq>T"
 
-  approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \<Rightarrow> bool"
+definition approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \<Rightarrow> bool" where
   "approx_loc G hp loc LT == list_all2 (approx_val G hp) loc LT"
 
-  approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \<Rightarrow> bool"
+definition approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \<Rightarrow> bool" where
   "approx_stk G hp stk ST == approx_loc G hp stk (map OK ST)"
 
-  correct_frame  :: "[jvm_prog,aheap,state_type,nat,bytecode] \<Rightarrow> frame \<Rightarrow> bool"
+definition correct_frame  :: "[jvm_prog,aheap,state_type,nat,bytecode] \<Rightarrow> frame \<Rightarrow> bool" where
   "correct_frame G hp == \<lambda>(ST,LT) maxl ins (stk,loc,C,sig,pc).
                          approx_stk G hp stk ST  \<and> approx_loc G hp loc LT \<and> 
                          pc < length ins \<and> length loc=length(snd sig)+maxl+1"
 
-
-consts
- correct_frames  :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \<Rightarrow> bool"
-primrec
-"correct_frames G hp phi rT0 sig0 [] = True"
+primrec correct_frames  :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \<Rightarrow> bool" where
+  "correct_frames G hp phi rT0 sig0 [] = True"
+| "correct_frames G hp phi rT0 sig0 (f#frs) =
+    (let (stk,loc,C,sig,pc) = f in
+    (\<exists>ST LT rT maxs maxl ins et.
+      phi C sig ! pc = Some (ST,LT) \<and> is_class G C \<and> 
+      method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>
+    (\<exists>C' mn pTs. ins!pc = (Invoke C' mn pTs) \<and> 
+           (mn,pTs) = sig0 \<and> 
+           (\<exists>apTs D ST' LT'.
+           (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \<and>
+           length apTs = length pTs \<and>
+           (\<exists>D' rT' maxs' maxl' ins' et'.
+             method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \<and>
+             G \<turnstile> rT0 \<preceq> rT') \<and>
+     correct_frame G hp (ST, LT) maxl ins f \<and> 
+     correct_frames G hp phi rT sig frs))))"
 
-"correct_frames G hp phi rT0 sig0 (f#frs) =
-  (let (stk,loc,C,sig,pc) = f in
-  (\<exists>ST LT rT maxs maxl ins et.
-    phi C sig ! pc = Some (ST,LT) \<and> is_class G C \<and> 
-    method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>
-  (\<exists>C' mn pTs. ins!pc = (Invoke C' mn pTs) \<and> 
-         (mn,pTs) = sig0 \<and> 
-         (\<exists>apTs D ST' LT'.
-         (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \<and>
-         length apTs = length pTs \<and>
-         (\<exists>D' rT' maxs' maxl' ins' et'.
-           method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \<and>
-           G \<turnstile> rT0 \<preceq> rT') \<and>
-   correct_frame G hp (ST, LT) maxl ins f \<and> 
-   correct_frames G hp phi rT sig frs))))"
-
-
-constdefs
- correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"
-                  ("_,_ |-JVM _ [ok]"  [51,51] 50)
+definition correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"
+                  ("_,_ |-JVM _ [ok]"  [51,51] 50) where
 "correct_state G phi == \<lambda>(xp,hp,frs).
    case xp of
      None \<Rightarrow> (case frs of
@@ -66,9 +59,8 @@
    | Some x \<Rightarrow> frs = []" 
 
 
-syntax (xsymbols)
- correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"
-                  ("_,_ \<turnstile>JVM _ \<surd>"  [51,51] 50)
+notation (xsymbols)
+ correct_state  ("_,_ \<turnstile>JVM _ \<surd>"  [51,51] 50)
 
 
 lemma sup_ty_opt_OK:
--- a/src/HOL/MicroJava/BV/Effect.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/Effect.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -13,106 +13,98 @@
   succ_type = "(p_count \<times> state_type option) list"
 
 text {* Program counter of successor instructions: *}
-consts
-  succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list"
-primrec 
+primrec succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list" where
   "succs (Load idx) pc         = [pc+1]"
-  "succs (Store idx) pc        = [pc+1]"
-  "succs (LitPush v) pc        = [pc+1]"
-  "succs (Getfield F C) pc     = [pc+1]"
-  "succs (Putfield F C) pc     = [pc+1]"
-  "succs (New C) pc            = [pc+1]"
-  "succs (Checkcast C) pc      = [pc+1]"
-  "succs Pop pc                = [pc+1]"
-  "succs Dup pc                = [pc+1]"
-  "succs Dup_x1 pc             = [pc+1]"
-  "succs Dup_x2 pc             = [pc+1]"
-  "succs Swap pc               = [pc+1]"
-  "succs IAdd pc               = [pc+1]"
-  "succs (Ifcmpeq b) pc        = [pc+1, nat (int pc + b)]"
-  "succs (Goto b) pc           = [nat (int pc + b)]"
-  "succs Return pc             = [pc]"  
-  "succs (Invoke C mn fpTs) pc = [pc+1]"
-  "succs Throw pc              = [pc]"
+| "succs (Store idx) pc        = [pc+1]"
+| "succs (LitPush v) pc        = [pc+1]"
+| "succs (Getfield F C) pc     = [pc+1]"
+| "succs (Putfield F C) pc     = [pc+1]"
+| "succs (New C) pc            = [pc+1]"
+| "succs (Checkcast C) pc      = [pc+1]"
+| "succs Pop pc                = [pc+1]"
+| "succs Dup pc                = [pc+1]"
+| "succs Dup_x1 pc             = [pc+1]"
+| "succs Dup_x2 pc             = [pc+1]"
+| "succs Swap pc               = [pc+1]"
+| "succs IAdd pc               = [pc+1]"
+| "succs (Ifcmpeq b) pc        = [pc+1, nat (int pc + b)]"
+| "succs (Goto b) pc           = [nat (int pc + b)]"
+| "succs Return pc             = [pc]"  
+| "succs (Invoke C mn fpTs) pc = [pc+1]"
+| "succs Throw pc              = [pc]"
 
 text "Effect of instruction on the state type:"
-consts 
-eff' :: "instr \<times> jvm_prog \<times> state_type \<Rightarrow> state_type"
 
-recdef eff' "{}"
-"eff' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)"
-"eff' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])"
-"eff' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)"
-"eff' (Getfield F C, G, (oT#ST, LT))    = (snd (the (field (G,C) F)) # ST, LT)"
-"eff' (Putfield F C, G, (vT#oT#ST, LT)) = (ST,LT)"
-"eff' (New C, G, (ST,LT))               = (Class C # ST, LT)"
-"eff' (Checkcast C, G, (RefT rt#ST,LT)) = (Class C # ST,LT)"
-"eff' (Pop, G, (ts#ST,LT))              = (ST,LT)"
-"eff' (Dup, G, (ts#ST,LT))              = (ts#ts#ST,LT)"
-"eff' (Dup_x1, G, (ts1#ts2#ST,LT))      = (ts1#ts2#ts1#ST,LT)"
-"eff' (Dup_x2, G, (ts1#ts2#ts3#ST,LT))  = (ts1#ts2#ts3#ts1#ST,LT)"
-"eff' (Swap, G, (ts1#ts2#ST,LT))        = (ts2#ts1#ST,LT)"
+fun eff' :: "instr \<times> jvm_prog \<times> state_type \<Rightarrow> state_type"
+where
+"eff' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)" |
+"eff' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])" |
+"eff' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)" |
+"eff' (Getfield F C, G, (oT#ST, LT))    = (snd (the (field (G,C) F)) # ST, LT)" |
+"eff' (Putfield F C, G, (vT#oT#ST, LT)) = (ST,LT)" |
+"eff' (New C, G, (ST,LT))               = (Class C # ST, LT)" |
+"eff' (Checkcast C, G, (RefT rt#ST,LT)) = (Class C # ST,LT)" |
+"eff' (Pop, G, (ts#ST,LT))              = (ST,LT)" |
+"eff' (Dup, G, (ts#ST,LT))              = (ts#ts#ST,LT)" |
+"eff' (Dup_x1, G, (ts1#ts2#ST,LT))      = (ts1#ts2#ts1#ST,LT)" |
+"eff' (Dup_x2, G, (ts1#ts2#ts3#ST,LT))  = (ts1#ts2#ts3#ts1#ST,LT)" |
+"eff' (Swap, G, (ts1#ts2#ST,LT))        = (ts2#ts1#ST,LT)" |
 "eff' (IAdd, G, (PrimT Integer#PrimT Integer#ST,LT)) 
-                                         = (PrimT Integer#ST,LT)"
-"eff' (Ifcmpeq b, G, (ts1#ts2#ST,LT))   = (ST,LT)"
-"eff' (Goto b, G, s)                    = s"
+                                         = (PrimT Integer#ST,LT)" |
+"eff' (Ifcmpeq b, G, (ts1#ts2#ST,LT))   = (ST,LT)" |
+"eff' (Goto b, G, s)                    = s" |
   -- "Return has no successor instruction in the same method"
-"eff' (Return, G, s)                    = s" 
+"eff' (Return, G, s)                    = s" |
   -- "Throw always terminates abruptly"
-"eff' (Throw, G, s)                     = s"
+"eff' (Throw, G, s)                     = s" |
 "eff' (Invoke C mn fpTs, G, (ST,LT))    = (let ST' = drop (length fpTs) ST 
   in  (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" 
 
 
-consts
-  match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
-primrec
+
+primrec match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" where
   "match_any G pc [] = []"
-  "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e;
+| "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e;
                                 es' = match_any G pc es 
                             in 
                             if start_pc <= pc \<and> pc < end_pc then catch_type#es' else es')"
 
-consts
-  match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
-primrec
+primrec match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" where
   "match G X pc [] = []"
-  "match G X pc (e#es) = 
+| "match G X pc (e#es) = 
   (if match_exception_entry G (Xcpt X) pc e then [Xcpt X] else match G X pc es)"
 
 lemma match_some_entry:
   "match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G (Xcpt X) pc e then [Xcpt X] else [])"
   by (induct et) auto
 
-consts
+fun
   xcpt_names :: "instr \<times> jvm_prog \<times> p_count \<times> exception_table \<Rightarrow> cname list" 
-recdef xcpt_names "{}"
+where
   "xcpt_names (Getfield F C, G, pc, et) = match G NullPointer pc et" 
-  "xcpt_names (Putfield F C, G, pc, et) = match G NullPointer pc et" 
-  "xcpt_names (New C, G, pc, et)        = match G OutOfMemory pc et"
-  "xcpt_names (Checkcast C, G, pc, et)  = match G ClassCast pc et"
-  "xcpt_names (Throw, G, pc, et)        = match_any G pc et"
-  "xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et" 
-  "xcpt_names (i, G, pc, et)            = []" 
+| "xcpt_names (Putfield F C, G, pc, et) = match G NullPointer pc et" 
+| "xcpt_names (New C, G, pc, et)        = match G OutOfMemory pc et"
+| "xcpt_names (Checkcast C, G, pc, et)  = match G ClassCast pc et"
+| "xcpt_names (Throw, G, pc, et)        = match_any G pc et"
+| "xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et" 
+| "xcpt_names (i, G, pc, et)            = []" 
 
 
-constdefs
-  xcpt_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_type option \<Rightarrow> exception_table \<Rightarrow> succ_type"
+definition xcpt_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_type option \<Rightarrow> exception_table \<Rightarrow> succ_type" where
   "xcpt_eff i G pc s et == 
    map (\<lambda>C. (the (match_exception_table G C pc et), case s of None \<Rightarrow> None | Some s' \<Rightarrow> Some ([Class C], snd s'))) 
        (xcpt_names (i,G,pc,et))"
 
-  norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option"
+definition norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option" where
   "norm_eff i G == Option.map (\<lambda>s. eff' (i,G,s))"
 
-  eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> succ_type"
+definition eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> succ_type" where
   "eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)"
 
-constdefs
-  isPrimT :: "ty \<Rightarrow> bool"
+definition isPrimT :: "ty \<Rightarrow> bool" where
   "isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False"
 
-  isRefT :: "ty \<Rightarrow> bool"
+definition isRefT :: "ty \<Rightarrow> bool" where
   "isRefT T == case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> True"
 
 lemma isPrimT [simp]:
@@ -127,61 +119,60 @@
 
 
 text "Conditions under which eff is applicable:"
-consts
+
+fun
 app' :: "instr \<times> jvm_prog \<times> p_count \<times> nat \<times> ty \<times> state_type \<Rightarrow> bool"
-
-recdef app' "{}"
+where
 "app' (Load idx, G, pc, maxs, rT, s) = 
-  (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)"
+  (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)" |
 "app' (Store idx, G, pc, maxs, rT, (ts#ST, LT)) = 
-  (idx < length LT)"
+  (idx < length LT)" |
 "app' (LitPush v, G, pc, maxs, rT, s) = 
-  (length (fst s) < maxs \<and> typeof (\<lambda>t. None) v \<noteq> None)"
+  (length (fst s) < maxs \<and> typeof (\<lambda>t. None) v \<noteq> None)" |
 "app' (Getfield F C, G, pc, maxs, rT, (oT#ST, LT)) = 
   (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and> 
-  G \<turnstile> oT \<preceq> (Class C))"
+  G \<turnstile> oT \<preceq> (Class C))" |
 "app' (Putfield F C, G, pc, maxs, rT, (vT#oT#ST, LT)) = 
   (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and>
-  G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> (snd (the (field (G,C) F))))" 
+  G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> (snd (the (field (G,C) F))))" |
 "app' (New C, G, pc, maxs, rT, s) = 
-  (is_class G C \<and> length (fst s) < maxs)"
+  (is_class G C \<and> length (fst s) < maxs)" |
 "app' (Checkcast C, G, pc, maxs, rT, (RefT rt#ST,LT)) = 
-  (is_class G C)"
+  (is_class G C)" |
 "app' (Pop, G, pc, maxs, rT, (ts#ST,LT)) = 
-  True"
+  True" |
 "app' (Dup, G, pc, maxs, rT, (ts#ST,LT)) = 
-  (1+length ST < maxs)"
+  (1+length ST < maxs)" |
 "app' (Dup_x1, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = 
-  (2+length ST < maxs)"
+  (2+length ST < maxs)" |
 "app' (Dup_x2, G, pc, maxs, rT, (ts1#ts2#ts3#ST,LT)) = 
-  (3+length ST < maxs)"
+  (3+length ST < maxs)" |
 "app' (Swap, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = 
-  True"
+  True" |
 "app' (IAdd, G, pc, maxs, rT, (PrimT Integer#PrimT Integer#ST,LT)) =
-  True"
+  True" |
 "app' (Ifcmpeq b, G, pc, maxs, rT, (ts#ts'#ST,LT)) = 
-  (0 \<le> int pc + b \<and> (isPrimT ts \<and> ts' = ts \<or> isRefT ts \<and> isRefT ts'))"
+  (0 \<le> int pc + b \<and> (isPrimT ts \<and> ts' = ts \<or> isRefT ts \<and> isRefT ts'))" |
 "app' (Goto b, G, pc, maxs, rT, s) = 
-  (0 \<le> int pc + b)"
+  (0 \<le> int pc + b)" |
 "app' (Return, G, pc, maxs, rT, (T#ST,LT)) = 
-  (G \<turnstile> T \<preceq> rT)"
+  (G \<turnstile> T \<preceq> rT)" |
 "app' (Throw, G, pc, maxs, rT, (T#ST,LT)) = 
-  isRefT T"
+  isRefT T" |
 "app' (Invoke C mn fpTs, G, pc, maxs, rT, s) = 
   (length fpTs < length (fst s) \<and> 
   (let apTs = rev (take (length fpTs) (fst s));
        X    = hd (drop (length fpTs) (fst s)) 
    in  
        G \<turnstile> X \<preceq> Class C \<and> is_class G C \<and> method (G,C) (mn,fpTs) \<noteq> None \<and>
-       list_all2 (\<lambda>x y. G \<turnstile> x \<preceq> y) apTs fpTs))"
+       list_all2 (\<lambda>x y. G \<turnstile> x \<preceq> y) apTs fpTs))" |
   
 "app' (i,G, pc,maxs,rT,s) = False"
 
-constdefs
-  xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool"
+definition xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool" where
   "xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C"
 
-  app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> bool"
+definition app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> bool" where
   "app i G maxs rT pc et s == case s of None \<Rightarrow> True | Some t \<Rightarrow> app' (i,G,pc,maxs,rT,t) \<and> xcpt_app i G pc et"
 
 
@@ -218,7 +209,7 @@
 qed auto
 
 lemma 2: "\<not>(2 < length a) \<Longrightarrow> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
-proof -;
+proof -
   assume "\<not>(2 < length a)"
   hence "length a < (Suc (Suc (Suc 0)))" by simp
   hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)" 
@@ -278,7 +269,7 @@
   "(app (Checkcast C) G maxs rT pc et (Some s)) =  
   (\<exists>rT ST LT. s = (RefT rT#ST,LT) \<and> is_class G C \<and>
   (\<forall>x \<in> set (match G ClassCast pc et). is_class G x))"
-  by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto)
+  by (cases s, cases "fst s", simp) (cases "hd (fst s)", auto)
 
 lemma appPop[simp]: 
 "(app Pop G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT))"
@@ -369,7 +360,7 @@
     assume app: "?app (a,b)"
     hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> 
            length fpTs < length a" (is "?a \<and> ?l") 
-      by (auto simp add: app_def)
+      by auto
     hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l") 
       by auto
     hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs" 
@@ -384,7 +375,7 @@
     hence "\<exists>apTs X ST. a = rev apTs @ X # ST \<and> length apTs = length fpTs" 
       by blast
     with app
-    show ?thesis by (unfold app_def, clarsimp) blast
+    show ?thesis by clarsimp blast
   qed
   with Pair 
   have "?app s \<Longrightarrow> ?P s" by (simp only:)
--- a/src/HOL/MicroJava/BV/JType.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/JType.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,19 +9,17 @@
 imports "../DFA/Semilattices" "../J/WellForm"
 begin
 
-constdefs
-  super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname"
+definition super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname" where
   "super G C == fst (the (class G C))"
 
 lemma superI:
   "G \<turnstile> C \<prec>C1 D \<Longrightarrow> super G C = D"
   by (unfold super_def) (auto dest: subcls1D)
 
-constdefs
-  is_ref :: "ty \<Rightarrow> bool"
+definition is_ref :: "ty \<Rightarrow> bool" where
   "is_ref T == case T of PrimT t \<Rightarrow> False | RefT r \<Rightarrow> True"
 
-  sup :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty err"
+definition sup :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty err" where
   "sup G T1 T2 ==
   case T1 of PrimT P1 \<Rightarrow> (case T2 of PrimT P2 \<Rightarrow> 
                          (if P1 = P2 then OK (PrimT P1) else Err) | RefT R \<Rightarrow> Err)
@@ -30,17 +28,16 @@
             | ClassT C \<Rightarrow> (case R2 of NullT \<Rightarrow> OK (Class C) 
                            | ClassT D \<Rightarrow> OK (Class (exec_lub (subcls1 G) (super G) C D)))))"
 
-  subtype :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"
+definition subtype :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" where
   "subtype G T1 T2 == G \<turnstile> T1 \<preceq> T2"
 
-  is_ty :: "'c prog \<Rightarrow> ty \<Rightarrow> bool"
+definition is_ty :: "'c prog \<Rightarrow> ty \<Rightarrow> bool" where
   "is_ty G T == case T of PrimT P \<Rightarrow> True | RefT R \<Rightarrow>
                (case R of NullT \<Rightarrow> True | ClassT C \<Rightarrow> (C, Object) \<in> (subcls1 G)^*)"
 
 abbreviation "types G == Collect (is_type G)"
 
-constdefs
-  esl :: "'c prog \<Rightarrow> ty esl"
+definition esl :: "'c prog \<Rightarrow> ty esl" where
   "esl G == (types G, subtype G, sup G)"
 
 lemma PrimT_PrimT: "(G \<turnstile> xb \<preceq> PrimT p) = (xb = PrimT p)"
--- a/src/HOL/MicroJava/BV/JVM.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/JVM.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,14 +9,13 @@
 imports Typing_Framework_JVM
 begin
 
-constdefs
-  kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
-             instr list \<Rightarrow> JVMType.state list \<Rightarrow> JVMType.state list"
+definition kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
+             instr list \<Rightarrow> JVMType.state list \<Rightarrow> JVMType.state list" where
   "kiljvm G maxs maxr rT et bs ==
   kildall (JVMType.le G maxs maxr) (JVMType.sup G maxs maxr) (exec G maxs rT et bs)"
 
-  wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
-             exception_table \<Rightarrow> instr list \<Rightarrow> bool"
+definition wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
+             exception_table \<Rightarrow> instr list \<Rightarrow> bool" where
   "wt_kil G C pTs rT mxs mxl et ins ==
    check_bounded ins et \<and> 0 < size ins \<and> 
    (let first  = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));
@@ -24,11 +23,10 @@
         result = kiljvm G mxs (1+size pTs+mxl) rT et ins start
     in \<forall>n < size ins. result!n \<noteq> Err)"
 
-  wt_jvm_prog_kildall :: "jvm_prog \<Rightarrow> bool"
+definition wt_jvm_prog_kildall :: "jvm_prog \<Rightarrow> bool" where
   "wt_jvm_prog_kildall G ==
   wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). wt_kil G C (snd sig) rT maxs maxl et b) G"
 
-
 theorem is_bcv_kiljvm:
   "\<lbrakk> wf_prog wf_mb G; bounded (exec G maxs rT et bs) (size bs) \<rbrakk> \<Longrightarrow>
       is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs)
--- a/src/HOL/MicroJava/BV/JVMType.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/JVMType.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -19,55 +19,47 @@
   prog_type    = "cname \<Rightarrow> class_type"
 
 
-constdefs
-  stk_esl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty list esl"
+definition stk_esl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty list esl" where
   "stk_esl S maxs == upto_esl maxs (JType.esl S)"
 
-  reg_sl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty err list sl"
+definition reg_sl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty err list sl" where
   "reg_sl S maxr == Listn.sl maxr (Err.sl (JType.esl S))"
 
-  sl :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state sl"
+definition sl :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state sl" where
   "sl S maxs maxr ==
   Err.sl(Opt.esl(Product.esl (stk_esl S maxs) (Err.esl(reg_sl S maxr))))"
 
-constdefs
-  states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state set"
+definition states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state set" where
   "states S maxs maxr == fst(sl S maxs maxr)"
 
-  le :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state ord"
+definition le :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state ord" where
   "le S maxs maxr == fst(snd(sl S maxs maxr))"
 
-  sup :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state binop"
+definition  sup :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> state binop" where
   "sup S maxs maxr == snd(snd(sl S maxs maxr))"
 
-
-constdefs
-  sup_ty_opt :: "['code prog,ty err,ty err] \<Rightarrow> bool" 
-                 ("_ |- _ <=o _" [71,71] 70)
+definition sup_ty_opt :: "['code prog,ty err,ty err] \<Rightarrow> bool"
+                 ("_ |- _ <=o _" [71,71] 70) where 
   "sup_ty_opt G == Err.le (subtype G)"
 
-  sup_loc :: "['code prog,locvars_type,locvars_type] \<Rightarrow> bool" 
-              ("_ |- _ <=l _"  [71,71] 70)
+definition sup_loc :: "['code prog,locvars_type,locvars_type] \<Rightarrow> bool" 
+              ("_ |- _ <=l _"  [71,71] 70) where
   "sup_loc G == Listn.le (sup_ty_opt G)"
 
-  sup_state :: "['code prog,state_type,state_type] \<Rightarrow> bool"   
-               ("_ |- _ <=s _"  [71,71] 70)
+definition sup_state :: "['code prog,state_type,state_type] \<Rightarrow> bool"   
+               ("_ |- _ <=s _"  [71,71] 70) where
   "sup_state G == Product.le (Listn.le (subtype G)) (sup_loc G)"
 
-  sup_state_opt :: "['code prog,state_type option,state_type option] \<Rightarrow> bool" 
-                   ("_ |- _ <=' _"  [71,71] 70)
+definition sup_state_opt :: "['code prog,state_type option,state_type option] \<Rightarrow> bool" 
+                   ("_ |- _ <=' _"  [71,71] 70) where
   "sup_state_opt G == Opt.le (sup_state G)"
 
 
-syntax (xsymbols)
-  sup_ty_opt    :: "['code prog,ty err,ty err] \<Rightarrow> bool" 
-                   ("_ \<turnstile> _ <=o _" [71,71] 70)
-  sup_loc       :: "['code prog,locvars_type,locvars_type] \<Rightarrow> bool" 
-                   ("_ \<turnstile> _ <=l _" [71,71] 70)
-  sup_state     :: "['code prog,state_type,state_type] \<Rightarrow> bool" 
-                   ("_ \<turnstile> _ <=s _" [71,71] 70)
-  sup_state_opt :: "['code prog,state_type option,state_type option] \<Rightarrow> bool"
-                   ("_ \<turnstile> _ <=' _" [71,71] 70)
+notation (xsymbols)
+  sup_ty_opt  ("_ \<turnstile> _ <=o _" [71,71] 70) and
+  sup_loc  ("_ \<turnstile> _ <=l _" [71,71] 70) and
+  sup_state  ("_ \<turnstile> _ <=s _" [71,71] 70) and
+  sup_state_opt  ("_ \<turnstile> _ <=' _" [71,71] 70)
                    
 
 lemma JVM_states_unfold: 
--- a/src/HOL/MicroJava/BV/LBVJVM.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/LBVJVM.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -11,18 +11,17 @@
 
 types prog_cert = "cname \<Rightarrow> sig \<Rightarrow> JVMType.state list"
 
-constdefs
-  check_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool"
+definition check_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" where
   "check_cert G mxs mxr n cert \<equiv> check_types G mxs mxr cert \<and> length cert = n+1 \<and>
                                  (\<forall>i<n. cert!i \<noteq> Err) \<and> cert!n = OK None"
 
-  lbvjvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
-             JVMType.state list \<Rightarrow> instr list \<Rightarrow> JVMType.state \<Rightarrow> JVMType.state"
+definition lbvjvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
+             JVMType.state list \<Rightarrow> instr list \<Rightarrow> JVMType.state \<Rightarrow> JVMType.state" where
   "lbvjvm G maxs maxr rT et cert bs \<equiv>
   wtl_inst_list bs cert  (JVMType.sup G maxs maxr) (JVMType.le G maxs maxr) Err (OK None) (exec G maxs rT et bs) 0"
 
-  wt_lbv :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
-             exception_table \<Rightarrow> JVMType.state list \<Rightarrow> instr list \<Rightarrow> bool"
+definition wt_lbv :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
+             exception_table \<Rightarrow> JVMType.state list \<Rightarrow> instr list \<Rightarrow> bool" where
   "wt_lbv G C pTs rT mxs mxl et cert ins \<equiv>
    check_bounded ins et \<and> 
    check_cert G mxs (1+size pTs+mxl) (length ins) cert \<and>
@@ -31,15 +30,15 @@
         result = lbvjvm G mxs (1+size pTs+mxl) rT et cert ins (OK start)
     in result \<noteq> Err)"
 
-  wt_jvm_prog_lbv :: "jvm_prog \<Rightarrow> prog_cert \<Rightarrow> bool"
+definition wt_jvm_prog_lbv :: "jvm_prog \<Rightarrow> prog_cert \<Rightarrow> bool" where
   "wt_jvm_prog_lbv G cert \<equiv>
   wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). wt_lbv G C (snd sig) rT maxs maxl et (cert C sig) b) G"
 
-  mk_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list 
-              \<Rightarrow> method_type \<Rightarrow> JVMType.state list"
+definition mk_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list 
+              \<Rightarrow> method_type \<Rightarrow> JVMType.state list" where
   "mk_cert G maxs rT et bs phi \<equiv> make_cert (exec G maxs rT et bs) (map OK phi) (OK None)"
 
-  prg_cert :: "jvm_prog \<Rightarrow> prog_type \<Rightarrow> prog_cert"
+definition prg_cert :: "jvm_prog \<Rightarrow> prog_type \<Rightarrow> prog_cert" where
   "prg_cert G phi C sig \<equiv> let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in 
                            mk_cert G maxs rT et ins (phi C sig)"
  
--- a/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,13 +9,11 @@
 imports "../DFA/Abstract_BV" JVMType EffectMono BVSpec
 begin
 
-constdefs
-  exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state step_type"
+definition exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state step_type" where
   "exec G maxs rT et bs == 
   err_step (size bs) (\<lambda>pc. app (bs!pc) G maxs rT pc et) (\<lambda>pc. eff (bs!pc) G pc et)"
 
-constdefs
-  opt_states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (ty list \<times> ty err list) option set"
+definition opt_states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (ty list \<times> ty err list) option set" where
   "opt_states G maxs maxr \<equiv> opt (\<Union>{list n (types G) |n. n \<le> maxs} \<times> list maxr (err (types G)))"
 
 
@@ -26,8 +24,7 @@
   "list_all'_rec P n []     = True"
   "list_all'_rec P n (x#xs) = (P x n \<and> list_all'_rec P (Suc n) xs)"
 
-constdefs
-  list_all' :: "('a \<Rightarrow> nat \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
+definition list_all' :: "('a \<Rightarrow> nat \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
   "list_all' P xs \<equiv> list_all'_rec P 0 xs"
 
 lemma list_all'_rec:
--- a/src/HOL/MicroJava/Comp/CorrComp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/CorrComp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -81,13 +81,12 @@
 
 (***********************************************************************)
 
-constdefs
-  progression :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
+definition progression :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
                  aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow>
                  bytecode \<Rightarrow>
                  aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow> 
                  bool"
-  ("{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60)
+  ("{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where
   "{G,C,S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1} ==
   \<forall> pre post frs.
   (gis (gmb G C S) = pre @ instrs @ post) \<longrightarrow>
@@ -161,10 +160,9 @@
 done
 
 (*****)
-constdefs
-  jump_fwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
+definition jump_fwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
                  aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> 
-                 instr \<Rightarrow> bytecode \<Rightarrow> bool"
+                 instr \<Rightarrow> bytecode \<Rightarrow> bool" where
   "jump_fwd G C S hp lvars os0 os1 instr instrs ==
   \<forall> pre post frs.
   (gis (gmb G C S) = pre @ instr # instrs @ post) \<longrightarrow>
@@ -216,10 +214,9 @@
 
 
 (* note: instrs and instr reversed wrt. jump_fwd *)
-constdefs
-  jump_bwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
+definition jump_bwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
                  aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> 
-                 bytecode \<Rightarrow> instr \<Rightarrow> bool"
+                 bytecode \<Rightarrow> instr \<Rightarrow> bool" where
   "jump_bwd G C S hp lvars os0 os1 instrs instr ==
   \<forall> pre post frs.
   (gis (gmb G C S) = pre @ instrs @ instr # post) \<longrightarrow>
@@ -258,14 +255,14 @@
 (**********************************************************************)
 
 (* class C with signature S is defined in program G *)
-constdefs class_sig_defined :: "'c prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> bool"
+definition class_sig_defined :: "'c prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> bool" where
   "class_sig_defined G C S == 
   is_class G C \<and> (\<exists> D rT mb. (method (G, C) S = Some (D, rT, mb)))"
 
 
 (* The environment of a java method body 
   (characterized by class and signature) *)
-constdefs env_of_jmb :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> java_mb env"
+definition env_of_jmb :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> java_mb env" where
   "env_of_jmb G C S == 
   (let (mn,pTs) = S;
        (D,rT,(pns,lvars,blk,res)) = the(method (G, C) S) in
@@ -331,11 +328,13 @@
 
 (**********************************************************************)
 
-constdefs wtpd_expr :: "java_mb env \<Rightarrow> expr \<Rightarrow> bool"
+definition wtpd_expr :: "java_mb env \<Rightarrow> expr \<Rightarrow> bool" where
   "wtpd_expr E e == (\<exists> T. E\<turnstile>e :: T)"
-  wtpd_exprs :: "java_mb env \<Rightarrow> (expr list) \<Rightarrow> bool"
+
+definition wtpd_exprs :: "java_mb env \<Rightarrow> (expr list) \<Rightarrow> bool" where
   "wtpd_exprs E e == (\<exists> T. E\<turnstile>e [::] T)"
-  wtpd_stmt :: "java_mb env \<Rightarrow> stmt \<Rightarrow> bool" 
+
+definition wtpd_stmt :: "java_mb env \<Rightarrow> stmt \<Rightarrow> bool" where
   "wtpd_stmt E c == (E\<turnstile>c \<surd>)"
 
 lemma wtpd_expr_newc: "wtpd_expr E (NewC C) \<Longrightarrow> is_class (prg E) C"
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -12,13 +12,13 @@
 
 (**********************************************************************)
 
-constdefs     
-   inited_LT :: "[cname, ty list, (vname \<times> ty) list] \<Rightarrow> locvars_type"
+definition inited_LT :: "[cname, ty list, (vname \<times> ty) list] \<Rightarrow> locvars_type" where
   "inited_LT C pTs lvars == (OK (Class C))#((map OK pTs))@(map (Fun.comp OK snd) lvars)"
-   is_inited_LT :: "[cname, ty list, (vname \<times> ty) list, locvars_type] \<Rightarrow> bool"
+
+definition is_inited_LT :: "[cname, ty list, (vname \<times> ty) list, locvars_type] \<Rightarrow> bool" where
   "is_inited_LT C pTs lvars LT == (LT = (inited_LT C pTs lvars))"
 
-  local_env :: "[java_mb prog, cname, sig, vname list,(vname \<times> ty) list] \<Rightarrow> java_mb env"
+definition local_env :: "[java_mb prog, cname, sig, vname list,(vname \<times> ty) list] \<Rightarrow> java_mb env" where
   "local_env G C S pns lvars == 
      let (mn, pTs) = S in (G,map_of lvars(pns[\<mapsto>]pTs)(This\<mapsto>Class C))"
 
@@ -536,13 +536,12 @@
 
 (**********************************************************************)
 
-constdefs
-  offset_xcentry :: "[nat, exception_entry] \<Rightarrow> exception_entry"
+definition offset_xcentry :: "[nat, exception_entry] \<Rightarrow> exception_entry" where
   "offset_xcentry == 
       \<lambda> n (start_pc, end_pc, handler_pc, catch_type).
           (start_pc + n, end_pc + n, handler_pc + n, catch_type)"
 
-  offset_xctable :: "[nat, exception_table] \<Rightarrow> exception_table"
+definition offset_xctable :: "[nat, exception_table] \<Rightarrow> exception_table" where
   "offset_xctable n == (map (offset_xcentry n))"
 
 lemma match_xcentry_offset [simp]: "
@@ -682,12 +681,11 @@
 (**********************************************************************)
 
 
-constdefs
-  start_sttp_resp_cons :: "[state_type \<Rightarrow> method_type \<times> state_type] \<Rightarrow> bool"
+definition start_sttp_resp_cons :: "[state_type \<Rightarrow> method_type \<times> state_type] \<Rightarrow> bool" where
   "start_sttp_resp_cons f == 
      (\<forall> sttp. let (mt', sttp') = (f sttp) in (\<exists>mt'_rest. mt' = Some sttp # mt'_rest))"
 
-  start_sttp_resp :: "[state_type \<Rightarrow> method_type \<times> state_type] \<Rightarrow> bool"
+definition start_sttp_resp :: "[state_type \<Rightarrow> method_type \<times> state_type] \<Rightarrow> bool" where
   "start_sttp_resp f == (f = comb_nil) \<or> (start_sttp_resp_cons f)"
 
 lemma start_sttp_resp_comb_nil [simp]: "start_sttp_resp comb_nil"
@@ -887,10 +885,9 @@
 
   (* ******************************************************************* *)
 
-constdefs
-   bc_mt_corresp :: "
+definition bc_mt_corresp :: "
   [bytecode, state_type \<Rightarrow> method_type \<times> state_type, state_type, jvm_prog, ty, nat, p_count]
-  \<Rightarrow> bool"
+  \<Rightarrow> bool" where
 
   "bc_mt_corresp bc f sttp0 cG rT mxr idx ==
   let (mt, sttp) = f sttp0 in
@@ -993,8 +990,7 @@
   (* ********************************************************************** *)
 
 
-constdefs
-  mt_sttp_flatten :: "method_type \<times> state_type \<Rightarrow> method_type"
+definition mt_sttp_flatten :: "method_type \<times> state_type \<Rightarrow> method_type" where
   "mt_sttp_flatten mt_sttp == (mt_of mt_sttp) @ [Some (sttp_of mt_sttp)]"
 
 
@@ -1473,8 +1469,7 @@
 
 
   (* ******************** *)
-constdefs 
-  contracting :: "(state_type \<Rightarrow> method_type \<times> state_type) \<Rightarrow> bool"
+definition contracting :: "(state_type \<Rightarrow> method_type \<times> state_type) \<Rightarrow> bool" where
   "contracting f == (\<forall> ST LT. 
                     let (ST', LT') = sttp_of (f (ST, LT)) 
                     in (length ST' \<le> length ST \<and> set ST' \<subseteq> set ST  \<and>
--- a/src/HOL/MicroJava/Comp/DefsComp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/DefsComp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/DefsComp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -10,50 +9,65 @@
 begin
 
 
-constdefs
-  method_rT :: "cname \<times> ty \<times> 'c \<Rightarrow> ty"
+definition method_rT :: "cname \<times> ty \<times> 'c \<Rightarrow> ty" where
   "method_rT mtd == (fst (snd mtd))"
 
 
-constdefs
 (* g = get *)
-  gx :: "xstate \<Rightarrow> val option"  "gx \<equiv> fst"
-  gs :: "xstate \<Rightarrow> state"  "gs \<equiv> snd"
-  gh :: "xstate \<Rightarrow> aheap"        "gh \<equiv> fst\<circ>snd"
-  gl :: "xstate \<Rightarrow> State.locals" "gl \<equiv> snd\<circ>snd"
+definition
+  gx :: "xstate \<Rightarrow> val option" where "gx \<equiv> fst"
+
+definition
+  gs :: "xstate \<Rightarrow> state" where "gs \<equiv> snd"
+
+definition
+  gh :: "xstate \<Rightarrow> aheap" where "gh \<equiv> fst\<circ>snd"
 
+definition
+  gl :: "xstate \<Rightarrow> State.locals" where "gl \<equiv> snd\<circ>snd"
+
+definition
   gmb :: "'a prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 'a"
-    "gmb G cn si \<equiv> snd(snd(the(method (G,cn) si)))"
+    where "gmb G cn si \<equiv> snd(snd(the(method (G,cn) si)))"
+
+definition
   gis :: "jvm_method \<Rightarrow> bytecode"
-    "gis \<equiv> fst \<circ> snd \<circ> snd"
+    where "gis \<equiv> fst \<circ> snd \<circ> snd"
 
 (* jmb = aus einem JavaMaethodBody *)
-  gjmb_pns  :: "java_mb \<Rightarrow> vname list"     "gjmb_pns \<equiv> fst"
-  gjmb_lvs  :: "java_mb \<Rightarrow> (vname\<times>ty)list" "gjmb_lvs \<equiv> fst\<circ>snd"
-  gjmb_blk  :: "java_mb \<Rightarrow> stmt"           "gjmb_blk \<equiv> fst\<circ>snd\<circ>snd"
-  gjmb_res  :: "java_mb \<Rightarrow> expr"           "gjmb_res \<equiv> snd\<circ>snd\<circ>snd"
+definition
+  gjmb_pns  :: "java_mb \<Rightarrow> vname list" where "gjmb_pns \<equiv> fst"
+
+definition
+  gjmb_lvs  :: "java_mb \<Rightarrow> (vname\<times>ty)list" where  "gjmb_lvs \<equiv> fst\<circ>snd"
+
+definition
+  gjmb_blk  :: "java_mb \<Rightarrow> stmt" where  "gjmb_blk \<equiv> fst\<circ>snd\<circ>snd"
+
+definition
+  gjmb_res  :: "java_mb \<Rightarrow> expr" where  "gjmb_res \<equiv> snd\<circ>snd\<circ>snd"
+
+definition
   gjmb_plns :: "java_mb \<Rightarrow> vname list"
-    "gjmb_plns \<equiv> \<lambda>jmb. gjmb_pns jmb @ map fst (gjmb_lvs jmb)"
+    where  "gjmb_plns \<equiv> \<lambda>jmb. gjmb_pns jmb @ map fst (gjmb_lvs jmb)"
 
+definition
   glvs :: "java_mb \<Rightarrow> State.locals \<Rightarrow> locvars"
-    "glvs jmb loc \<equiv> map (the\<circ>loc) (gjmb_plns jmb)"
+    where "glvs jmb loc \<equiv> map (the\<circ>loc) (gjmb_plns jmb)"
   
 lemmas gdefs = gx_def gh_def gl_def gmb_def gis_def glvs_def
 lemmas gjmbdefs = gjmb_pns_def gjmb_lvs_def gjmb_blk_def gjmb_res_def gjmb_plns_def
 
 lemmas galldefs = gdefs gjmbdefs
 
-
-
-constdefs 
-  locvars_locals :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> State.locals \<Rightarrow> locvars"
+definition locvars_locals :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> State.locals \<Rightarrow> locvars" where
   "locvars_locals G C S lvs == the (lvs This) # glvs (gmb G C S) lvs"
 
-  locals_locvars :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> locvars \<Rightarrow> State.locals"
+definition locals_locvars :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> locvars \<Rightarrow> State.locals" where
   "locals_locvars G C S lvs == 
   empty ((gjmb_plns (gmb G C S))[\<mapsto>](tl lvs)) (This\<mapsto>(hd lvs))"
 
-  locvars_xstate :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> xstate \<Rightarrow> locvars"
+definition locvars_xstate :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> xstate \<Rightarrow> locvars" where
   "locvars_xstate G C S xs == locvars_locals G C S (gl xs)"
 
 
--- a/src/HOL/MicroJava/Comp/Index.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/Index.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,8 +9,7 @@
 begin
 
 (*indexing a variable name among all variable declarations in a method body*)
-constdefs
- index :: "java_mb => vname => nat"
+definition index :: "java_mb => vname => nat" where
  "index ==  \<lambda> (pn,lv,blk,res) v.
   if v = This
   then 0 
@@ -99,8 +98,7 @@
 
 (* The following def should replace the conditions in WellType.thy / wf_java_mdecl
 *)
-constdefs 
-  disjoint_varnames :: "[vname list, (vname \<times> ty) list] \<Rightarrow> bool"
+definition disjoint_varnames :: "[vname list, (vname \<times> ty) list] \<Rightarrow> bool" where
 (* This corresponds to the original def in wf_java_mdecl:
   "disjoint_varnames pns lvars \<equiv> 
   nodups pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
--- a/src/HOL/MicroJava/Comp/TranslComp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/TranslComp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -98,17 +98,16 @@
 (*compiling methods, classes and programs*) 
 
 (*initialising a single variable*)
-constdefs
- load_default_val :: "ty => instr"
+definition load_default_val :: "ty => instr" where
 "load_default_val ty == LitPush (default_val ty)"
 
- compInit :: "java_mb => (vname * ty) => instr list"
+definition compInit :: "java_mb => (vname * ty) => instr list" where
 "compInit jmb == \<lambda> (vn,ty). [load_default_val ty, Store (index jmb vn)]"
 
- compInitLvars :: "[java_mb, (vname \<times> ty) list] \<Rightarrow> bytecode"
+definition compInitLvars :: "[java_mb, (vname \<times> ty) list] \<Rightarrow> bytecode" where
  "compInitLvars jmb lvars == concat (map (compInit jmb) lvars)"
 
-  compMethod :: "java_mb prog \<Rightarrow> cname \<Rightarrow> java_mb mdecl \<Rightarrow> jvm_method mdecl"
+definition compMethod :: "java_mb prog \<Rightarrow> cname \<Rightarrow> java_mb mdecl \<Rightarrow> jvm_method mdecl" where
   "compMethod G C jmdl == let (sig, rT, jmb) = jmdl;
                         (pns,lvars,blk,res) = jmb;
                         mt = (compTpMethod G C jmdl);
@@ -117,10 +116,10 @@
                              [Return]
                   in (sig, rT, max_ssize mt, length lvars, bc, [])"
 
-  compClass :: "java_mb prog => java_mb cdecl=> jvm_method cdecl"
+definition compClass :: "java_mb prog => java_mb cdecl=> jvm_method cdecl" where
   "compClass G == \<lambda> (C,cno,fdls,jmdls). (C,cno,fdls, map (compMethod G C) jmdls)"
 
-  comp :: "java_mb prog => jvm_prog"
+definition comp :: "java_mb prog => jvm_prog" where
   "comp G == map (compClass G) G"
 
 end
--- a/src/HOL/MicroJava/Comp/TranslCompTp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -6,22 +6,18 @@
 imports Index "../BV/JVMType"
 begin
 
-
-
 (**********************************************************************)
 
-
-constdefs
-  comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" 
+definition comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" where 
   "comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0; 
                             (xs2, x2) = f2 x1 
                         in  (xs1 @ xs2, x2))"
-  comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a"
+
+definition comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a" where
   "comb_nil a == ([], a)"
 
-syntax (xsymbols)
-  "comb" :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd"    
-            (infixr "\<box>" 55)
+notation (xsymbols)
+  comb  (infixr "\<box>" 55)
 
 lemma comb_nil_left [simp]: "comb_nil \<box> f = f"
 by (simp add: comb_def comb_nil_def split_beta)
@@ -59,23 +55,26 @@
   compTpStmt  :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> stmt
                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 
-
-constdefs
-  nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type"
+definition nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type" where
   "nochangeST sttp == ([Some sttp], sttp)"
-  pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type"
+
+definition pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type" where
   "pushST tps == (\<lambda> (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))"
-  dupST :: "state_type \<Rightarrow> method_type \<times> state_type"
+
+definition dupST :: "state_type \<Rightarrow> method_type \<times> state_type" where
   "dupST == (\<lambda> (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))"
-  dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type"
+
+definition dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type" where
   "dup_x1ST == (\<lambda> (ST, LT). ([Some (ST, LT)], 
                              (hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))"
-  popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type"
+
+definition popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type" where
   "popST n == (\<lambda> (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))"
-  replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
+
+definition replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where
   "replST n tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))"
 
-  storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
+definition storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where
   "storeST i tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tl ST, LT [i:= OK tp])))"
 
 
@@ -138,9 +137,8 @@
       (pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>
       (compTpStmt jmb G c) \<box> nochangeST"
 
-constdefs
-  compTpInit  :: "java_mb \<Rightarrow> (vname * ty)
-                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
+definition compTpInit :: "java_mb \<Rightarrow> (vname * ty)
+                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" where
   "compTpInit jmb == (\<lambda> (vn,ty). (pushST [ty]) \<box>  (storeST (index jmb vn) ty))"
 
 consts
@@ -151,14 +149,13 @@
   "compTpInitLvars jmb [] = comb_nil"
   "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \<box> (compTpInitLvars jmb lvars)"
 
-constdefs
-   start_ST :: "opstack_type"
+definition start_ST :: "opstack_type" where
   "start_ST == []"
 
-   start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type"
+definition start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type" where
   "start_LT C pTs n ==  (OK (Class C))#((map OK pTs))@(replicate n Err)"
 
-  compTpMethod  :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type"
+definition compTpMethod  :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type" where
   "compTpMethod G C == \<lambda> ((mn,pTs),rT, jmb). 
                          let (pns,lvars,blk,res) = jmb
                          in (mt_of
@@ -168,7 +165,7 @@
                               nochangeST)
                                 (start_ST, start_LT C pTs (length lvars))))"
 
-  compTp :: "java_mb prog => prog_type"
+definition compTp :: "java_mb prog => prog_type" where
   "compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig))
                       in compTpMethod G C (sig, rT, jmb)"
 
@@ -177,15 +174,13 @@
 (**********************************************************************)
   (* Computing the maximum stack size from the method_type *)
 
-constdefs
-  ssize_sto :: "(state_type option) \<Rightarrow> nat"
+definition ssize_sto :: "(state_type option) \<Rightarrow> nat" where
   "ssize_sto sto ==  case sto of None \<Rightarrow> 0 | (Some (ST, LT)) \<Rightarrow> length ST"
 
-  max_of_list :: "nat list \<Rightarrow> nat"
+definition max_of_list :: "nat list \<Rightarrow> nat" where
   "max_of_list xs == foldr max xs 0"
 
-  max_ssize :: "method_type \<Rightarrow> nat"
+definition max_ssize :: "method_type \<Rightarrow> nat" where
   "max_ssize mt == max_of_list (map ssize_sto mt)"
 
-
 end
--- a/src/HOL/MicroJava/Comp/TypeInf.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/Comp/TypeInf.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/TypeInf.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -169,10 +168,10 @@
 
   
 
-constdefs 
-  inferred_tp  :: "[java_mb env, expr] \<Rightarrow> ty"
+definition inferred_tp :: "[java_mb env, expr] \<Rightarrow> ty" where
   "inferred_tp E e == (SOME T. E\<turnstile>e :: T)"
-  inferred_tps :: "[java_mb env, expr list] \<Rightarrow> ty list"
+
+definition inferred_tps :: "[java_mb env, expr list] \<Rightarrow> ty list" where
   "inferred_tps E es == (SOME Ts. E\<turnstile>es [::] Ts)"
 
 (* get inferred type(s) for well-typed term *)
--- a/src/HOL/MicroJava/DFA/Err.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Err.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -14,35 +14,32 @@
 types 'a ebinop = "'a \<Rightarrow> 'a \<Rightarrow> 'a err"
       'a esl =    "'a set * 'a ord * 'a ebinop"
 
-consts
-  ok_val :: "'a err \<Rightarrow> 'a"
-primrec
+primrec ok_val :: "'a err \<Rightarrow> 'a" where
   "ok_val (OK x) = x"
 
-constdefs
- lift :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
+definition lift :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)" where
 "lift f e == case e of Err \<Rightarrow> Err | OK x \<Rightarrow> f x"
 
- lift2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a err \<Rightarrow> 'b err \<Rightarrow> 'c err"
+definition lift2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a err \<Rightarrow> 'b err \<Rightarrow> 'c err" where
 "lift2 f e1 e2 ==
  case e1 of Err  \<Rightarrow> Err
           | OK x \<Rightarrow> (case e2 of Err \<Rightarrow> Err | OK y \<Rightarrow> f x y)"
 
- le :: "'a ord \<Rightarrow> 'a err ord"
+definition le :: "'a ord \<Rightarrow> 'a err ord" where
 "le r e1 e2 ==
         case e2 of Err \<Rightarrow> True |
                    OK y \<Rightarrow> (case e1 of Err \<Rightarrow> False | OK x \<Rightarrow> x <=_r y)"
 
- sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a err \<Rightarrow> 'b err \<Rightarrow> 'c err)"
+definition sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a err \<Rightarrow> 'b err \<Rightarrow> 'c err)" where
 "sup f == lift2(%x y. OK(x +_f y))"
 
- err :: "'a set \<Rightarrow> 'a err set"
+definition err :: "'a set \<Rightarrow> 'a err set" where
 "err A == insert Err {x . ? y:A. x = OK y}"
 
- esl :: "'a sl \<Rightarrow> 'a esl"
+definition esl :: "'a sl \<Rightarrow> 'a esl" where
 "esl == %(A,r,f). (A,r, %x y. OK(f x y))"
 
- sl :: "'a esl \<Rightarrow> 'a err sl"
+definition sl :: "'a esl \<Rightarrow> 'a err sl" where
 "sl == %(A,r,f). (err A, le r, lift2 f)"
 
 abbreviation
--- a/src/HOL/MicroJava/DFA/Kildall.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Kildall.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,37 +9,28 @@
 imports SemilatAlg While_Combinator
 begin
 
-
-consts
- iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow>
-          's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set"
- propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set"
+primrec propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set" where
+  "propa f []      ss w = (ss,w)"
+| "propa f (q'#qs) ss w = (let (q,t) = q';
+                               u = t +_f ss!q;
+                               w' = (if u = ss!q then w else insert q w)
+                           in propa f qs (ss[q := u]) w')"
 
-primrec
-"propa f []      ss w = (ss,w)"
-"propa f (q'#qs) ss w = (let (q,t) = q';
-                             u = t +_f ss!q;
-                             w' = (if u = ss!q then w else insert q w)
-                         in propa f qs (ss[q := u]) w')"
-
-defs iter_def:
-"iter f step ss w ==
- while (%(ss,w). w \<noteq> {})
+definition iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set" where
+  "iter f step ss w == while (%(ss,w). w \<noteq> {})
        (%(ss,w). let p = SOME p. p \<in> w
                  in propa f (step p (ss!p)) ss (w-{p}))
        (ss,w)"
 
-constdefs
- unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set"
+definition unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set" where
 "unstables r step ss == {p. p < size ss \<and> \<not>stable r step ss p}"
 
- kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list"
+definition kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list" where
 "kildall r f step ss == fst(iter f step ss (unstables r step ss))"
 
-consts merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list"
-primrec
-"merges f []      ss = ss"
-"merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))"
+primrec merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list" where
+  "merges f []      ss = ss"
+| "merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))"
 
 
 lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric]
--- a/src/HOL/MicroJava/DFA/LBVComplete.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/LBVComplete.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,12 +9,11 @@
 imports LBVSpec Typing_Framework
 begin
 
-constdefs
-  is_target :: "['s step_type, 's list, nat] \<Rightarrow> bool" 
+definition is_target :: "['s step_type, 's list, nat] \<Rightarrow> bool" where 
   "is_target step phi pc' \<equiv>
      \<exists>pc s'. pc' \<noteq> pc+1 \<and> pc < length phi \<and> (pc',s') \<in> set (step pc (phi!pc))"
 
-  make_cert :: "['s step_type, 's list, 's] \<Rightarrow> 's certificate"
+definition make_cert :: "['s step_type, 's list, 's] \<Rightarrow> 's certificate" where
   "make_cert step phi B \<equiv> 
      map (\<lambda>pc. if is_target step phi pc then phi!pc else B) [0..<length phi] @ [B]"
 
--- a/src/HOL/MicroJava/DFA/LBVSpec.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/LBVSpec.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -12,46 +12,38 @@
 types
   's certificate = "'s list"   
 
-consts
-merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's"
-primrec
-"merge cert f r T pc []     x = x"
-"merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in 
+primrec merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's" where
+  "merge cert f r T pc []     x = x"
+| "merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in 
                                   if pc'=pc+1 then s' +_f x
                                   else if s' <=_r (cert!pc') then x
                                   else T)"
 
-constdefs
-wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow>
-             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
+definition wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow>
+             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
 "wtl_inst cert f r T step pc s \<equiv> merge cert f r T pc (step pc s) (cert!(pc+1))"
 
-wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
-             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
+definition wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
+             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
 "wtl_cert cert f r T B step pc s \<equiv>
   if cert!pc = B then 
     wtl_inst cert f r T step pc s
   else
     if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T"
 
-consts 
-wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
-                  's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
-primrec
-"wtl_inst_list []     cert f r T B step pc s = s"
-"wtl_inst_list (i#is) cert f r T B step pc s = 
+primrec wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
+                  's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
+  "wtl_inst_list []     cert f r T B step pc s = s"
+| "wtl_inst_list (i#is) cert f r T B step pc s = 
     (let s' = wtl_cert cert f r T B step pc s in
       if s' = T \<or> s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')"
 
-constdefs
-  cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool"
+definition cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool" where
   "cert_ok cert n T B A \<equiv> (\<forall>i < n. cert!i \<in> A \<and> cert!i \<noteq> T) \<and> (cert!n = B)"
 
-constdefs
-  bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+definition bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool" where
   "bottom r B \<equiv> \<forall>x. B <=_r x"
 
-
 locale lbv = Semilat +
   fixes T :: "'a" ("\<top>") 
   fixes B :: "'a" ("\<bottom>") 
--- a/src/HOL/MicroJava/DFA/Listn.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Listn.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,12 +9,10 @@
 imports Err
 begin
 
-constdefs
-
- list :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a list set"
+definition list :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a list set" where
 "list n A == {xs. length xs = n & set xs <= A}"
 
- le :: "'a ord \<Rightarrow> ('a list)ord"
+definition le :: "'a ord \<Rightarrow> ('a list)ord" where
 "le r == list_all2 (%x y. x <=_r y)"
 
 abbreviation
@@ -27,8 +25,7 @@
        ("(_ /<[_] _)" [50, 0, 51] 50)
   where "x <[r] y == x <_(le r) y"
 
-constdefs
- map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
+definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
 "map2 f == (%xs ys. map (split f) (zip xs ys))"
 
 abbreviation
@@ -36,19 +33,17 @@
        ("(_ /+[_] _)" [65, 0, 66] 65)
   where "x +[f] y == x +_(map2 f) y"
 
-consts coalesce :: "'a err list \<Rightarrow> 'a list err"
-primrec
-"coalesce [] = OK[]"
-"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)"
+primrec coalesce :: "'a err list \<Rightarrow> 'a list err" where
+  "coalesce [] = OK[]"
+| "coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)"
 
-constdefs
- sl :: "nat \<Rightarrow> 'a sl \<Rightarrow> 'a list sl"
+definition sl :: "nat \<Rightarrow> 'a sl \<Rightarrow> 'a list sl" where
 "sl n == %(A,r,f). (list n A, le r, map2 f)"
 
- sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list err"
+definition sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list err" where
 "sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err"
 
- upto_esl :: "nat \<Rightarrow> 'a esl \<Rightarrow> 'a list esl"
+definition upto_esl :: "nat \<Rightarrow> 'a esl \<Rightarrow> 'a list esl" where
 "upto_esl m == %(A,r,f). (Union{list n A |n. n <= m}, le r, sup f)"
 
 lemmas [simp] = set_update_subsetI
--- a/src/HOL/MicroJava/DFA/Opt.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Opt.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,21 +9,20 @@
 imports Err
 begin
 
-constdefs
- le :: "'a ord \<Rightarrow> 'a option ord"
+definition le :: "'a ord \<Rightarrow> 'a option ord" where
 "le r o1 o2 == case o2 of None \<Rightarrow> o1=None |
                               Some y \<Rightarrow> (case o1 of None \<Rightarrow> True
                                                   | Some x \<Rightarrow> x <=_r y)"
 
- opt :: "'a set \<Rightarrow> 'a option set"
+definition opt :: "'a set \<Rightarrow> 'a option set" where
 "opt A == insert None {x . ? y:A. x = Some y}"
 
- sup :: "'a ebinop \<Rightarrow> 'a option ebinop"
+definition sup :: "'a ebinop \<Rightarrow> 'a option ebinop" where
 "sup f o1 o2 ==  
  case o1 of None \<Rightarrow> OK o2 | Some x \<Rightarrow> (case o2 of None \<Rightarrow> OK o1
      | Some y \<Rightarrow> (case f x y of Err \<Rightarrow> Err | OK z \<Rightarrow> OK (Some z)))"
 
- esl :: "'a esl \<Rightarrow> 'a option esl"
+definition esl :: "'a esl \<Rightarrow> 'a option esl" where
 "esl == %(A,r,f). (opt A, le r, sup f)"
 
 lemma unfold_le_opt:
--- a/src/HOL/MicroJava/DFA/Product.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Product.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,14 +9,13 @@
 imports Err
 begin
 
-constdefs
- le :: "'a ord \<Rightarrow> 'b ord \<Rightarrow> ('a * 'b) ord"
+definition le :: "'a ord \<Rightarrow> 'b ord \<Rightarrow> ('a * 'b) ord" where
 "le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'"
 
- sup :: "'a ebinop \<Rightarrow> 'b ebinop \<Rightarrow> ('a * 'b)ebinop"
+definition sup :: "'a ebinop \<Rightarrow> 'b ebinop \<Rightarrow> ('a * 'b)ebinop" where
 "sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)"
 
- esl :: "'a esl \<Rightarrow> 'b esl \<Rightarrow> ('a * 'b ) esl"
+definition esl :: "'a esl \<Rightarrow> 'b esl \<Rightarrow> ('a * 'b ) esl" where
 "esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)"
 
 abbreviation
--- a/src/HOL/MicroJava/DFA/Semilat.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Semilat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -32,16 +32,19 @@
   "lesssub"  ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
   "plussub"  ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
 (*<*)
-syntax
- (* allow \<sub> instead of \<bsub>..\<esub> *)  
-  "_lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
-  "_lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
-  "_plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
+(* allow \<sub> instead of \<bsub>..\<esub> *)
+
+abbreviation (input)
+  lesub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
+  where "x \<sqsubseteq>\<^sub>r y == x \<sqsubseteq>\<^bsub>r\<^esub> y"
 
-translations
-  "x \<sqsubseteq>\<^sub>r y" => "x \<sqsubseteq>\<^bsub>r\<^esub> y"
-  "x \<sqsubset>\<^sub>r y" => "x \<sqsubset>\<^bsub>r\<^esub> y" 
-  "x \<squnion>\<^sub>f y" => "x \<squnion>\<^bsub>f\<^esub> y" 
+abbreviation (input)
+  lesssub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
+  where "x \<sqsubset>\<^sub>r y == x \<sqsubset>\<^bsub>r\<^esub> y"
+
+abbreviation (input)
+  plussub1 :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
+  where "x \<squnion>\<^sub>f y == x \<squnion>\<^bsub>f\<^esub> y"
 (*>*)
 
 defs
@@ -49,36 +52,34 @@
   lesssub_def: "x \<sqsubset>\<^sub>r y \<equiv> x \<sqsubseteq>\<^sub>r y \<and> x \<noteq> y"
   plussub_def: "x \<squnion>\<^sub>f y \<equiv> f x y"
 
-constdefs
-  ord :: "('a \<times> 'a) set \<Rightarrow> 'a ord"
+definition ord :: "('a \<times> 'a) set \<Rightarrow> 'a ord" where
   "ord r \<equiv> \<lambda>x y. (x,y) \<in> r"
 
-  order :: "'a ord \<Rightarrow> bool"
+definition order :: "'a ord \<Rightarrow> bool" where
   "order r \<equiv> (\<forall>x. x \<sqsubseteq>\<^sub>r x) \<and> (\<forall>x y. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r x \<longrightarrow> x=y) \<and> (\<forall>x y z. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<sqsubseteq>\<^sub>r z)"
 
-  top :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+definition top :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool" where
   "top r T \<equiv> \<forall>x. x \<sqsubseteq>\<^sub>r T"
   
-  acc :: "'a ord \<Rightarrow> bool"
+definition acc :: "'a ord \<Rightarrow> bool" where
   "acc r \<equiv> wf {(y,x). x \<sqsubset>\<^sub>r y}"
 
-  closed :: "'a set \<Rightarrow> 'a binop \<Rightarrow> bool"
+definition closed :: "'a set \<Rightarrow> 'a binop \<Rightarrow> bool" where
   "closed A f \<equiv> \<forall>x\<in>A. \<forall>y\<in>A. x \<squnion>\<^sub>f y \<in> A"
 
-  semilat :: "'a sl \<Rightarrow> bool"
+definition semilat :: "'a sl \<Rightarrow> bool" where
   "semilat \<equiv> \<lambda>(A,r,f). order r \<and> closed A f \<and> 
                        (\<forall>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and>
                        (\<forall>x\<in>A. \<forall>y\<in>A. y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and>
                        (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z)"
 
-
-  is_ub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+definition is_ub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
   "is_ub r x y u \<equiv> (x,u)\<in>r \<and> (y,u)\<in>r"
 
-  is_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+definition is_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
   "is_lub r x y u \<equiv> is_ub r x y u \<and> (\<forall>z. is_ub r x y z \<longrightarrow> (u,z)\<in>r)"
 
-  some_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+definition some_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   "some_lub r x y \<equiv> SOME z. is_lub r x y z"
 
 locale Semilat =
@@ -304,8 +305,7 @@
 
 subsection{*An executable lub-finder*}
 
-constdefs
- exec_lub :: "('a * 'a) set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a binop"
+definition exec_lub :: "('a * 'a) set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a binop" where
 "exec_lub r f x y \<equiv> while (\<lambda>z. (x,z) \<notin> r\<^sup>*) f y"
 
 lemma exec_lub_refl: "exec_lub r f T T = T"
--- a/src/HOL/MicroJava/DFA/SemilatAlg.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,29 +9,24 @@
 imports Typing_Framework Product
 begin
 
-constdefs 
-  lesubstep_type :: "(nat \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool"
-                    ("(_ /<=|_| _)" [50, 0, 51] 50)
+definition lesubstep_type :: "(nat \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool"
+                    ("(_ /<=|_| _)" [50, 0, 51] 50) where
   "x <=|r| y \<equiv> \<forall>(p,s) \<in> set x. \<exists>s'. (p,s') \<in> set y \<and> s <=_r s'"
 
-consts
-  plusplussub :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65)
-primrec
+primrec plusplussub :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) where
   "[] ++_f y = y"
-  "(x#xs) ++_f y = xs ++_f (x +_f y)"
+| "(x#xs) ++_f y = xs ++_f (x +_f y)"
 
-constdefs
- bounded :: "'s step_type \<Rightarrow> nat \<Rightarrow> bool"
+definition bounded :: "'s step_type \<Rightarrow> nat \<Rightarrow> bool" where
 "bounded step n == !p<n. !s. !(q,t):set(step p s). q<n"  
 
- pres_type :: "'s step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+definition pres_type :: "'s step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool" where
 "pres_type step n A == \<forall>s\<in>A. \<forall>p<n. \<forall>(q,s')\<in>set (step p s). s' \<in> A"
 
- mono :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+definition mono :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool" where
 "mono r step n A ==
  \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> step p s <=|r| step p t"
 
-
 lemma pres_typeD:
   "\<lbrakk> pres_type step n A; s\<in>A; p<n; (q,s')\<in>set (step p s) \<rbrakk> \<Longrightarrow> s' \<in> A"
   by (unfold pres_type_def, blast)
--- a/src/HOL/MicroJava/DFA/Typing_Framework.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Typing_Framework.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -15,20 +15,19 @@
 types
   's step_type = "nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list"
 
-constdefs
- stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool"
+definition stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool" where
 "stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q"
 
- stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+definition stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool" where
 "stables r step ss == !p<size ss. stable r step ss p"
 
- wt_step ::
-"'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+definition wt_step ::
+"'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool" where
 "wt_step r T step ts ==
  !p<size(ts). ts!p ~= T & stable r step ts p"
 
- is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
-           \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool"  
+definition is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
+           \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool" where
 "is_bcv r T step n A bcv == !ss : list n A.
    (!p<n. (bcv ss)!p ~= T) =
    (? ts: list n A. ss <=[r] ts & wt_step r T step ts)"
--- a/src/HOL/MicroJava/DFA/Typing_Framework_err.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/DFA/Typing_Framework_err.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,28 +9,26 @@
 imports Typing_Framework SemilatAlg
 begin
 
-constdefs
-
-wt_err_step :: "'s ord \<Rightarrow> 's err step_type \<Rightarrow> 's err list \<Rightarrow> bool"
+definition wt_err_step :: "'s ord \<Rightarrow> 's err step_type \<Rightarrow> 's err list \<Rightarrow> bool" where
 "wt_err_step r step ts \<equiv> wt_step (Err.le r) Err step ts"
 
-wt_app_eff :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+definition wt_app_eff :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool" where
 "wt_app_eff r app step ts \<equiv>
   \<forall>p < size ts. app p (ts!p) \<and> (\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q)"
 
-map_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'c) list"
+definition map_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'c) list" where
 "map_snd f \<equiv> map (\<lambda>(x,y). (x, f y))"
 
-error :: "nat \<Rightarrow> (nat \<times> 'a err) list"
+definition error :: "nat \<Rightarrow> (nat \<times> 'a err) list" where
 "error n \<equiv> map (\<lambda>x. (x,Err)) [0..<n]"
 
-err_step :: "nat \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's err step_type"
+definition err_step :: "nat \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's err step_type" where
 "err_step n app step p t \<equiv> 
   case t of 
     Err   \<Rightarrow> error n
   | OK t' \<Rightarrow> if app p t' then map_snd OK (step p t') else error n"
 
-app_mono :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+definition app_mono :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool" where
 "app_mono r app n A \<equiv>
  \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> app p s"
 
--- a/src/HOL/MicroJava/J/Conform.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/Conform.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/J/Conform.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -10,52 +9,39 @@
 
 types 'c env' = "'c prog \<times> (vname \<rightharpoonup> ty)"  -- "same as @{text env} of @{text WellType.thy}"
 
-constdefs
-
-  hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50)
+definition hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) where
  "h<=|h' == \<forall>a C fs. h a = Some(C,fs) --> (\<exists>fs'. h' a = Some(C,fs'))"
 
-  conf :: "'c prog => aheap => val => ty => bool" 
-                                   ("_,_ |- _ ::<= _"  [51,51,51,51] 50)
+definition conf :: "'c prog => aheap => val => ty => bool" 
+                                   ("_,_ |- _ ::<= _"  [51,51,51,51] 50) where
  "G,h|-v::<=T == \<exists>T'. typeof (Option.map obj_ty o h) v = Some T' \<and> G\<turnstile>T'\<preceq>T"
 
-  lconf :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
-                                   ("_,_ |- _ [::<=] _" [51,51,51,51] 50)
+definition lconf :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
+                                   ("_,_ |- _ [::<=] _" [51,51,51,51] 50) where
  "G,h|-vs[::<=]Ts == \<forall>n T. Ts n = Some T --> (\<exists>v. vs n = Some v \<and> G,h|-v::<=T)"
 
-  oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50)
+definition oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) where
  "G,h|-obj [ok] == G,h|-snd obj[::<=]map_of (fields (G,fst obj))"
 
-  hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50)
+definition hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) where
  "G|-h h [ok]    == \<forall>a obj. h a = Some obj --> G,h|-obj [ok]"
  
-  xconf :: "aheap \<Rightarrow> val option \<Rightarrow> bool"
+definition xconf :: "aheap \<Rightarrow> val option \<Rightarrow> bool" where
   "xconf hp vo  == preallocated hp \<and> (\<forall> v. (vo = Some v) \<longrightarrow> (\<exists> xc. v = (Addr (XcptRef xc))))"
 
-  conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50)
+definition conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) where
  "s::<=E == prg E|-h heap (store s) [ok] \<and> 
             prg E,heap (store s)|-locals (store s)[::<=]localT E \<and> 
             xconf (heap (store s)) (abrupt s)"
 
 
-syntax (xsymbols)
-  hext     :: "aheap => aheap => bool"
-              ("_ \<le>| _" [51,51] 50)
-
-  conf     :: "'c prog => aheap => val => ty => bool"
-              ("_,_ \<turnstile> _ ::\<preceq> _" [51,51,51,51] 50)
-
-  lconf    :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
-              ("_,_ \<turnstile> _ [::\<preceq>] _" [51,51,51,51] 50)
-
-  oconf    :: "'c prog => aheap => obj => bool"
-              ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50)
-
-  hconf    :: "'c prog => aheap => bool"
-              ("_ \<turnstile>h _ \<surd>" [51,51] 50)
-
-  conforms :: "state => java_mb env' => bool"
-              ("_ ::\<preceq> _" [51,51] 50)
+notation (xsymbols)
+  hext  ("_ \<le>| _" [51,51] 50) and
+  conf  ("_,_ \<turnstile> _ ::\<preceq> _" [51,51,51,51] 50) and
+  lconf  ("_,_ \<turnstile> _ [::\<preceq>] _" [51,51,51,51] 50) and
+  oconf  ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50) and
+  hconf  ("_ \<turnstile>h _ \<surd>" [51,51] 50) and
+  conforms  ("_ ::\<preceq> _" [51,51] 50)
 
 
 section "hext"
--- a/src/HOL/MicroJava/J/Decl.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/Decl.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/J/Decl.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -32,11 +31,10 @@
   "prog  c" <= (type) "(c cdecl) list"
 
 
-constdefs
-  "class" :: "'c prog => (cname \<rightharpoonup> 'c class)"
+definition "class" :: "'c prog => (cname \<rightharpoonup> 'c class)" where
   "class \<equiv> map_of"
 
-  is_class :: "'c prog => cname => bool"
+definition is_class :: "'c prog => cname => bool" where
   "is_class G C \<equiv> class G C \<noteq> None"
 
 
@@ -46,10 +44,8 @@
 apply (rule finite_dom_map_of)
 done
 
-consts
-  is_type :: "'c prog => ty    => bool"
-primrec
+primrec is_type :: "'c prog => ty => bool" where
   "is_type G (PrimT pt) = True"
-  "is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)"
+| "is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)"
 
 end
--- a/src/HOL/MicroJava/J/Eval.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/Eval.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/J/Eval.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -11,22 +10,16 @@
 
   -- "Auxiliary notions"
 
-constdefs
-  fits    :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
+definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
  "G,s\<turnstile>a' fits T  \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
 
-constdefs
-  catch ::"java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60)
+definition catch :: "java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
  "G,s\<turnstile>catch C\<equiv>  case abrupt s of None \<Rightarrow> False | Some a \<Rightarrow> G,store s\<turnstile> a fits Class C"
 
-
-
-constdefs
-  lupd       :: "vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"        ("lupd'(_\<mapsto>_')"[10,10]1000)
+definition lupd :: "vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" ("lupd'(_\<mapsto>_')"[10,10]1000) where
  "lupd vn v   \<equiv> \<lambda> (hp,loc). (hp, (loc(vn\<mapsto>v)))"
 
-constdefs
-  new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate"
+definition new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate" where
  "new_xcpt_var vn \<equiv>  \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)"
 
 
--- a/src/HOL/MicroJava/J/Exceptions.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/Exceptions.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -6,11 +6,10 @@
 theory Exceptions imports State begin
 
 text {* a new, blank object with default values in all fields: *}
-constdefs
-  blank :: "'c prog \<Rightarrow> cname \<Rightarrow> obj"
+definition blank :: "'c prog \<Rightarrow> cname \<Rightarrow> obj" where
   "blank G C \<equiv> (C,init_vars (fields(G,C)))" 
 
-  start_heap :: "'c prog \<Rightarrow> aheap"
+definition start_heap :: "'c prog \<Rightarrow> aheap" where
   "start_heap G \<equiv> empty (XcptRef NullPointer \<mapsto> blank G (Xcpt NullPointer))
                         (XcptRef ClassCast \<mapsto> blank G (Xcpt ClassCast))
                         (XcptRef OutOfMemory \<mapsto> blank G (Xcpt OutOfMemory))"
@@ -21,8 +20,7 @@
   where "cname_of hp v == fst (the (hp (the_Addr v)))"
 
 
-constdefs
-  preallocated :: "aheap \<Rightarrow> bool"
+definition preallocated :: "aheap \<Rightarrow> bool" where
   "preallocated hp \<equiv> \<forall>x. \<exists>fs. hp (XcptRef x) = Some (Xcpt x, fs)"
 
 lemma preallocatedD:
--- a/src/HOL/MicroJava/J/JBasis.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/JBasis.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -15,8 +15,7 @@
 
 section "unique"
  
-constdefs
-  unique  :: "('a \<times> 'b) list => bool"
+definition unique :: "('a \<times> 'b) list => bool" where
   "unique  == distinct \<circ> map fst"
 
 
--- a/src/HOL/MicroJava/J/JListExample.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/JListExample.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -21,29 +21,28 @@
   l3_nam :: vnam
   l4_nam :: vnam
 
-constdefs
-  val_name :: vname
+definition val_name :: vname where
   "val_name == VName val_nam"
 
-  next_name :: vname
+definition next_name :: vname where
   "next_name == VName next_nam"
 
-  l_name :: vname
+definition l_name :: vname where
   "l_name == VName l_nam"
 
-  l1_name :: vname
+definition l1_name :: vname where
   "l1_name == VName l1_nam"
 
-  l2_name :: vname
+definition l2_name :: vname where
   "l2_name == VName l2_nam"
 
-  l3_name :: vname
+definition l3_name :: vname where
   "l3_name == VName l3_nam"
 
-  l4_name :: vname
+definition l4_name :: vname where
   "l4_name == VName l4_nam"
 
-  list_class :: "java_mb class"
+definition list_class :: "java_mb class" where
   "list_class ==
     (Object,
      [(val_name, PrimT Integer), (next_name, RefT (ClassT list_name))],
@@ -56,7 +55,7 @@
            append_name({[RefT (ClassT list_name)]}[LAcc l_name])), 
        Lit Unit))])"
 
-  example_prg :: "java_mb prog"
+definition example_prg :: "java_mb prog" where
   "example_prg == [ObjectC, (list_name, list_class)]"
 
 types_code
--- a/src/HOL/MicroJava/J/State.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/State.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -14,11 +14,10 @@
 
   obj = "cname \<times> fields'"    -- "class instance with class name and fields"
 
-constdefs
-  obj_ty  :: "obj => ty"
+definition obj_ty :: "obj => ty" where
  "obj_ty obj  == Class (fst obj)"
 
-  init_vars :: "('a \<times> ty) list => ('a \<rightharpoonup> val)"
+definition init_vars :: "('a \<times> ty) list => ('a \<rightharpoonup> val)" where
  "init_vars == map_of o map (\<lambda>(n,T). (n,default_val T))"
 
 types aheap  = "loc \<rightharpoonup> obj"    -- {* "@{text heap}" used in a translation below *}
@@ -49,21 +48,19 @@
   lookup_obj :: "state \<Rightarrow> val \<Rightarrow> obj"
   where "lookup_obj s a' == the (heap s (the_Addr a'))"
 
-
-constdefs
-  raise_if :: "bool \<Rightarrow> xcpt \<Rightarrow> val option \<Rightarrow> val option"
+definition raise_if :: "bool \<Rightarrow> xcpt \<Rightarrow> val option \<Rightarrow> val option" where
   "raise_if b x xo \<equiv> if b \<and>  (xo = None) then Some (Addr (XcptRef x)) else xo"
 
-  new_Addr  :: "aheap => loc \<times> val option"
+definition new_Addr  :: "aheap => loc \<times> val option" where
   "new_Addr h \<equiv> SOME (a,x). (h a = None \<and>  x = None) |  x = Some (Addr (XcptRef OutOfMemory))"
 
-  np    :: "val => val option => val option"
+definition np    :: "val => val option => val option" where
  "np v == raise_if (v = Null) NullPointer"
 
-  c_hupd  :: "aheap => xstate => xstate"
+definition c_hupd  :: "aheap => xstate => xstate" where
  "c_hupd h'== \<lambda>(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))"
 
-  cast_ok :: "'c prog => cname => aheap => val => bool"
+definition cast_ok :: "'c prog => cname => aheap => val => bool" where
  "cast_ok G C h v == v = Null \<or> G\<turnstile>obj_ty (the (h (the_Addr v)))\<preceq> Class C"
 
 lemma obj_ty_def2 [simp]: "obj_ty (C,fs) = Class C"
--- a/src/HOL/MicroJava/J/SystemClasses.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/SystemClasses.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/J/SystemClasses.thy
-    ID:         $Id$
     Author:     Gerwin Klein
     Copyright   2002 Technische Universitaet Muenchen
 *)
@@ -13,20 +12,19 @@
   and the system exceptions.
 *}
 
-constdefs
-  ObjectC :: "'c cdecl"
+definition ObjectC :: "'c cdecl" where
   "ObjectC \<equiv> (Object, (undefined,[],[]))"
 
-  NullPointerC :: "'c cdecl"
+definition NullPointerC :: "'c cdecl" where
   "NullPointerC \<equiv> (Xcpt NullPointer, (Object,[],[]))"
 
-  ClassCastC :: "'c cdecl"
+definition ClassCastC :: "'c cdecl" where
   "ClassCastC \<equiv> (Xcpt ClassCast, (Object,[],[]))"
 
-  OutOfMemoryC :: "'c cdecl"
+definition OutOfMemoryC :: "'c cdecl" where
   "OutOfMemoryC \<equiv> (Xcpt OutOfMemory, (Object,[],[]))"
 
-  SystemClasses :: "'c cdecl list"
+definition SystemClasses :: "'c cdecl list" where
   "SystemClasses \<equiv> [ObjectC, NullPointerC, ClassCastC, OutOfMemoryC]"
 
 end
--- a/src/HOL/MicroJava/J/TypeRel.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/TypeRel.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -54,9 +54,8 @@
 apply  auto
 done
 
-constdefs
-  class_rec :: "'c prog \<Rightarrow> cname \<Rightarrow> 'a \<Rightarrow>
-    (cname \<Rightarrow> fdecl list \<Rightarrow> 'c mdecl list \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a"
+definition class_rec :: "'c prog \<Rightarrow> cname \<Rightarrow> 'a \<Rightarrow>
+    (cname \<Rightarrow> fdecl list \<Rightarrow> 'c mdecl list \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a" where
   "class_rec G == wfrec ((subcls1 G)^-1)
     (\<lambda>r C t f. case class G C of
          None \<Rightarrow> undefined
--- a/src/HOL/MicroJava/J/WellForm.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/WellForm.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -27,45 +27,44 @@
 *}
 types 'c wf_mb = "'c prog => cname => 'c mdecl => bool"
 
-constdefs
- wf_syscls :: "'c prog => bool"
+definition wf_syscls :: "'c prog => bool" where
 "wf_syscls G == let cs = set G in Object \<in> fst ` cs \<and> (\<forall>x. Xcpt x \<in> fst ` cs)"
 
- wf_fdecl :: "'c prog => fdecl => bool"
+definition wf_fdecl :: "'c prog => fdecl => bool" where
 "wf_fdecl G == \<lambda>(fn,ft). is_type G ft"
 
- wf_mhead :: "'c prog => sig => ty => bool"
+definition wf_mhead :: "'c prog => sig => ty => bool" where
 "wf_mhead G == \<lambda>(mn,pTs) rT. (\<forall>T\<in>set pTs. is_type G T) \<and> is_type G rT"
 
- ws_cdecl :: "'c prog => 'c cdecl => bool"
+definition ws_cdecl :: "'c prog => 'c cdecl => bool" where
 "ws_cdecl G ==
    \<lambda>(C,(D,fs,ms)).
   (\<forall>f\<in>set fs. wf_fdecl G         f) \<and>  unique fs \<and>
   (\<forall>(sig,rT,mb)\<in>set ms. wf_mhead G sig rT) \<and> unique ms \<and>
   (C \<noteq> Object \<longrightarrow> is_class G D \<and>  \<not>G\<turnstile>D\<preceq>C C)"
 
- ws_prog :: "'c prog => bool"
+definition ws_prog :: "'c prog => bool" where
 "ws_prog G == 
   wf_syscls G \<and> (\<forall>c\<in>set G. ws_cdecl G c) \<and> unique G"
 
- wf_mrT   :: "'c prog => 'c cdecl => bool"
+definition wf_mrT   :: "'c prog => 'c cdecl => bool" where
 "wf_mrT G ==
    \<lambda>(C,(D,fs,ms)).
   (C \<noteq> Object \<longrightarrow> (\<forall>(sig,rT,b)\<in>set ms. \<forall>D' rT' b'.
                       method(G,D) sig = Some(D',rT',b') --> G\<turnstile>rT\<preceq>rT'))"
 
- wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool"
+definition wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where
 "wf_cdecl_mdecl wf_mb G ==
    \<lambda>(C,(D,fs,ms)). (\<forall>m\<in>set ms. wf_mb G C m)"
 
- wf_prog :: "'c wf_mb => 'c prog => bool"
+definition wf_prog :: "'c wf_mb => 'c prog => bool" where
 "wf_prog wf_mb G == 
      ws_prog G \<and> (\<forall>c\<in> set G. wf_mrT G c \<and> wf_cdecl_mdecl wf_mb G c)"
 
- wf_mdecl :: "'c wf_mb => 'c wf_mb"
+definition wf_mdecl :: "'c wf_mb => 'c wf_mb" where
 "wf_mdecl wf_mb G C == \<lambda>(sig,rT,mb). wf_mhead G sig rT \<and> wf_mb G C (sig,rT,mb)"
 
- wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool"
+definition wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where
 "wf_cdecl wf_mb G ==
    \<lambda>(C,(D,fs,ms)).
   (\<forall>f\<in>set fs. wf_fdecl G         f) \<and>  unique fs \<and>
--- a/src/HOL/MicroJava/J/WellType.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/J/WellType.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -193,9 +193,7 @@
         E\<turnstile>While(e) s\<surd>"
 
 
-constdefs
-
- wf_java_mdecl :: "'c prog => cname => java_mb mdecl => bool"
+definition wf_java_mdecl :: "'c prog => cname => java_mb mdecl => bool" where
 "wf_java_mdecl G C == \<lambda>((mn,pTs),rT,(pns,lvars,blk,res)).
   length pTs = length pns \<and>
   distinct pns \<and>
--- a/src/HOL/MicroJava/JVM/JVMDefensive.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/JVM/JVMDefensive.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -4,7 +4,9 @@
 
 header {* \isaheader{A Defensive JVM} *}
 
-theory JVMDefensive imports JVMExec begin
+theory JVMDefensive
+imports JVMExec
+begin
 
 text {*
   Extend the state space by one element indicating a type error (or
@@ -16,39 +18,32 @@
   fifth :: "'a \<times> 'b \<times> 'c \<times> 'd \<times> 'e \<times> 'f \<Rightarrow> 'e"
   where "fifth x == fst(snd(snd(snd(snd x))))"
 
-
-consts isAddr :: "val \<Rightarrow> bool"
-recdef isAddr "{}"
+fun isAddr :: "val \<Rightarrow> bool" where
   "isAddr (Addr loc) = True"
-  "isAddr v          = False"
+| "isAddr v          = False"
 
-consts isIntg :: "val \<Rightarrow> bool"
-recdef isIntg "{}"
+fun isIntg :: "val \<Rightarrow> bool" where
   "isIntg (Intg i) = True"
-  "isIntg v        = False"
+| "isIntg v        = False"
 
-constdefs
-  isRef :: "val \<Rightarrow> bool"
+definition isRef :: "val \<Rightarrow> bool" where
   "isRef v \<equiv> v = Null \<or> isAddr v"
 
-
-consts
-  check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, 
-                  cname, sig, p_count, nat, frame list] \<Rightarrow> bool"
-primrec 
+primrec check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, 
+                  cname, sig, p_count, nat, frame list] \<Rightarrow> bool" where
   "check_instr (Load idx) G hp stk vars C sig pc mxs frs = 
   (idx < length vars \<and> size stk < mxs)"
 
-  "check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = 
+| "check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = 
   (0 < length stk \<and> idx < length vars)"
 
-  "check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = 
+| "check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = 
   (\<not>isAddr v \<and> size stk < mxs)"
 
-  "check_instr (New C) G hp stk vars Cl sig pc mxs frs = 
+| "check_instr (New C) G hp stk vars Cl sig pc mxs frs = 
   (is_class G C \<and> size stk < mxs)"
 
-  "check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = 
+| "check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = 
   (0 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> 
   (let (C', T) = the (field (G,C) F); ref = hd stk in 
     C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> 
@@ -56,7 +51,7 @@
       (let (D,vs) = the (hp (the_Addr ref)) in 
         G \<turnstile> D \<preceq>C C \<and> vs (F,C) \<noteq> None \<and> G,hp \<turnstile> the (vs (F,C)) ::\<preceq> T))))" 
 
-  "check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = 
+| "check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = 
   (1 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> 
   (let (C', T) = the (field (G,C) F); v = hd stk; ref = hd (tl stk) in 
     C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> 
@@ -64,10 +59,10 @@
       (let (D,vs) = the (hp (the_Addr ref)) in 
         G \<turnstile> D \<preceq>C C \<and> G,hp \<turnstile> v ::\<preceq> T))))" 
 
-  "check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs =
+| "check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs =
   (0 < length stk \<and> is_class G C \<and> isRef (hd stk))"
 
-  "check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs =
+| "check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs =
   (length ps < length stk \<and> 
   (let n = length ps; v = stk!n in
   isRef v \<and> (v \<noteq> Null \<longrightarrow> 
@@ -75,41 +70,40 @@
     method (G,cname_of hp v) (mn,ps) \<noteq> None \<and>
     list_all2 (\<lambda>v T. G,hp \<turnstile> v ::\<preceq> T) (rev (take n stk)) ps)))"
   
-  "check_instr Return G hp stk0 vars Cl sig0 pc mxs frs =
+| "check_instr Return G hp stk0 vars Cl sig0 pc mxs frs =
   (0 < length stk0 \<and> (0 < length frs \<longrightarrow> 
     method (G,Cl) sig0 \<noteq> None \<and>    
     (let v = hd stk0;  (C, rT, body) = the (method (G,Cl) sig0) in
     Cl = C \<and> G,hp \<turnstile> v ::\<preceq> rT)))"
  
-  "check_instr Pop G hp stk vars Cl sig pc mxs frs = 
+| "check_instr Pop G hp stk vars Cl sig pc mxs frs = 
   (0 < length stk)"
 
-  "check_instr Dup G hp stk vars Cl sig pc mxs frs = 
+| "check_instr Dup G hp stk vars Cl sig pc mxs frs = 
   (0 < length stk \<and> size stk < mxs)"
 
-  "check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = 
+| "check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = 
   (1 < length stk \<and> size stk < mxs)"
 
-  "check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = 
+| "check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = 
   (2 < length stk \<and> size stk < mxs)"
 
-  "check_instr Swap G hp stk vars Cl sig pc mxs frs =
+| "check_instr Swap G hp stk vars Cl sig pc mxs frs =
   (1 < length stk)"
 
-  "check_instr IAdd G hp stk vars Cl sig pc mxs frs =
+| "check_instr IAdd G hp stk vars Cl sig pc mxs frs =
   (1 < length stk \<and> isIntg (hd stk) \<and> isIntg (hd (tl stk)))"
 
-  "check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs =
+| "check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs =
   (1 < length stk \<and> 0 \<le> int pc+b)"
 
-  "check_instr (Goto b) G hp stk vars Cl sig pc mxs frs =
+| "check_instr (Goto b) G hp stk vars Cl sig pc mxs frs =
   (0 \<le> int pc+b)"
 
-  "check_instr Throw G hp stk vars Cl sig pc mxs frs =
+| "check_instr Throw G hp stk vars Cl sig pc mxs frs =
   (0 < length stk \<and> isRef (hd stk))"
 
-constdefs
-  check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool"
+definition check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool" where
   "check G s \<equiv> let (xcpt, hp, frs) = s in
                (case frs of [] \<Rightarrow> True | (stk,loc,C,sig,pc)#frs' \<Rightarrow> 
                 (let  (C',rt,mxs,mxl,ins,et) = the (method (G,C) sig); i = ins!pc in
@@ -117,25 +111,21 @@
                  check_instr i G hp stk loc C sig pc mxs frs'))"
 
 
-  exec_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state option type_error"
+definition exec_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state option type_error" where
   "exec_d G s \<equiv> case s of 
       TypeError \<Rightarrow> TypeError 
     | Normal s' \<Rightarrow> if check G s' then Normal (exec (G, s')) else TypeError"
 
 
-consts
-  "exec_all_d" :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool" 
+constdefs
+  exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool" 
                    ("_ |- _ -jvmd-> _" [61,61,61]60)
+  "G |- s -jvmd-> t \<equiv>
+         (s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union>
+                  {(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
 
-syntax (xsymbols)
-  "exec_all_d" :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool" 
-                   ("_ \<turnstile> _ -jvmd\<rightarrow> _" [61,61,61]60)  
- 
-defs
-  exec_all_d_def:
-  "G \<turnstile> s -jvmd\<rightarrow> t \<equiv>
-         (s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union> 
-                  {(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
+notation (xsymbols)
+  exec_all_d  ("_ \<turnstile> _ -jvmd\<rightarrow> _" [61,61,61]60)
 
 
 declare split_paired_All [simp del]
--- a/src/HOL/MicroJava/JVM/JVMExceptions.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/JVM/JVMExceptions.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,8 +7,7 @@
 
 theory JVMExceptions imports JVMInstructions begin
 
-constdefs
-  match_exception_entry :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_entry \<Rightarrow> bool"
+definition match_exception_entry :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_entry \<Rightarrow> bool" where
   "match_exception_entry G cn pc ee == 
                  let (start_pc, end_pc, handler_pc, catch_type) = ee in
                  start_pc <= pc \<and> pc < end_pc \<and> G\<turnstile> cn \<preceq>C catch_type"
--- a/src/HOL/MicroJava/JVM/JVMExec.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/JVM/JVMExec.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/JVM/JVMExec.thy
-    ID:         $Id$
     Author:     Cornelia Pusch, Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -9,32 +8,28 @@
 theory JVMExec imports JVMExecInstr JVMExceptions begin
 
 
-consts
+fun
   exec :: "jvm_prog \<times> jvm_state => jvm_state option"
-
-
--- "exec is not recursive. recdef is just used for pattern matching"
-recdef exec "{}"
+-- "exec is not recursive. fun is just used for pattern matching"
+where
   "exec (G, xp, hp, []) = None"
 
-  "exec (G, None, hp, (stk,loc,C,sig,pc)#frs) =
+| "exec (G, None, hp, (stk,loc,C,sig,pc)#frs) =
   (let 
      i = fst(snd(snd(snd(snd(the(method (G,C) sig)))))) ! pc;
      (xcpt', hp', frs') = exec_instr i G hp stk loc C sig pc frs
    in Some (find_handler G xcpt' hp' frs'))"
 
-  "exec (G, Some xp, hp, frs) = None" 
+| "exec (G, Some xp, hp, frs) = None" 
 
 
-constdefs
-  exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool"
-              ("_ |- _ -jvm-> _" [61,61,61]60)
+definition exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool"
+              ("_ |- _ -jvm-> _" [61,61,61]60) where
   "G |- s -jvm-> t == (s,t) \<in> {(s,t). exec(G,s) = Some t}^*"
 
 
-syntax (xsymbols)
-  exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool"
-              ("_ \<turnstile> _ -jvm\<rightarrow> _" [61,61,61]60)
+notation (xsymbols)
+  exec_all  ("_ \<turnstile> _ -jvm\<rightarrow> _" [61,61,61]60)
 
 text {*
   The start configuration of the JVM: in the start heap, we call a 
@@ -42,8 +37,7 @@
   @{text this} pointer of the frame is set to @{text Null} to simulate
   a static method invokation.
 *}
-constdefs  
-  start_state :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> jvm_state"
+definition start_state :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> jvm_state" where
   "start_state G C m \<equiv>
   let (C',rT,mxs,mxl,i,et) = the (method (G,C) (m,[])) in
     (None, start_heap G, [([], Null # replicate mxl undefined, C, (m,[]), 0)])"
--- a/src/HOL/MicroJava/JVM/JVMListExample.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/JVM/JVMListExample.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -16,20 +16,19 @@
   val_nam :: vnam
   next_nam :: vnam
 
-constdefs
-  list_name :: cname
+definition list_name :: cname where
   "list_name == Cname list_nam"
   
-  test_name :: cname
+definition test_name :: cname where
   "test_name == Cname test_nam"
 
-  val_name :: vname
+definition val_name :: vname where
   "val_name == VName val_nam"
 
-  next_name :: vname
+definition next_name :: vname where
   "next_name == VName next_nam"
 
-  append_ins :: bytecode
+definition append_ins :: bytecode where
   "append_ins == 
        [Load 0,
         Getfield next_name list_name,
@@ -46,14 +45,14 @@
         LitPush Unit,
         Return]"
 
-  list_class :: "jvm_method class"
+definition list_class :: "jvm_method class" where
   "list_class ==
     (Object,
      [(val_name, PrimT Integer), (next_name, Class list_name)],
      [((append_name, [Class list_name]), PrimT Void,
         (3, 0, append_ins,[(1,2,8,Xcpt NullPointer)]))])"
 
-  make_list_ins :: bytecode
+definition make_list_ins :: bytecode where
   "make_list_ins ==
        [New list_name,
         Dup,
@@ -79,12 +78,12 @@
         Invoke list_name append_name [Class list_name],
         Return]"
 
-  test_class :: "jvm_method class"
+definition test_class :: "jvm_method class" where
   "test_class ==
     (Object, [],
      [((makelist_name, []), PrimT Void, (3, 2, make_list_ins,[]))])"
 
-  E :: jvm_prog
+definition E :: jvm_prog where
   "E == SystemClasses @ [(list_name, list_class), (test_name, test_class)]"
 
 
--- a/src/HOL/MicroJava/JVM/JVMState.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/MicroJava/JVM/JVMState.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -33,8 +33,7 @@
 
 
 section {* Exceptions *}
-constdefs
-  raise_system_xcpt :: "bool \<Rightarrow> xcpt \<Rightarrow> val option"
+definition raise_system_xcpt :: "bool \<Rightarrow> xcpt \<Rightarrow> val option" where
   "raise_system_xcpt b x \<equiv> raise_if b x None"
 
 section {* Runtime State *}
--- a/src/HOL/Modelcheck/CTL.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Modelcheck/CTL.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Modelcheck/CTL.thy
-    ID:         $Id$
     Author:     Olaf Mueller, Jan Philipps, Robert Sandner
     Copyright   1997  TU Muenchen
 *)
@@ -11,10 +10,10 @@
 types
   'a trans  = "('a * 'a) set"
 
-constdefs
-  CEX ::"['a trans,'a pred, 'a]=>bool"
+definition CEX :: "['a trans,'a pred, 'a]=>bool" where
   "CEX N f u == (? v. (f v & (u,v):N))"
-  EG ::"['a trans,'a pred]=> 'a pred"
+
+definition EG ::"['a trans,'a pred]=> 'a pred" where
   "EG N f == nu (% Q. % u.(f u & CEX N Q u))"
 
 end
--- a/src/HOL/Modelcheck/EindhovenExample.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Modelcheck/EindhovenExample.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Modelcheck/EindhovenExample.thy
-    ID:         $Id$
     Author:     Olaf Mueller, Jan Philipps, Robert Sandner
     Copyright   1997  TU Muenchen
 *)
@@ -11,17 +10,16 @@
 types
   state = "bool * bool * bool"
 
-constdefs
-  INIT :: "state pred"
+definition INIT :: "state pred" where
   "INIT x == ~(fst x)&~(fst (snd x))&~(snd (snd x))"
 
-  N :: "[state,state] => bool"
+definition N :: "[state,state] => bool" where
   "N == % (x1,x2,x3) (y1,y2,y3).
       (~x1 & ~x2 & ~x3 &   y1 & ~y2 & ~y3) |
       ( x1 & ~x2 & ~x3 &  ~y1 & ~y2 & ~y3) |
       ( x1 & ~x2 & ~x3 &   y1 &  y2 &  y3)"
 
-  reach:: "state pred"
+definition reach:: "state pred" where
   "reach  == mu (%Q x. INIT x | (? y. Q y & N y x))"
 
 lemma init_state: "INIT (a, b, c) = (~a & ~b &~c)"
--- a/src/HOL/Modelcheck/MuCalculus.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Modelcheck/MuCalculus.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Modelcheck/MuCalculus.thy
-    ID:         $Id$
     Author:     Olaf Mueller, Jan Philipps, Robert Sandner
     Copyright   1997  TU Muenchen
 *)
@@ -11,17 +10,16 @@
 types
  'a pred = "'a=>bool"
 
-constdefs
-  Charfun :: "'a set => 'a pred"
+definition Charfun :: "'a set => 'a pred" where
   "Charfun == (% A.% x. x:A)"
 
-  monoP  :: "('a pred => 'a pred) => bool"
+definition monoP  :: "('a pred => 'a pred) => bool" where
   "monoP f == mono(Collect o f o Charfun)"
 
-  mu :: "('a pred => 'a pred) => 'a pred"    (binder "Mu " 10)
+definition mu :: "('a pred => 'a pred) => 'a pred" (binder "Mu " 10) where
   "mu f == Charfun(lfp(Collect o f o Charfun))"
 
-  nu :: "('a pred => 'a pred) => 'a pred"    (binder "Nu " 10)
+definition nu :: "('a pred => 'a pred) => 'a pred" (binder "Nu " 10) where
   "nu f == Charfun(gfp(Collect o f o Charfun))"
 
 end
--- a/src/HOL/Modelcheck/MuckeExample1.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Modelcheck/MuckeExample1.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -11,8 +11,7 @@
 types
   state = "bool * bool * bool"
 
-constdefs
-  INIT :: "state pred"
+definition INIT :: "state pred" where
   "INIT x ==  ~(fst x)&~(fst (snd x))&~(snd (snd x))"
   N    :: "[state,state] => bool"
   "N x y == let x1 = fst(x); x2 = fst(snd(x)); x3 = snd(snd(x));
--- a/src/HOL/Modelcheck/MuckeExample2.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Modelcheck/MuckeExample2.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,8 +8,7 @@
 imports MuckeSyn
 begin
 
-constdefs
-  Init  :: "bool pred"
+definition Init :: "bool pred" where
   "Init x == x"
   R     :: "[bool,bool] => bool"
   "R x y == (x & ~y) | (~x & y)"
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -27,10 +27,10 @@
 
 parse_translation {*
 let
-  fun cart t u = Syntax.const @{type_name cart} $ t $ u;   (* FIXME @{type_syntax} *)
+  fun cart t u = Syntax.const @{type_syntax cart} $ t $ u;
   fun finite_cart_tr [t, u as Free (x, _)] =
         if Syntax.is_tid x then
-          cart t (Syntax.const @{syntax_const "_ofsort"} $ u $ Syntax.const (hd @{sort finite}))
+          cart t (Syntax.const @{syntax_const "_ofsort"} $ u $ Syntax.const @{class_syntax finite})
         else cart t u
     | finite_cart_tr [t, u] = cart t u
 in
--- a/src/HOL/Mutabelle/mutabelle.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Mutabelle/mutabelle.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -361,7 +361,7 @@
    val t' = canonize_term t comms;
    val u' = canonize_term u comms;
  in 
-   if s mem comms andalso TermOrd.termless (u', t')
+   if s mem comms andalso Term_Ord.termless (u', t')
    then Const (s, T) $ u' $ t'
    else Const (s, T) $ t' $ u'
  end
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -12,9 +12,9 @@
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
 type timing = (string * int) list
 
-type mtd = string * (theory -> term -> outcome * timing)
+type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
 
-type mutant_subentry = term * (string * (outcome * timing)) list
+type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
 type subentry = string * int * int * int * int * int * int
@@ -54,7 +54,7 @@
 
 (* quickcheck options *)
 (*val quickcheck_generator = "SML"*)
-val iterations = 1
+val iterations = 100
 val size = 5
 
 exception RANDOM;
@@ -79,9 +79,9 @@
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
 type timing = (string * int) list
 
-type mtd = string * (theory -> term -> outcome * timing)
+type mtd = string * (theory -> term -> outcome * (timing * (int * Quickcheck.report list) list option))
 
-type mutant_subentry = term * (string * (outcome * timing)) list
+type mutant_subentry = term * (string * (outcome * (timing * Quickcheck.report option))) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
 type subentry = string * int * int * int * int * int * int
@@ -97,11 +97,11 @@
 fun invoke_quickcheck quickcheck_generator thy t =
   TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
       (fn _ =>
-          case Quickcheck.timed_test_term (ProofContext.init thy) false (SOME quickcheck_generator)
+          case Quickcheck.gen_test_term (ProofContext.init thy) true true (SOME quickcheck_generator)
                                     size iterations (preprocess thy [] t) of
-            (NONE, time_report) => (NoCex, time_report)
-          | (SOME _, time_report) => (GenuineCex, time_report)) ()
-  handle TimeLimit.TimeOut => (Timeout, [("timelimit", !Auto_Counterexample.time_limit)])
+            (NONE, (time_report, report)) => (NoCex, (time_report, report))
+          | (SOME _, (time_report, report)) => (GenuineCex, (time_report, report))) ()
+  handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Counterexample.time_limit)], NONE))
 
 fun quickcheck_mtd quickcheck_generator =
   ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
@@ -258,13 +258,13 @@
 fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   let
     val _ = priority ("Invoking " ^ mtd_name)
-    val ((res, timing), time) = cpu_time "total time"
+    val ((res, (timing, reports)), time) = cpu_time "total time"
       (fn () => case try (invoke_mtd thy) t of
-          SOME (res, timing) => (res, timing)
+          SOME (res, (timing, reports)) => (res, (timing, reports))
         | NONE => (priority ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
-           (Error , [])))
+           (Error , ([], NONE))))
     val _ = priority (" Done")
-  in (res, time :: timing) end
+  in (res, (time :: timing, reports)) end
 
 (* theory -> term list -> mtd -> subentry *)
 (*
@@ -341,10 +341,21 @@
   "\n"
 
 fun string_of_mutant_subentry' thy thm_name (t, results) =
-  "mutant of " ^ thm_name ^ ":" ^
-    cat_lines (map (fn (mtd_name, (outcome, timing)) =>
+  let
+    fun string_of_report (Quickcheck.Report {iterations = i, raised_match_errors = e,
+      satisfied_assms = s, positive_concl_tests = p}) =
+      "errors: " ^ string_of_int e ^ "; conclusion tests: " ^ string_of_int p
+    fun string_of_reports NONE = ""
+      | string_of_reports (SOME reports) =
+        cat_lines (map (fn (size, [report]) =>
+          "size " ^ string_of_int size ^ ": " ^ string_of_report report) (rev reports))
+    fun string_of_mtd_result (mtd_name, (outcome, (timing, reports))) =
       mtd_name ^ ": " ^ string_of_outcome outcome ^ "; " ^
-      space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing)) results)
+      space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing)
+      ^ "\n" ^ string_of_reports reports
+  in
+    "mutant of " ^ thm_name ^ ":\n" ^ cat_lines (map string_of_mtd_result results)
+  end
 
 fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
    thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
--- a/src/HOL/NanoJava/Decl.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/NanoJava/Decl.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -50,11 +50,10 @@
   Object  :: cname      --{* name of root class *}
 
 
-constdefs
- "class"     :: "cname \<rightharpoonup> class"
+definition "class" :: "cname \<rightharpoonup> class" where
  "class      \<equiv> map_of Prog"
 
-  is_class   :: "cname => bool"
+definition is_class   :: "cname => bool" where
  "is_class C \<equiv> class C \<noteq> None"
 
 lemma finite_is_class: "finite {C. is_class C}"
--- a/src/HOL/NanoJava/Equivalence.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/NanoJava/Equivalence.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/NanoJava/Equivalence.thy
-    ID:         $Id$
     Author:     David von Oheimb
     Copyright   2001 Technische Universitaet Muenchen
 *)
@@ -10,37 +9,35 @@
 
 subsection "Validity"
 
-constdefs
-  valid   :: "[assn,stmt, assn] => bool"  ("|= {(1_)}/ (_)/ {(1_)}" [3,90,3] 60)
+definition valid :: "[assn,stmt, assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where
  "|=  {P} c {Q} \<equiv> \<forall>s   t. P s --> (\<exists>n. s -c  -n\<rightarrow> t) --> Q   t"
 
- evalid   :: "[assn,expr,vassn] => bool" ("|=e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60)
+definition evalid   :: "[assn,expr,vassn] => bool" ("|=e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where
  "|=e {P} e {Q} \<equiv> \<forall>s v t. P s --> (\<exists>n. s -e\<succ>v-n\<rightarrow> t) --> Q v t"
 
-
- nvalid   :: "[nat, triple    ] => bool" ("|=_: _"  [61,61] 60)
+definition nvalid   :: "[nat, triple    ] => bool" ("|=_: _"  [61,61] 60) where
  "|=n:  t \<equiv> let (P,c,Q) = t in \<forall>s   t. s -c  -n\<rightarrow> t --> P s --> Q   t"
 
-envalid   :: "[nat,etriple    ] => bool" ("|=_:e _" [61,61] 60)
+definition envalid   :: "[nat,etriple    ] => bool" ("|=_:e _" [61,61] 60) where
  "|=n:e t \<equiv> let (P,e,Q) = t in \<forall>s v t. s -e\<succ>v-n\<rightarrow> t --> P s --> Q v t"
 
-  nvalids :: "[nat,       triple set] => bool" ("||=_: _" [61,61] 60)
+definition nvalids :: "[nat,       triple set] => bool" ("||=_: _" [61,61] 60) where
  "||=n: T \<equiv> \<forall>t\<in>T. |=n: t"
 
- cnvalids :: "[triple set,triple set] => bool" ("_ ||=/ _"  [61,61] 60)
+definition cnvalids :: "[triple set,triple set] => bool" ("_ ||=/ _"  [61,61] 60) where
  "A ||=  C \<equiv> \<forall>n. ||=n: A --> ||=n: C"
 
-cenvalid  :: "[triple set,etriple   ] => bool" ("_ ||=e/ _" [61,61] 60)
+definition cenvalid  :: "[triple set,etriple   ] => bool" ("_ ||=e/ _" [61,61] 60) where
  "A ||=e t \<equiv> \<forall>n. ||=n: A --> |=n:e t"
 
-syntax (xsymbols)
-   valid  :: "[assn,stmt, assn] => bool" ( "\<Turnstile> {(1_)}/ (_)/ {(1_)}" [3,90,3] 60)
-  evalid  :: "[assn,expr,vassn] => bool" ("\<Turnstile>\<^sub>e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60)
-  nvalid  :: "[nat, triple          ] => bool" ("\<Turnstile>_: _"  [61,61] 60)
- envalid  :: "[nat,etriple          ] => bool" ("\<Turnstile>_:\<^sub>e _" [61,61] 60)
-  nvalids :: "[nat,       triple set] => bool" ("|\<Turnstile>_: _"  [61,61] 60)
- cnvalids :: "[triple set,triple set] => bool" ("_ |\<Turnstile>/ _" [61,61] 60)
-cenvalid  :: "[triple set,etriple   ] => bool" ("_ |\<Turnstile>\<^sub>e/ _"[61,61] 60)
+notation (xsymbols)
+  valid  ("\<Turnstile> {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) and
+  evalid  ("\<Turnstile>\<^sub>e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) and
+  nvalid  ("\<Turnstile>_: _" [61,61] 60) and
+  envalid  ("\<Turnstile>_:\<^sub>e _" [61,61] 60) and
+  nvalids  ("|\<Turnstile>_: _" [61,61] 60) and
+  cnvalids  ("_ |\<Turnstile>/ _" [61,61] 60) and
+  cenvalid  ("_ |\<Turnstile>\<^sub>e/ _"[61,61] 60)
 
 
 lemma nvalid_def2: "\<Turnstile>n: (P,c,Q) \<equiv> \<forall>s t. s -c-n\<rightarrow> t \<longrightarrow> P s \<longrightarrow> Q t"
@@ -160,14 +157,16 @@
 
 subsection "(Relative) Completeness"
 
-constdefs MGT    :: "stmt => state =>  triple"
+definition MGT :: "stmt => state => triple" where
          "MGT  c Z \<equiv> (\<lambda>s. Z = s, c, \<lambda>  t. \<exists>n. Z -c-  n\<rightarrow> t)"
-          MGTe   :: "expr => state => etriple"
+
+definition MGTe   :: "expr => state => etriple" where
          "MGTe e Z \<equiv> (\<lambda>s. Z = s, e, \<lambda>v t. \<exists>n. Z -e\<succ>v-n\<rightarrow> t)"
-syntax (xsymbols)
-         MGTe    :: "expr => state => etriple" ("MGT\<^sub>e")
-syntax (HTML output)
-         MGTe    :: "expr => state => etriple" ("MGT\<^sub>e")
+
+notation (xsymbols)
+  MGTe  ("MGT\<^sub>e")
+notation (HTML output)
+  MGTe  ("MGT\<^sub>e")
 
 lemma MGF_implies_complete:
  "\<forall>Z. {} |\<turnstile> { MGT c Z} \<Longrightarrow> \<Turnstile>  {P} c {Q} \<Longrightarrow> {} \<turnstile>  {P} c {Q}"
--- a/src/HOL/NanoJava/State.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/NanoJava/State.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,9 +7,7 @@
 
 theory State imports TypeRel begin
 
-constdefs
-
-  body :: "cname \<times> mname => stmt"
+definition body :: "cname \<times> mname => stmt" where
  "body \<equiv> \<lambda>(C,m). bdy (the (method C m))"
 
 text {* Locations, i.e.\ abstract references to objects *}
@@ -29,9 +27,7 @@
   "fields" \<leftharpoondown> (type)"fname => val option"
   "obj"    \<leftharpoondown> (type)"cname \<times> fields"
 
-constdefs
-
-  init_vars:: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> val)"
+definition init_vars :: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> val)" where
  "init_vars m == Option.map (\<lambda>T. Null) o m"
   
 text {* private: *}
@@ -49,54 +45,49 @@
   "locals" \<leftharpoondown> (type)"vname => val option"
   "state" \<leftharpoondown> (type)"(|heap :: heap, locals :: locals|)"
 
-constdefs
-
-  del_locs     :: "state => state"
+definition del_locs :: "state => state" where
  "del_locs s \<equiv> s (| locals := empty |)"
 
-  init_locs     :: "cname => mname => state => state"
+definition init_locs     :: "cname => mname => state => state" where
  "init_locs C m s \<equiv> s (| locals := locals s ++ 
                          init_vars (map_of (lcl (the (method C m)))) |)"
 
 text {* The first parameter of @{term set_locs} is of type @{typ state} 
         rather than @{typ locals} in order to keep @{typ locals} private.*}
-constdefs
-  set_locs  :: "state => state => state"
+definition set_locs :: "state => state => state" where
  "set_locs s s' \<equiv> s' (| locals := locals s |)"
 
-  get_local     :: "state => vname => val" ("_<_>" [99,0] 99)
+definition get_local     :: "state => vname => val" ("_<_>" [99,0] 99) where
  "get_local s x  \<equiv> the (locals s x)"
 
 --{* local function: *}
-  get_obj       :: "state => loc => obj"
+definition get_obj       :: "state => loc => obj" where
  "get_obj s a \<equiv> the (heap s a)"
 
-  obj_class     :: "state => loc => cname"
+definition obj_class     :: "state => loc => cname" where
  "obj_class s a \<equiv> fst (get_obj s a)"
 
-  get_field     :: "state => loc => fname => val"
+definition get_field     :: "state => loc => fname => val" where
  "get_field s a f \<equiv> the (snd (get_obj s a) f)"
 
 --{* local function: *}
-  hupd       :: "loc => obj => state => state"   ("hupd'(_|->_')" [10,10] 1000)
+definition hupd       :: "loc => obj => state => state"   ("hupd'(_|->_')" [10,10] 1000) where
  "hupd a obj s \<equiv> s (| heap   := ((heap   s)(a\<mapsto>obj))|)"
 
-  lupd       :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000)
+definition lupd       :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000) where
  "lupd x v s   \<equiv> s (| locals := ((locals s)(x\<mapsto>v  ))|)"
 
-syntax (xsymbols)
-  hupd       :: "loc => obj => state => state"   ("hupd'(_\<mapsto>_')" [10,10] 1000)
-  lupd       :: "vname => val => state => state" ("lupd'(_\<mapsto>_')" [10,10] 1000)
+notation (xsymbols)
+  hupd  ("hupd'(_\<mapsto>_')" [10,10] 1000) and
+  lupd  ("lupd'(_\<mapsto>_')" [10,10] 1000)
 
-constdefs
-
-  new_obj    :: "loc => cname => state => state"
+definition new_obj :: "loc => cname => state => state" where
  "new_obj a C   \<equiv> hupd(a\<mapsto>(C,init_vars (field C)))"
 
-  upd_obj    :: "loc => fname => val => state => state"
+definition upd_obj    :: "loc => fname => val => state => state" where
  "upd_obj a f v s \<equiv> let (C,fs) = the (heap s a) in hupd(a\<mapsto>(C,fs(f\<mapsto>v))) s"
 
-  new_Addr      :: "state => val"
+definition new_Addr      :: "state => val" where
  "new_Addr s == SOME v. (\<exists>a. v = Addr a \<and> (heap s) a = None) | v = Null"
 
 
--- a/src/HOL/NanoJava/TypeRel.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/NanoJava/TypeRel.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -66,9 +66,7 @@
 apply  auto
 done
 
-constdefs
-
-  ws_prog  :: "bool"
+definition ws_prog :: "bool" where
  "ws_prog \<equiv> \<forall>(C,c)\<in>set Prog. C\<noteq>Object \<longrightarrow> 
                               is_class (super c) \<and> (super c,C)\<notin>subcls1^+"
 
--- a/src/HOL/Nat.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -45,8 +45,7 @@
   nat = Nat
   by (rule exI, unfold mem_def, rule Nat.Zero_RepI)
 
-constdefs
-  Suc ::   "nat => nat"
+definition Suc :: "nat => nat" where
   Suc_def: "Suc == (%n. Abs_Nat (Suc_Rep (Rep_Nat n)))"
 
 local
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1000,7 +1000,7 @@
 oops
 
 lemma "(Q\<Colon>nat set) (Eps Q)"
-nitpick [expect = none]
+nitpick [expect = none] (* unfortunate *)
 oops
 
 lemma "\<not> Q (Eps Q)"
@@ -1053,11 +1053,11 @@
 
 lemma "Q = {x\<Colon>'a} \<Longrightarrow> (Q\<Colon>'a set) (The Q)"
 nitpick [expect = none]
-oops
+sorry
 
 lemma "Q = {x\<Colon>nat} \<Longrightarrow> (Q\<Colon>nat set) (The Q)"
 nitpick [expect = none]
-oops
+sorry
 
 subsection {* Destructors and Recursors *}
 
--- a/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -69,7 +69,7 @@
 oops
 
 lemma "fs (Pd ((a, b), (c, d))) = (a, b)"
-nitpick [card = 1\<midarrow>10, expect = none]
+nitpick [card = 1\<midarrow>9, expect = none]
 sorry
 
 lemma "fs (Pd ((a, b), (c, d))) = (c, d)"
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -30,8 +30,7 @@
    special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [],
    wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []}
 (* term -> bool *)
-val is_mono = Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a}
-                                              Nitpick_Mono.Plus [] []
+fun is_mono t = Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a} ([t], [])
 fun is_const t =
   let val T = fastype_of t in
     is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t),
@@ -47,26 +46,26 @@
 ML {* const @{term "(A::'a set) = A"} *}
 ML {* const @{term "(A::'a set set) = A"} *}
 ML {* const @{term "(\<lambda>x::'a set. x a)"} *}
-ML {* const @{term "{{a}} = C"} *}
+ML {* const @{term "{{a::'a}} = C"} *}
 ML {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *}
-ML {* const @{term "A \<union> B"} *}
+ML {* const @{term "A \<union> (B::'a set)"} *}
 ML {* const @{term "P (a::'a)"} *}
 ML {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *}
 ML {* const @{term "\<forall>A::'a set. A a"} *}
 ML {* const @{term "\<forall>A::'a set. P A"} *}
 ML {* const @{term "P \<or> Q"} *}
-ML {* const @{term "A \<union> B = C"} *}
+ML {* const @{term "A \<union> B = (C::'a set)"} *}
 ML {* const @{term "(if P then (A::'a set) else B) = C"} *}
-ML {* const @{term "let A = C in A \<union> B"} *}
+ML {* const @{term "let A = (C::'a set) in A \<union> B"} *}
 ML {* const @{term "THE x::'b. P x"} *}
-ML {* const @{term "{}::'a set"} *}
+ML {* const @{term "(\<lambda>x::'a. False)"} *}
 ML {* const @{term "(\<lambda>x::'a. True)"} *}
-ML {* const @{term "Let a A"} *}
+ML {* const @{term "Let (a::'a) A"} *}
 ML {* const @{term "A (a::'a)"} *}
-ML {* const @{term "insert a A = B"} *}
+ML {* const @{term "insert (a::'a) A = B"} *}
 ML {* const @{term "- (A::'a set)"} *}
-ML {* const @{term "finite A"} *}
-ML {* const @{term "\<not> finite A"} *}
+ML {* const @{term "finite (A::'a set)"} *}
+ML {* const @{term "\<not> finite (A::'a set)"} *}
 ML {* const @{term "finite (A::'a set set)"} *}
 ML {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *}
 ML {* const @{term "A < (B::'a set)"} *}
@@ -75,27 +74,25 @@
 ML {* const @{term "[a::'a set]"} *}
 ML {* const @{term "[A \<union> (B::'a set)]"} *}
 ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
-ML {* const @{term "\<forall>P. P a"} *}
 
-ML {* nonconst @{term "{%x. True}"} *}
-ML {* nonconst @{term "{(%x. x = a)} = C"} *}
+ML {* nonconst @{term "{(\<lambda>x::'a. x = a)} = C"} *}
 ML {* nonconst @{term "\<forall>P (a::'a). P a"} *}
 ML {* nonconst @{term "\<forall>a::'a. P a"} *}
 ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
-ML {* nonconst @{term "THE x. P x"} *}
-ML {* nonconst @{term "SOME x. P x"} *}
+ML {* nonconst @{term "THE x::'a. P x"} *}
+ML {* nonconst @{term "SOME x::'a. P x"} *}
 
 ML {* mono @{prop "Q (\<forall>x::'a set. P x)"} *}
 ML {* mono @{prop "P (a::'a)"} *}
-ML {* mono @{prop "{a} = {b}"} *}
+ML {* mono @{prop "{a} = {b::'a}"} *}
 ML {* mono @{prop "P (a::'a) \<and> P \<union> P = P"} *}
 ML {* mono @{prop "\<forall>F::'a set set. P"} *}
 ML {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h)"} *}
 ML {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *}
-ML {* mono @{prop "\<not> (\<forall>x. P x)"} *}
+ML {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *}
 
-ML {* nonmono @{prop "\<forall>x. P x"} *}
+ML {* nonmono @{prop "\<forall>x::'a. P x"} *}
+ML {* nonmono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *}
 ML {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *}
-ML {* nonmono @{prop "myall P = (P = (\<lambda>x. True))"} *}
 
 end
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -244,12 +244,15 @@
 
 text {* ``The transitive closure of an arbitrary relation is non-empty.'' *}
 
-constdefs
-"trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 "trans P \<equiv> (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)"
-"subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+
+definition
+"subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 "subset P Q \<equiv> (ALL x y. P x y \<longrightarrow> Q x y)"
-"trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+
+definition
+"trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
 "trans_closure P Q \<equiv> (subset Q P) \<and> (trans P) \<and> (ALL R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)"
 
 lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)"
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,7 +8,7 @@
 header {* Examples Featuring Nitpick Applied to Typedefs *}
 
 theory Typedef_Nits
-imports Main Rational
+imports Complex_Main
 begin
 
 nitpick_params [card = 1\<midarrow>4, sat_solver = MiniSat_JNI, max_threads = 1,
@@ -64,7 +64,7 @@
 oops
 
 lemma "x \<noteq> (y\<Colon>bool bounded) \<Longrightarrow> z = x \<or> z = y"
-nitpick [expect = none]
+nitpick [fast_descrs (* ### FIXME *), expect = none]
 sorry
 
 lemma "x \<noteq> (y\<Colon>(bool \<times> bool) bounded) \<Longrightarrow> z = x \<or> z = y"
--- a/src/HOL/Nominal/Examples/Class.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/Examples/Class.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -10289,17 +10289,16 @@
 
 text {* set operators *}
 
-constdefs
-  AXIOMSn::"ty \<Rightarrow> ntrm set"
+definition AXIOMSn :: "ty \<Rightarrow> ntrm set" where
   "AXIOMSn B \<equiv> { (x):(Ax y b) | x y b. True }"
 
-  AXIOMSc::"ty \<Rightarrow> ctrm set"
+definition AXIOMSc::"ty \<Rightarrow> ctrm set" where
   "AXIOMSc B \<equiv> { <a>:(Ax y b) | a y b. True }"
 
-  BINDINGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
+definition BINDINGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set" where
   "BINDINGn B X \<equiv> { (x):M | x M. \<forall>a P. <a>:P\<in>X \<longrightarrow> SNa (M{x:=<a>.P})}"
 
-  BINDINGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
+definition BINDINGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set" where
   "BINDINGc B X \<equiv> { <a>:M | a M. \<forall>x P. (x):P\<in>X \<longrightarrow> SNa (M{a:=(x).P})}"
 
 lemma BINDINGn_decreasing:
@@ -16540,8 +16539,7 @@
 apply(fast)
 done
 
-constdefs
-  SNa_Redu :: "(trm \<times> trm) set"
+definition SNa_Redu :: "(trm \<times> trm) set" where
   "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV <*> SNa_set)"
 
 lemma wf_SNa_Redu:
--- a/src/HOL/Nominal/Examples/Fsub.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/Examples/Fsub.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -223,8 +223,7 @@
   in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the 
   @{text "support"} of @{term "S"}. *}
 
-constdefs
-  "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100)
+definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
   "S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)"
 
 lemma closed_in_eqvt[eqvt]:
@@ -718,8 +717,7 @@
   another. This generalization seems to make the proof for weakening to be
   smoother than if we had strictly adhered to the version in the POPLmark-paper. *}
 
-constdefs 
-  extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100)
+definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
   "\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>"
 
 lemma extends_ty_dom:
--- a/src/HOL/Nominal/Examples/LocalWeakening.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/Examples/LocalWeakening.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -58,8 +58,7 @@
 by (induct t arbitrary: n rule: llam.induct)
    (simp_all add: perm_nat_def)
 
-constdefs
-  freshen :: "llam \<Rightarrow> name \<Rightarrow> llam"
+definition freshen :: "llam \<Rightarrow> name \<Rightarrow> llam" where
   "freshen t p \<equiv> vsub t 0 (lPar p)"
 
 lemma freshen_eqvt[eqvt]:
--- a/src/HOL/Nominal/Examples/SN.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/Examples/SN.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,5 @@
 theory SN
-  imports Lam_Funs
+imports Lam_Funs
 begin
 
 text {* Strong Normalisation proof from the Proofs and Types book *}
@@ -158,8 +158,7 @@
 
 subsection {* a fact about beta *}
 
-constdefs
-  "NORMAL" :: "lam \<Rightarrow> bool"
+definition "NORMAL" :: "lam \<Rightarrow> bool" where
   "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')"
 
 lemma NORMAL_Var:
@@ -234,8 +233,7 @@
 by (rule TrueI)+
 
 text {* neutral terms *}
-constdefs
-  NEUT :: "lam \<Rightarrow> bool"
+definition NEUT :: "lam \<Rightarrow> bool" where
   "NEUT t \<equiv> (\<exists>a. t = Var a) \<or> (\<exists>t1 t2. t = App t1 t2)" 
 
 (* a slight hack to get the first element of applications *)
@@ -274,20 +272,19 @@
 
 section {* Candidates *}
 
-constdefs
-  "CR1" :: "ty \<Rightarrow> bool"
+definition "CR1" :: "ty \<Rightarrow> bool" where
   "CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)"
 
-  "CR2" :: "ty \<Rightarrow> bool"
+definition "CR2" :: "ty \<Rightarrow> bool" where
   "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>"
 
-  "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool"
+definition "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool" where
   "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow>  t'\<in>RED \<tau>" 
 
-  "CR3" :: "ty \<Rightarrow> bool"
+definition "CR3" :: "ty \<Rightarrow> bool" where
   "CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>"
    
-  "CR4" :: "ty \<Rightarrow> bool"
+definition "CR4" :: "ty \<Rightarrow> bool" where
   "CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>"
 
 lemma CR3_implies_CR4: 
--- a/src/HOL/Nominal/Nominal.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/Nominal.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -31,7 +31,7 @@
 
 (* an auxiliary constant for the decision procedure involving *) 
 (* permutations (to avoid loops when using perm-compositions)  *)
-constdefs
+definition
   "perm_aux pi x \<equiv> pi\<bullet>x"
 
 (* overloaded permutation operations *)
@@ -187,20 +187,18 @@
 section {* permutation equality *}
 (*==============================*)
 
-constdefs
-  prm_eq :: "'x prm \<Rightarrow> 'x prm \<Rightarrow> bool"  (" _ \<triangleq> _ " [80,80] 80)
+definition prm_eq :: "'x prm \<Rightarrow> 'x prm \<Rightarrow> bool" (" _ \<triangleq> _ " [80,80] 80) where
   "pi1 \<triangleq> pi2 \<equiv> \<forall>a::'x. pi1\<bullet>a = pi2\<bullet>a"
 
 section {* Support, Freshness and Supports*}
 (*========================================*)
-constdefs
-   supp :: "'a \<Rightarrow> ('x set)"  
+definition supp :: "'a \<Rightarrow> ('x set)" where  
    "supp x \<equiv> {a . (infinite {b . [(a,b)]\<bullet>x \<noteq> x})}"
 
-   fresh :: "'x \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [80,80] 80)
+definition fresh :: "'x \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [80,80] 80) where
    "a \<sharp> x \<equiv> a \<notin> supp x"
 
-   supports :: "'x set \<Rightarrow> 'a \<Rightarrow> bool" (infixl "supports" 80)
+definition supports :: "'x set \<Rightarrow> 'a \<Rightarrow> bool" (infixl "supports" 80) where
    "S supports x \<equiv> \<forall>a b. (a\<notin>S \<and> b\<notin>S \<longrightarrow> [(a,b)]\<bullet>x=x)"
 
 (* lemmas about supp *)
@@ -400,14 +398,14 @@
 (*=========================================================*)
 
 (* properties for being a permutation type *)
-constdefs 
+definition
   "pt TYPE('a) TYPE('x) \<equiv> 
      (\<forall>(x::'a). ([]::'x prm)\<bullet>x = x) \<and> 
      (\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). (pi1@pi2)\<bullet>x = pi1\<bullet>(pi2\<bullet>x)) \<and> 
      (\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). pi1 \<triangleq> pi2 \<longrightarrow> pi1\<bullet>x = pi2\<bullet>x)"
 
 (* properties for being an atom type *)
-constdefs 
+definition
   "at TYPE('x) \<equiv> 
      (\<forall>(x::'x). ([]::'x prm)\<bullet>x = x) \<and>
      (\<forall>(a::'x) (b::'x) (pi::'x prm) (x::'x). ((a,b)#(pi::'x prm))\<bullet>x = swap (a,b) (pi\<bullet>x)) \<and> 
@@ -415,18 +413,18 @@
      (infinite (UNIV::'x set))"
 
 (* property of two atom-types being disjoint *)
-constdefs
+definition
   "disjoint TYPE('x) TYPE('y) \<equiv> 
        (\<forall>(pi::'x prm)(x::'y). pi\<bullet>x = x) \<and> 
        (\<forall>(pi::'y prm)(x::'x). pi\<bullet>x = x)"
 
 (* composition property of two permutation on a type 'a *)
-constdefs
+definition
   "cp TYPE ('a) TYPE('x) TYPE('y) \<equiv> 
       (\<forall>(pi2::'y prm) (pi1::'x prm) (x::'a) . pi1\<bullet>(pi2\<bullet>x) = (pi1\<bullet>pi2)\<bullet>(pi1\<bullet>x))" 
 
 (* property of having finite support *)
-constdefs 
+definition
   "fs TYPE('a) TYPE('x) \<equiv> \<forall>(x::'a). finite ((supp x)::'x set)"
 
 section {* Lemmas about the atom-type properties*}
@@ -2216,8 +2214,7 @@
 section {* Facts about the support of finite sets of finitely supported things *}
 (*=============================================================================*)
 
-constdefs
-  X_to_Un_supp :: "('a set) \<Rightarrow> 'x set"
+definition X_to_Un_supp :: "('a set) \<Rightarrow> 'x set" where
   "X_to_Un_supp X \<equiv> \<Union>x\<in>X. ((supp x)::'x set)"
 
 lemma UNION_f_eqvt:
@@ -2838,8 +2835,7 @@
 qed
 
 -- "packaging the freshness lemma into a function"
-constdefs
-  fresh_fun :: "('x\<Rightarrow>'a)\<Rightarrow>'a"
+definition fresh_fun :: "('x\<Rightarrow>'a)\<Rightarrow>'a" where
   "fresh_fun (h) \<equiv> THE fr. (\<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr)"
 
 lemma fresh_fun_app:
@@ -2970,8 +2966,7 @@
   shows "pt TYPE('x\<Rightarrow>('a noption)) TYPE('x)"
   by (rule pt_fun_inst[OF at_pt_inst[OF at],OF pt_noption_inst[OF pt],OF at])
 
-constdefs
-  abs_fun :: "'x\<Rightarrow>'a\<Rightarrow>('x\<Rightarrow>('a noption))" ("[_]._" [100,100] 100)
+definition abs_fun :: "'x\<Rightarrow>'a\<Rightarrow>('x\<Rightarrow>('a noption))" ("[_]._" [100,100] 100) where 
   "[a].x \<equiv> (\<lambda>b. (if b=a then nSome(x) else (if b\<sharp>x then nSome([(a,b)]\<bullet>x) else nNone)))"
 
 (* FIXME: should be called perm_if and placed close to the definition of permutations on bools *)
@@ -3403,13 +3398,13 @@
   where
   ABS_in: "(abs_fun a x)\<in>ABS_set"
 
-typedef (ABS) ('x,'a) ABS = "ABS_set::('x\<Rightarrow>('a noption)) set"
+typedef (ABS) ('x,'a) ABS ("\<guillemotleft>_\<guillemotright>_" [1000,1000] 1000) =
+  "ABS_set::('x\<Rightarrow>('a noption)) set"
 proof 
   fix x::"'a" and a::"'x"
   show "(abs_fun a x)\<in> ABS_set" by (rule ABS_in)
 qed
 
-syntax ABS :: "type \<Rightarrow> type \<Rightarrow> type" ("\<guillemotleft>_\<guillemotright>_" [1000,1000] 1000)
 
 section {* lemmas for deciding permutation equations *}
 (*===================================================*)
--- a/src/HOL/Nominal/nominal_datatype.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/nominal_datatype.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -1016,7 +1016,7 @@
               (HOLogic.mk_Trueprop (HOLogic.mk_eq
                 (supp c,
                  if null dts then HOLogic.mk_set atomT []
-                 else foldr1 (HOLogic.mk_binop @{const_name union}) (map supp args2)))))
+                 else foldr1 (HOLogic.mk_binop @{const_abbrev union}) (map supp args2)))))
             (fn _ =>
               simp_tac (HOL_basic_ss addsimps (supp_def ::
                  Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
@@ -2079,7 +2079,7 @@
 local structure P = OuterParse and K = OuterKeyword in
 
 val datatype_decl =
-  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
+  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_mixfix --
     (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
 
 fun mk_datatype args =
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -182,7 +182,7 @@
      (Scan.succeed false);
 
 fun setup_generate_fresh x =
-  (Args.goal_spec -- Args.tyname >>
+  (Args.goal_spec -- Args.type_name true >>
     (fn (quant, s) => K (SIMPLE_METHOD'' quant (generate_fresh_tac s)))) x;
 
 fun setup_fresh_fun_simp x =
--- a/src/HOL/Number_Theory/MiscAlgebra.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Number_Theory/MiscAlgebra.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -31,8 +31,7 @@
 *)
 
 
-constdefs 
-  units_of :: "('a, 'b) monoid_scheme => 'a monoid"
+definition units_of :: "('a, 'b) monoid_scheme => 'a monoid" where
   "units_of G == (| carrier = Units G,
      Group.monoid.mult = Group.monoid.mult G,
      one  = one G |)";
--- a/src/HOL/Number_Theory/Residues.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Number_Theory/Residues.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -22,8 +22,7 @@
 
 *)
 
-constdefs 
-  residue_ring :: "int => int ring"
+definition residue_ring :: "int => int ring" where
   "residue_ring m == (| 
     carrier =       {0..m - 1}, 
     mult =          (%x y. (x * y) mod m),
@@ -287,8 +286,7 @@
 
 (* the definition of the phi function *)
 
-constdefs
-  phi :: "int => nat"
+definition phi :: "int => nat" where
   "phi m == card({ x. 0 < x & x < m & gcd x m = 1})" 
 
 lemma phi_zero [simp]: "phi 0 = 0"
--- a/src/HOL/Number_Theory/UniqueFactorization.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Number_Theory/UniqueFactorization.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -67,8 +67,7 @@
    "ALL i :# M. P i"? 
 *)
 
-constdefs
-  msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b"
+definition msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b" where
   "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)"
 
 syntax
@@ -214,8 +213,7 @@
   thus ?thesis by (simp add:multiset_eq_conv_count_eq)
 qed
 
-constdefs
-  multiset_prime_factorization :: "nat => nat multiset"
+definition multiset_prime_factorization :: "nat => nat multiset" where
   "multiset_prime_factorization n ==
      if n > 0 then (THE M. ((ALL p : set_of M. prime p) & 
        n = (PROD i :# M. i)))
--- a/src/HOL/Old_Number_Theory/EulerFermat.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Old_Number_Theory/EulerFermat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -25,20 +25,18 @@
   | insert: "A \<in> RsetR m ==> zgcd a m = 1 ==>
       \<forall>a'. a' \<in> A --> \<not> zcong a a' m ==> insert a A \<in> RsetR m"
 
-consts
-  BnorRset :: "int * int => int set"
-
-recdef BnorRset
-  "measure ((\<lambda>(a, m). nat a) :: int * int => nat)"
-  "BnorRset (a, m) =
+fun
+  BnorRset :: "int \<Rightarrow> int => int set"
+where
+  "BnorRset a m =
    (if 0 < a then
-    let na = BnorRset (a - 1, m)
+    let na = BnorRset (a - 1) m
     in (if zgcd a m = 1 then insert a na else na)
     else {})"
 
 definition
   norRRset :: "int => int set" where
-  "norRRset m = BnorRset (m - 1, m)"
+  "norRRset m = BnorRset (m - 1) m"
 
 definition
   noXRRset :: "int => int => int set" where
@@ -74,28 +72,27 @@
 
 lemma BnorRset_induct:
   assumes "!!a m. P {} a m"
-    and "!!a m. 0 < (a::int) ==> P (BnorRset (a - 1, m::int)) (a - 1) m
-      ==> P (BnorRset(a,m)) a m"
-  shows "P (BnorRset(u,v)) u v"
+    and "!!a m :: int. 0 < a ==> P (BnorRset (a - 1) m) (a - 1) m
+      ==> P (BnorRset a m) a m"
+  shows "P (BnorRset u v) u v"
   apply (rule BnorRset.induct)
-  apply safe
-   apply (case_tac [2] "0 < a")
-    apply (rule_tac [2] prems)
+   apply (case_tac "0 < a")
+    apply (rule_tac assms)
      apply simp_all
-   apply (simp_all add: BnorRset.simps prems)
+   apply (simp_all add: BnorRset.simps assms)
   done
 
-lemma Bnor_mem_zle [rule_format]: "b \<in> BnorRset (a, m) \<longrightarrow> b \<le> a"
+lemma Bnor_mem_zle [rule_format]: "b \<in> BnorRset a m \<longrightarrow> b \<le> a"
   apply (induct a m rule: BnorRset_induct)
    apply simp
   apply (subst BnorRset.simps)
    apply (unfold Let_def, auto)
   done
 
-lemma Bnor_mem_zle_swap: "a < b ==> b \<notin> BnorRset (a, m)"
+lemma Bnor_mem_zle_swap: "a < b ==> b \<notin> BnorRset a m"
   by (auto dest: Bnor_mem_zle)
 
-lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> 0 < b"
+lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset a m --> 0 < b"
   apply (induct a m rule: BnorRset_induct)
    prefer 2
    apply (subst BnorRset.simps)
@@ -103,7 +100,7 @@
   done
 
 lemma Bnor_mem_if [rule_format]:
-    "zgcd b m = 1 --> 0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
+    "zgcd b m = 1 --> 0 < b --> b \<le> a --> b \<in> BnorRset a m"
   apply (induct a m rule: BnorRset.induct, auto)
    apply (subst BnorRset.simps)
    defer
@@ -111,7 +108,7 @@
    apply (unfold Let_def, auto)
   done
 
-lemma Bnor_in_RsetR [rule_format]: "a < m --> BnorRset (a, m) \<in> RsetR m"
+lemma Bnor_in_RsetR [rule_format]: "a < m --> BnorRset a m \<in> RsetR m"
   apply (induct a m rule: BnorRset_induct, simp)
   apply (subst BnorRset.simps)
   apply (unfold Let_def, auto)
@@ -124,7 +121,7 @@
         apply (rule_tac [5] Bnor_mem_zg, auto)
   done
 
-lemma Bnor_fin: "finite (BnorRset (a, m))"
+lemma Bnor_fin: "finite (BnorRset a m)"
   apply (induct a m rule: BnorRset_induct)
    prefer 2
    apply (subst BnorRset.simps)
@@ -258,8 +255,8 @@
 by (unfold inj_on_def, auto)
 
 lemma Bnor_prod_power [rule_format]:
-  "x \<noteq> 0 ==> a < m --> \<Prod>((\<lambda>a. a * x) ` BnorRset (a, m)) =
-      \<Prod>(BnorRset(a, m)) * x^card (BnorRset (a, m))"
+  "x \<noteq> 0 ==> a < m --> \<Prod>((\<lambda>a. a * x) ` BnorRset a m) =
+      \<Prod>(BnorRset a m) * x^card (BnorRset a m)"
   apply (induct a m rule: BnorRset_induct)
    prefer 2
    apply (simplesubst BnorRset.simps)  --{*multiple redexes*}
@@ -284,7 +281,7 @@
   done
 
 lemma Bnor_prod_zgcd [rule_format]:
-    "a < m --> zgcd (\<Prod>(BnorRset(a, m))) m = 1"
+    "a < m --> zgcd (\<Prod>(BnorRset a m)) m = 1"
   apply (induct a m rule: BnorRset_induct)
    prefer 2
    apply (subst BnorRset.simps)
@@ -299,13 +296,13 @@
   apply (case_tac "x = 0")
    apply (case_tac [2] "m = 1")
     apply (rule_tac [3] iffD1)
-     apply (rule_tac [3] k = "\<Prod>(BnorRset(m - 1, m))"
+     apply (rule_tac [3] k = "\<Prod>(BnorRset (m - 1) m)"
        in zcong_cancel2)
       prefer 5
       apply (subst Bnor_prod_power [symmetric])
         apply (rule_tac [7] Bnor_prod_zgcd, simp_all)
   apply (rule bijzcong_zcong_prod)
-  apply (fold norRRset_def noXRRset_def)
+  apply (fold norRRset_def, fold noXRRset_def)
   apply (subst RRset2norRR_eq_norR [symmetric])
     apply (rule_tac [3] inj_func_bijR, auto)
      apply (unfold zcongm_def)
@@ -319,12 +316,12 @@
   done
 
 lemma Bnor_prime:
-  "\<lbrakk> zprime p; a < p \<rbrakk> \<Longrightarrow> card (BnorRset (a, p)) = nat a"
+  "\<lbrakk> zprime p; a < p \<rbrakk> \<Longrightarrow> card (BnorRset a p) = nat a"
   apply (induct a p rule: BnorRset.induct)
   apply (subst BnorRset.simps)
   apply (unfold Let_def, auto simp add:zless_zprime_imp_zrelprime)
-  apply (subgoal_tac "finite (BnorRset (a - 1,m))")
-   apply (subgoal_tac "a ~: BnorRset (a - 1,m)")
+  apply (subgoal_tac "finite (BnorRset (a - 1) m)")
+   apply (subgoal_tac "a ~: BnorRset (a - 1) m")
     apply (auto simp add: card_insert_disjoint Suc_nat_eq_nat_zadd1)
    apply (frule Bnor_mem_zle, arith)
   apply (frule Bnor_fin)
--- a/src/HOL/Old_Number_Theory/IntFact.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Old_Number_Theory/IntFact.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -14,14 +14,14 @@
   \bigskip
 *}
 
-consts
+fun
   zfact :: "int => int"
-  d22set :: "int => int set"
-
-recdef zfact  "measure ((\<lambda>n. nat n) :: int => nat)"
+where
   "zfact n = (if n \<le> 0 then 1 else n * zfact (n - 1))"
 
-recdef d22set  "measure ((\<lambda>a. nat a) :: int => nat)"
+fun
+  d22set :: "int => int set"
+where
   "d22set a = (if 1 < a then insert a (d22set (a - 1)) else {})"
 
 
@@ -38,12 +38,10 @@
     and "!!a. 1 < (a::int) ==> P (d22set (a - 1)) (a - 1) ==> P (d22set a) a"
   shows "P (d22set u) u"
   apply (rule d22set.induct)
-  apply safe
-   prefer 2
-   apply (case_tac "1 < a")
-    apply (rule_tac prems)
-     apply (simp_all (no_asm_simp))
-   apply (simp_all (no_asm_simp) add: d22set.simps prems)
+  apply (case_tac "1 < a")
+   apply (rule_tac assms)
+    apply (simp_all (no_asm_simp))
+  apply (simp_all (no_asm_simp) add: d22set.simps assms)
   done
 
 lemma d22set_g_1 [rule_format]: "b \<in> d22set a --> 1 < b"
@@ -66,7 +64,8 @@
 lemma d22set_mem: "1 < b \<Longrightarrow> b \<le> a \<Longrightarrow> b \<in> d22set a"
   apply (induct a rule: d22set.induct)
   apply auto
-   apply (simp_all add: d22set.simps)
+  apply (subst d22set.simps)
+  apply (case_tac "b < a", auto)
   done
 
 lemma d22set_fin: "finite (d22set a)"
@@ -81,8 +80,6 @@
 
 lemma d22set_prod_zfact: "\<Prod>(d22set a) = zfact a"
   apply (induct a rule: d22set.induct)
-  apply safe
-   apply (simp add: d22set.simps zfact.simps)
   apply (subst d22set.simps)
   apply (subst zfact.simps)
   apply (case_tac "1 < a")
--- a/src/HOL/Old_Number_Theory/IntPrimes.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Old_Number_Theory/IntPrimes.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -19,17 +19,14 @@
 
 subsection {* Definitions *}
 
-consts
-  xzgcda :: "int * int * int * int * int * int * int * int => int * int * int"
-
-recdef xzgcda
-  "measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r)
-    :: int * int * int * int *int * int * int * int => nat)"
-  "xzgcda (m, n, r', r, s', s, t', t) =
+fun
+  xzgcda :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int => (int * int * int)"
+where
+  "xzgcda m n r' r s' s t' t =
         (if r \<le> 0 then (r', s', t')
-         else xzgcda (m, n, r, r' mod r, 
-                      s, s' - (r' div r) * s, 
-                      t, t' - (r' div r) * t))"
+         else xzgcda m n r (r' mod r) 
+                      s (s' - (r' div r) * s) 
+                      t (t' - (r' div r) * t))"
 
 definition
   zprime :: "int \<Rightarrow> bool" where
@@ -37,7 +34,7 @@
 
 definition
   xzgcd :: "int => int => int * int * int" where
-  "xzgcd m n = xzgcda (m, n, m, n, 1, 0, 0, 1)"
+  "xzgcd m n = xzgcda m n m n 1 0 0 1"
 
 definition
   zcong :: "int => int => int => bool"  ("(1[_ = _] '(mod _'))") where
@@ -307,9 +304,8 @@
 
 lemma xzgcd_correct_aux1:
   "zgcd r' r = k --> 0 < r -->
-    (\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn))"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
+    (\<exists>sn tn. xzgcda m n r' r s' s t' t = (k, sn, tn))"
+  apply (induct m n r' r s' s t' t rule: xzgcda.induct)
   apply (subst zgcd_eq)
   apply (subst xzgcda.simps, auto)
   apply (case_tac "r' mod r = 0")
@@ -321,17 +317,16 @@
   done
 
 lemma xzgcd_correct_aux2:
-  "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> 0 < r -->
+  "(\<exists>sn tn. xzgcda m n r' r s' s t' t = (k, sn, tn)) --> 0 < r -->
     zgcd r' r = k"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
+  apply (induct m n r' r s' s t' t rule: xzgcda.induct)
   apply (subst zgcd_eq)
   apply (subst xzgcda.simps)
   apply (auto simp add: linorder_not_le)
   apply (case_tac "r' mod r = 0")
    prefer 2
    apply (frule_tac a = "r'" in pos_mod_sign, auto)
-  apply (metis Pair_eq simps zle_refl)
+  apply (metis Pair_eq xzgcda.simps zle_refl)
   done
 
 lemma xzgcd_correct:
@@ -362,10 +357,9 @@
   by (rule iffD2 [OF order_less_le conjI])
 
 lemma xzgcda_linear [rule_format]:
-  "0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) -->
+  "0 < r --> xzgcda m n r' r s' s t' t = (rn, sn, tn) -->
     r' = s' * m + t' * n -->  r = s * m + t * n --> rn = sn * m + tn * n"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
+  apply (induct m n r' r s' s t' t rule: xzgcda.induct)
   apply (subst xzgcda.simps)
   apply (simp (no_asm))
   apply (rule impI)+
--- a/src/HOL/Old_Number_Theory/WilsonRuss.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -17,14 +17,12 @@
   inv :: "int => int => int" where
   "inv p a = (a^(nat (p - 2))) mod p"
 
-consts
-  wset :: "int * int => int set"
-
-recdef wset
-  "measure ((\<lambda>(a, p). nat a) :: int * int => nat)"
-  "wset (a, p) =
+fun
+  wset :: "int \<Rightarrow> int => int set"
+where
+  "wset a p =
     (if 1 < a then
-      let ws = wset (a - 1, p)
+      let ws = wset (a - 1) p
       in (if a \<in> ws then ws else insert a (insert (inv p a) ws)) else {})"
 
 
@@ -163,35 +161,33 @@
 lemma wset_induct:
   assumes "!!a p. P {} a p"
     and "!!a p. 1 < (a::int) \<Longrightarrow>
-      P (wset (a - 1, p)) (a - 1) p ==> P (wset (a, p)) a p"
-  shows "P (wset (u, v)) u v"
-  apply (rule wset.induct, safe)
-   prefer 2
-   apply (case_tac "1 < a")
-    apply (rule prems)
-     apply simp_all
-   apply (simp_all add: wset.simps prems)
+      P (wset (a - 1) p) (a - 1) p ==> P (wset a p) a p"
+  shows "P (wset u v) u v"
+  apply (rule wset.induct)
+  apply (case_tac "1 < a")
+   apply (rule assms)
+    apply (simp_all add: wset.simps assms)
   done
 
 lemma wset_mem_imp_or [rule_format]:
-  "1 < a \<Longrightarrow> b \<notin> wset (a - 1, p)
-    ==> b \<in> wset (a, p) --> b = a \<or> b = inv p a"
+  "1 < a \<Longrightarrow> b \<notin> wset (a - 1) p
+    ==> b \<in> wset a p --> b = a \<or> b = inv p a"
   apply (subst wset.simps)
   apply (unfold Let_def, simp)
   done
 
-lemma wset_mem_mem [simp]: "1 < a ==> a \<in> wset (a, p)"
+lemma wset_mem_mem [simp]: "1 < a ==> a \<in> wset a p"
   apply (subst wset.simps)
   apply (unfold Let_def, simp)
   done
 
-lemma wset_subset: "1 < a \<Longrightarrow> b \<in> wset (a - 1, p) ==> b \<in> wset (a, p)"
+lemma wset_subset: "1 < a \<Longrightarrow> b \<in> wset (a - 1) p ==> b \<in> wset a p"
   apply (subst wset.simps)
   apply (unfold Let_def, auto)
   done
 
 lemma wset_g_1 [rule_format]:
-    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> 1 < b"
+    "zprime p --> a < p - 1 --> b \<in> wset a p --> 1 < b"
   apply (induct a p rule: wset_induct, auto)
   apply (case_tac "b = a")
    apply (case_tac [2] "b = inv p a")
@@ -203,7 +199,7 @@
   done
 
 lemma wset_less [rule_format]:
-    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> b < p - 1"
+    "zprime p --> a < p - 1 --> b \<in> wset a p --> b < p - 1"
   apply (induct a p rule: wset_induct, auto)
   apply (case_tac "b = a")
    apply (case_tac [2] "b = inv p a")
@@ -216,7 +212,7 @@
 
 lemma wset_mem [rule_format]:
   "zprime p -->
-    a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset (a, p)"
+    a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset a p"
   apply (induct a p rule: wset.induct, auto)
   apply (rule_tac wset_subset)
   apply (simp (no_asm_simp))
@@ -224,8 +220,8 @@
   done
 
 lemma wset_mem_inv_mem [rule_format]:
-  "zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset (a, p)
-    --> inv p b \<in> wset (a, p)"
+  "zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset a p
+    --> inv p b \<in> wset a p"
   apply (induct a p rule: wset_induct, auto)
    apply (case_tac "b = a")
     apply (subst wset.simps)
@@ -240,13 +236,13 @@
 
 lemma wset_inv_mem_mem:
   "zprime p \<Longrightarrow> 5 \<le> p \<Longrightarrow> a < p - 1 \<Longrightarrow> 1 < b \<Longrightarrow> b < p - 1
-    \<Longrightarrow> inv p b \<in> wset (a, p) \<Longrightarrow> b \<in> wset (a, p)"
+    \<Longrightarrow> inv p b \<in> wset a p \<Longrightarrow> b \<in> wset a p"
   apply (rule_tac s = "inv p (inv p b)" and t = b in subst)
    apply (rule_tac [2] wset_mem_inv_mem)
       apply (rule inv_inv, simp_all)
   done
 
-lemma wset_fin: "finite (wset (a, p))"
+lemma wset_fin: "finite (wset a p)"
   apply (induct a p rule: wset_induct)
    prefer 2
    apply (subst wset.simps)
@@ -255,27 +251,27 @@
 
 lemma wset_zcong_prod_1 [rule_format]:
   "zprime p -->
-    5 \<le> p --> a < p - 1 --> [(\<Prod>x\<in>wset(a, p). x) = 1] (mod p)"
+    5 \<le> p --> a < p - 1 --> [(\<Prod>x\<in>wset a p. x) = 1] (mod p)"
   apply (induct a p rule: wset_induct)
    prefer 2
    apply (subst wset.simps)
-   apply (unfold Let_def, auto)
+   apply (auto, unfold Let_def, auto)
   apply (subst setprod_insert)
     apply (tactic {* stac (thm "setprod_insert") 3 *})
       apply (subgoal_tac [5]
-        "zcong (a * inv p a * (\<Prod>x\<in> wset(a - 1, p). x)) (1 * 1) p")
+        "zcong (a * inv p a * (\<Prod>x\<in>wset (a - 1) p. x)) (1 * 1) p")
        prefer 5
        apply (simp add: zmult_assoc)
       apply (rule_tac [5] zcong_zmult)
        apply (rule_tac [5] inv_is_inv)
          apply (tactic "clarify_tac @{claset} 4")
-         apply (subgoal_tac [4] "a \<in> wset (a - 1, p)")
+         apply (subgoal_tac [4] "a \<in> wset (a - 1) p")
           apply (rule_tac [5] wset_inv_mem_mem)
                apply (simp_all add: wset_fin)
   apply (rule inv_distinct, auto)
   done
 
-lemma d22set_eq_wset: "zprime p ==> d22set (p - 2) = wset (p - 2, p)"
+lemma d22set_eq_wset: "zprime p ==> d22set (p - 2) = wset (p - 2) p"
   apply safe
    apply (erule wset_mem)
      apply (rule_tac [2] d22set_g_1)
--- a/src/HOL/Orderings.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Orderings.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -657,13 +657,14 @@
 
   fun matches_bound v t =
     (case t of
-      Const ("_bound", _) $ Free (v', _) => v = v'
+      Const (@{syntax_const "_bound"}, _) $ Free (v', _) => v = v'
     | _ => false);
   fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false);
   fun mk v c n P = Syntax.const c $ Syntax.mark_bound v $ n $ P;
 
   fun tr' q = (q,
-    fn [Const ("_bound", _) $ Free (v, _), Const (c, _) $ (Const (d, _) $ t $ u) $ P] =>
+    fn [Const (@{syntax_const "_bound"}, _) $ Free (v, _),
+        Const (c, _) $ (Const (d, _) $ t $ u) $ P] =>
         (case AList.lookup (op =) trans (q, c, d) of
           NONE => raise Match
         | SOME (l, g) =>
--- a/src/HOL/PReal.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/PReal.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -9,7 +9,7 @@
 header {* Positive real numbers *}
 
 theory PReal
-imports Rational 
+imports Rat
 begin
 
 text{*Could be generalized and moved to @{text Groups}*}
--- a/src/HOL/Product_Type.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Product_Type.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -448,44 +448,44 @@
 *}
 
 ML {*
-
 local
-  val cond_split_eta_ss = HOL_basic_ss addsimps [thm "cond_split_eta"]
-  fun  Pair_pat k 0 (Bound m) = (m = k)
-  |    Pair_pat k i (Const ("Pair",  _) $ Bound m $ t) = i > 0 andalso
-                        m = k+i andalso Pair_pat k (i-1) t
-  |    Pair_pat _ _ _ = false;
-  fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t
-  |   no_args k i (t $ u) = no_args k i t andalso no_args k i u
-  |   no_args k i (Bound m) = m < k orelse m > k+i
-  |   no_args _ _ _ = true;
-  fun split_pat tp i (Abs  (_,_,t)) = if tp 0 i t then SOME (i,t) else NONE
-  |   split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t
-  |   split_pat tp i _ = NONE;
+  val cond_split_eta_ss = HOL_basic_ss addsimps @{thms cond_split_eta};
+  fun Pair_pat k 0 (Bound m) = (m = k)
+    | Pair_pat k i (Const (@{const_name Pair},  _) $ Bound m $ t) =
+        i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t
+    | Pair_pat _ _ _ = false;
+  fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t
+    | no_args k i (t $ u) = no_args k i t andalso no_args k i u
+    | no_args k i (Bound m) = m < k orelse m > k + i
+    | no_args _ _ _ = true;
+  fun split_pat tp i (Abs  (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE
+    | split_pat tp i (Const (@{const_name split}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
+    | split_pat tp i _ = NONE;
   fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] []
-        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs)))
+        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
         (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1)));
 
-  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t
-  |   beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse
-                        (beta_term_pat k i t andalso beta_term_pat k i u)
-  |   beta_term_pat k i t = no_args k i t;
-  fun  eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
-  |    eta_term_pat _ _ _ = false;
+  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t
+    | beta_term_pat k i (t $ u) =
+        Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u)
+    | beta_term_pat k i t = no_args k i t;
+  fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
+    | eta_term_pat _ _ _ = false;
   fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t)
-  |   subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg
-                              else (subst arg k i t $ subst arg k i u)
-  |   subst arg k i t = t;
-  fun beta_proc ss (s as Const ("split", _) $ Abs (_, _, t) $ arg) =
+    | subst arg k i (t $ u) =
+        if Pair_pat k i (t $ u) then incr_boundvars k arg
+        else (subst arg k i t $ subst arg k i u)
+    | subst arg k i t = t;
+  fun beta_proc ss (s as Const (@{const_name split}, _) $ Abs (_, _, t) $ arg) =
         (case split_pat beta_term_pat 1 t of
-        SOME (i,f) => SOME (metaeq ss s (subst arg 0 i f))
+          SOME (i, f) => SOME (metaeq ss s (subst arg 0 i f))
         | NONE => NONE)
-  |   beta_proc _ _ = NONE;
-  fun eta_proc ss (s as Const ("split", _) $ Abs (_, _, t)) =
+    | beta_proc _ _ = NONE;
+  fun eta_proc ss (s as Const (@{const_name split}, _) $ Abs (_, _, t)) =
         (case split_pat eta_term_pat 1 t of
-          SOME (_,ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end))
+          SOME (_, ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end))
         | NONE => NONE)
-  |   eta_proc _ _ = NONE;
+    | eta_proc _ _ = NONE;
 in
   val split_beta_proc = Simplifier.simproc @{theory} "split_beta" ["split f z"] (K beta_proc);
   val split_eta_proc = Simplifier.simproc @{theory} "split_eta" ["split f"] (K eta_proc);
@@ -565,11 +565,11 @@
 
 ML {*
 local (* filtering with exists_p_split is an essential optimization *)
-  fun exists_p_split (Const ("split",_) $ _ $ (Const ("Pair",_)$_$_)) = true
+  fun exists_p_split (Const (@{const_name split},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
     | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
     | exists_p_split (Abs (_, _, t)) = exists_p_split t
     | exists_p_split _ = false;
-  val ss = HOL_basic_ss addsimps [thm "split_conv"];
+  val ss = HOL_basic_ss addsimps @{thms split_conv};
 in
 val split_conv_tac = SUBGOAL (fn (t, i) =>
     if exists_p_split t then safe_full_simp_tac ss i else no_tac);
@@ -634,10 +634,11 @@
 lemma internal_split_conv: "internal_split c (a, b) = c a b"
   by (simp only: internal_split_def split_conv)
 
+use "Tools/split_rule.ML"
+setup Split_Rule.setup
+
 hide const internal_split
 
-use "Tools/split_rule.ML"
-setup SplitRule.setup
 
 lemmas prod_caseI = prod.cases [THEN iffD2, standard]
 
@@ -1049,7 +1050,6 @@
   "Pair"    ("(_,/ _)")
 
 setup {*
-
 let
 
 fun strip_abs_split 0 t = ([], t)
@@ -1058,16 +1058,18 @@
         val s' = Codegen.new_name t s;
         val v = Free (s', T)
       in apfst (cons v) (strip_abs_split (i-1) (subst_bound (v, t))) end
-  | strip_abs_split i (u as Const ("split", _) $ t) = (case strip_abs_split (i+1) t of
+  | strip_abs_split i (u as Const (@{const_name split}, _) $ t) =
+      (case strip_abs_split (i+1) t of
         (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u)
       | _ => ([], u))
   | strip_abs_split i t =
       strip_abs_split i (Abs ("x", hd (binder_types (fastype_of t)), t $ Bound 0));
 
-fun let_codegen thy defs dep thyname brack t gr = (case strip_comb t of
-    (t1 as Const ("Let", _), t2 :: t3 :: ts) =>
+fun let_codegen thy defs dep thyname brack t gr =
+  (case strip_comb t of
+    (t1 as Const (@{const_name Let}, _), t2 :: t3 :: ts) =>
     let
-      fun dest_let (l as Const ("Let", _) $ t $ u) =
+      fun dest_let (l as Const (@{const_name Let}, _) $ t $ u) =
           (case strip_abs_split 1 u of
              ([p], u') => apfst (cons (p, t)) (dest_let u')
            | _ => ([], l))
@@ -1098,7 +1100,7 @@
   | _ => NONE);
 
 fun split_codegen thy defs dep thyname brack t gr = (case strip_comb t of
-    (t1 as Const ("split", _), t2 :: ts) =>
+    (t1 as Const (@{const_name split}, _), t2 :: ts) =>
       let
         val ([p], u) = strip_abs_split 1 (t1 $ t2);
         val (q, gr1) = Codegen.invoke_codegen thy defs dep thyname false p gr;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Rat.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -0,0 +1,1210 @@
+(*  Title:  HOL/Rat.thy
+    Author: Markus Wenzel, TU Muenchen
+*)
+
+header {* Rational numbers *}
+
+theory Rat
+imports GCD Archimedean_Field
+uses ("Tools/float_syntax.ML")
+begin
+
+subsection {* Rational numbers as quotient *}
+
+subsubsection {* Construction of the type of rational numbers *}
+
+definition
+  ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" where
+  "ratrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
+
+lemma ratrel_iff [simp]:
+  "(x, y) \<in> ratrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
+  by (simp add: ratrel_def)
+
+lemma refl_on_ratrel: "refl_on {x. snd x \<noteq> 0} ratrel"
+  by (auto simp add: refl_on_def ratrel_def)
+
+lemma sym_ratrel: "sym ratrel"
+  by (simp add: ratrel_def sym_def)
+
+lemma trans_ratrel: "trans ratrel"
+proof (rule transI, unfold split_paired_all)
+  fix a b a' b' a'' b'' :: int
+  assume A: "((a, b), (a', b')) \<in> ratrel"
+  assume B: "((a', b'), (a'', b'')) \<in> ratrel"
+  have "b' * (a * b'') = b'' * (a * b')" by simp
+  also from A have "a * b' = a' * b" by auto
+  also have "b'' * (a' * b) = b * (a' * b'')" by simp
+  also from B have "a' * b'' = a'' * b'" by auto
+  also have "b * (a'' * b') = b' * (a'' * b)" by simp
+  finally have "b' * (a * b'') = b' * (a'' * b)" .
+  moreover from B have "b' \<noteq> 0" by auto
+  ultimately have "a * b'' = a'' * b" by simp
+  with A B show "((a, b), (a'', b'')) \<in> ratrel" by auto
+qed
+  
+lemma equiv_ratrel: "equiv {x. snd x \<noteq> 0} ratrel"
+  by (rule equiv.intro [OF refl_on_ratrel sym_ratrel trans_ratrel])
+
+lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
+lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]
+
+lemma equiv_ratrel_iff [iff]: 
+  assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
+  shows "ratrel `` {x} = ratrel `` {y} \<longleftrightarrow> (x, y) \<in> ratrel"
+  by (rule eq_equiv_class_iff, rule equiv_ratrel) (auto simp add: assms)
+
+typedef (Rat) rat = "{x. snd x \<noteq> 0} // ratrel"
+proof
+  have "(0::int, 1::int) \<in> {x. snd x \<noteq> 0}" by simp
+  then show "ratrel `` {(0, 1)} \<in> {x. snd x \<noteq> 0} // ratrel" by (rule quotientI)
+qed
+
+lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel `` {x} \<in> Rat"
+  by (simp add: Rat_def quotientI)
+
+declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp]
+
+
+subsubsection {* Representation and basic operations *}
+
+definition
+  Fract :: "int \<Rightarrow> int \<Rightarrow> rat" where
+  "Fract a b = Abs_Rat (ratrel `` {if b = 0 then (0, 1) else (a, b)})"
+
+lemma eq_rat:
+  shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
+  and "\<And>a. Fract a 0 = Fract 0 1"
+  and "\<And>a c. Fract 0 a = Fract 0 c"
+  by (simp_all add: Fract_def)
+
+lemma Rat_cases [case_names Fract, cases type: rat]:
+  assumes "\<And>a b. q = Fract a b \<Longrightarrow> b > 0 \<Longrightarrow> coprime a b \<Longrightarrow> C"
+  shows C
+proof -
+  obtain a b :: int where "q = Fract a b" and "b \<noteq> 0"
+    by (cases q) (clarsimp simp add: Fract_def Rat_def quotient_def)
+  let ?a = "a div gcd a b"
+  let ?b = "b div gcd a b"
+  from `b \<noteq> 0` have "?b * gcd a b = b"
+    by (simp add: dvd_div_mult_self)
+  with `b \<noteq> 0` have "?b \<noteq> 0" by auto
+  from `q = Fract a b` `b \<noteq> 0` `?b \<noteq> 0` have q: "q = Fract ?a ?b"
+    by (simp add: eq_rat dvd_div_mult mult_commute [of a])
+  from `b \<noteq> 0` have coprime: "coprime ?a ?b"
+    by (auto intro: div_gcd_coprime_int)
+  show C proof (cases "b > 0")
+    case True
+    note assms
+    moreover note q
+    moreover from True have "?b > 0" by (simp add: nonneg1_imp_zdiv_pos_iff)
+    moreover note coprime
+    ultimately show C .
+  next
+    case False
+    note assms
+    moreover from q have "q = Fract (- ?a) (- ?b)" by (simp add: Fract_def)
+    moreover from False `b \<noteq> 0` have "- ?b > 0" by (simp add: pos_imp_zdiv_neg_iff)
+    moreover from coprime have "coprime (- ?a) (- ?b)" by simp
+    ultimately show C .
+  qed
+qed
+
+lemma Rat_induct [case_names Fract, induct type: rat]:
+  assumes "\<And>a b. b > 0 \<Longrightarrow> coprime a b \<Longrightarrow> P (Fract a b)"
+  shows "P q"
+  using assms by (cases q) simp
+
+instantiation rat :: comm_ring_1
+begin
+
+definition
+  Zero_rat_def: "0 = Fract 0 1"
+
+definition
+  One_rat_def: "1 = Fract 1 1"
+
+definition
+  add_rat_def:
+  "q + r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
+    ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
+
+lemma add_rat [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
+proof -
+  have "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
+    respects2 ratrel"
+  by (rule equiv_ratrel [THEN congruent2_commuteI]) (simp_all add: left_distrib)
+  with assms show ?thesis by (simp add: Fract_def add_rat_def UN_ratrel2)
+qed
+
+definition
+  minus_rat_def:
+  "- q = Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel `` {(- fst x, snd x)})"
+
+lemma minus_rat [simp]: "- Fract a b = Fract (- a) b"
+proof -
+  have "(\<lambda>x. ratrel `` {(- fst x, snd x)}) respects ratrel"
+    by (simp add: congruent_def)
+  then show ?thesis by (simp add: Fract_def minus_rat_def UN_ratrel)
+qed
+
+lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b"
+  by (cases "b = 0") (simp_all add: eq_rat)
+
+definition
+  diff_rat_def: "q - r = q + - (r::rat)"
+
+lemma diff_rat [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
+  using assms by (simp add: diff_rat_def)
+
+definition
+  mult_rat_def:
+  "q * r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
+    ratrel``{(fst x * fst y, snd x * snd y)})"
+
+lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)"
+proof -
+  have "(\<lambda>x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel"
+    by (rule equiv_ratrel [THEN congruent2_commuteI]) simp_all
+  then show ?thesis by (simp add: Fract_def mult_rat_def UN_ratrel2)
+qed
+
+lemma mult_rat_cancel:
+  assumes "c \<noteq> 0"
+  shows "Fract (c * a) (c * b) = Fract a b"
+proof -
+  from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def)
+  then show ?thesis by (simp add: mult_rat [symmetric])
+qed
+
+instance proof
+  fix q r s :: rat show "(q * r) * s = q * (r * s)" 
+    by (cases q, cases r, cases s) (simp add: eq_rat)
+next
+  fix q r :: rat show "q * r = r * q"
+    by (cases q, cases r) (simp add: eq_rat)
+next
+  fix q :: rat show "1 * q = q"
+    by (cases q) (simp add: One_rat_def eq_rat)
+next
+  fix q r s :: rat show "(q + r) + s = q + (r + s)"
+    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
+next
+  fix q r :: rat show "q + r = r + q"
+    by (cases q, cases r) (simp add: eq_rat)
+next
+  fix q :: rat show "0 + q = q"
+    by (cases q) (simp add: Zero_rat_def eq_rat)
+next
+  fix q :: rat show "- q + q = 0"
+    by (cases q) (simp add: Zero_rat_def eq_rat)
+next
+  fix q r :: rat show "q - r = q + - r"
+    by (cases q, cases r) (simp add: eq_rat)
+next
+  fix q r s :: rat show "(q + r) * s = q * s + r * s"
+    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
+next
+  show "(0::rat) \<noteq> 1" by (simp add: Zero_rat_def One_rat_def eq_rat)
+qed
+
+end
+
+lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1"
+  by (induct k) (simp_all add: Zero_rat_def One_rat_def)
+
+lemma of_int_rat: "of_int k = Fract k 1"
+  by (cases k rule: int_diff_cases) (simp add: of_nat_rat)
+
+lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
+  by (rule of_nat_rat [symmetric])
+
+lemma Fract_of_int_eq: "Fract k 1 = of_int k"
+  by (rule of_int_rat [symmetric])
+
+instantiation rat :: number_ring
+begin
+
+definition
+  rat_number_of_def: "number_of w = Fract w 1"
+
+instance proof
+qed (simp add: rat_number_of_def of_int_rat)
+
+end
+
+lemma rat_number_collapse:
+  "Fract 0 k = 0"
+  "Fract 1 1 = 1"
+  "Fract (number_of k) 1 = number_of k"
+  "Fract k 0 = 0"
+  by (cases "k = 0")
+    (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def)
+
+lemma rat_number_expand [code_unfold]:
+  "0 = Fract 0 1"
+  "1 = Fract 1 1"
+  "number_of k = Fract (number_of k) 1"
+  by (simp_all add: rat_number_collapse)
+
+lemma iszero_rat [simp]:
+  "iszero (number_of k :: rat) \<longleftrightarrow> iszero (number_of k :: int)"
+  by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat)
+
+lemma Rat_cases_nonzero [case_names Fract 0]:
+  assumes Fract: "\<And>a b. q = Fract a b \<Longrightarrow> b > 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> C"
+  assumes 0: "q = 0 \<Longrightarrow> C"
+  shows C
+proof (cases "q = 0")
+  case True then show C using 0 by auto
+next
+  case False
+  then obtain a b where "q = Fract a b" and "b > 0" and "coprime a b" by (cases q) auto
+  moreover with False have "0 \<noteq> Fract a b" by simp
+  with `b > 0` have "a \<noteq> 0" by (simp add: Zero_rat_def eq_rat)
+  with Fract `q = Fract a b` `b > 0` `coprime a b` show C by blast
+qed
+
+subsubsection {* Function @{text normalize} *}
+
+lemma Fract_coprime: "Fract (a div gcd a b) (b div gcd a b) = Fract a b"
+proof (cases "b = 0")
+  case True then show ?thesis by (simp add: eq_rat)
+next
+  case False
+  moreover have "b div gcd a b * gcd a b = b"
+    by (rule dvd_div_mult_self) simp
+  ultimately have "b div gcd a b \<noteq> 0" by auto
+  with False show ?thesis by (simp add: eq_rat dvd_div_mult mult_commute [of a])
+qed
+
+definition normalize :: "int \<times> int \<Rightarrow> int \<times> int" where
+  "normalize p = (if snd p > 0 then (let a = gcd (fst p) (snd p) in (fst p div a, snd p div a))
+    else if snd p = 0 then (0, 1)
+    else (let a = - gcd (fst p) (snd p) in (fst p div a, snd p div a)))"
+
+lemma normalize_crossproduct:
+  assumes "q \<noteq> 0" "s \<noteq> 0"
+  assumes "normalize (p, q) = normalize (r, s)"
+  shows "p * s = r * q"
+proof -
+  have aux: "p * gcd r s = sgn (q * s) * r * gcd p q \<Longrightarrow> q * gcd r s = sgn (q * s) * s * gcd p q \<Longrightarrow> p * s = q * r"
+  proof -
+    assume "p * gcd r s = sgn (q * s) * r * gcd p q" and "q * gcd r s = sgn (q * s) * s * gcd p q"
+    then have "(p * gcd r s) * (sgn (q * s) * s * gcd p q) = (q * gcd r s) * (sgn (q * s) * r * gcd p q)" by simp
+    with assms show "p * s = q * r" by (auto simp add: mult_ac sgn_times sgn_0_0)
+  qed
+  from assms show ?thesis
+    by (auto simp add: normalize_def Let_def dvd_div_div_eq_mult mult_commute sgn_times split: if_splits intro: aux)
+qed
+
+lemma normalize_eq: "normalize (a, b) = (p, q) \<Longrightarrow> Fract p q = Fract a b"
+  by (auto simp add: normalize_def Let_def Fract_coprime dvd_div_neg rat_number_collapse
+    split:split_if_asm)
+
+lemma normalize_denom_pos: "normalize r = (p, q) \<Longrightarrow> q > 0"
+  by (auto simp add: normalize_def Let_def dvd_div_neg pos_imp_zdiv_neg_iff nonneg1_imp_zdiv_pos_iff
+    split:split_if_asm)
+
+lemma normalize_coprime: "normalize r = (p, q) \<Longrightarrow> coprime p q"
+  by (auto simp add: normalize_def Let_def dvd_div_neg div_gcd_coprime_int
+    split:split_if_asm)
+
+lemma normalize_stable [simp]:
+  "q > 0 \<Longrightarrow> coprime p q \<Longrightarrow> normalize (p, q) = (p, q)"
+  by (simp add: normalize_def)
+
+lemma normalize_denom_zero [simp]:
+  "normalize (p, 0) = (0, 1)"
+  by (simp add: normalize_def)
+
+lemma normalize_negative [simp]:
+  "q < 0 \<Longrightarrow> normalize (p, q) = normalize (- p, - q)"
+  by (simp add: normalize_def Let_def dvd_div_neg dvd_neg_div)
+
+text{*
+  Decompose a fraction into normalized, i.e. coprime numerator and denominator:
+*}
+
+definition quotient_of :: "rat \<Rightarrow> int \<times> int" where
+  "quotient_of x = (THE pair. x = Fract (fst pair) (snd pair) &
+                   snd pair > 0 & coprime (fst pair) (snd pair))"
+
+lemma quotient_of_unique:
+  "\<exists>!p. r = Fract (fst p) (snd p) \<and> snd p > 0 \<and> coprime (fst p) (snd p)"
+proof (cases r)
+  case (Fract a b)
+  then have "r = Fract (fst (a, b)) (snd (a, b)) \<and> snd (a, b) > 0 \<and> coprime (fst (a, b)) (snd (a, b))" by auto
+  then show ?thesis proof (rule ex1I)
+    fix p
+    obtain c d :: int where p: "p = (c, d)" by (cases p)
+    assume "r = Fract (fst p) (snd p) \<and> snd p > 0 \<and> coprime (fst p) (snd p)"
+    with p have Fract': "r = Fract c d" "d > 0" "coprime c d" by simp_all
+    have "c = a \<and> d = b"
+    proof (cases "a = 0")
+      case True with Fract Fract' show ?thesis by (simp add: eq_rat)
+    next
+      case False
+      with Fract Fract' have *: "c * b = a * d" and "c \<noteq> 0" by (auto simp add: eq_rat)
+      then have "c * b > 0 \<longleftrightarrow> a * d > 0" by auto
+      with `b > 0` `d > 0` have "a > 0 \<longleftrightarrow> c > 0" by (simp add: zero_less_mult_iff)
+      with `a \<noteq> 0` `c \<noteq> 0` have sgn: "sgn a = sgn c" by (auto simp add: not_less)
+      from `coprime a b` `coprime c d` have "\<bar>a\<bar> * \<bar>d\<bar> = \<bar>c\<bar> * \<bar>b\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> \<bar>d\<bar> = \<bar>b\<bar>"
+        by (simp add: coprime_crossproduct_int)
+      with `b > 0` `d > 0` have "\<bar>a\<bar> * d = \<bar>c\<bar> * b \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> d = b" by simp
+      then have "a * sgn a * d = c * sgn c * b \<longleftrightarrow> a * sgn a = c * sgn c \<and> d = b" by (simp add: abs_sgn)
+      with sgn * show ?thesis by (auto simp add: sgn_0_0)
+    qed
+    with p show "p = (a, b)" by simp
+  qed
+qed
+
+lemma quotient_of_Fract [code]:
+  "quotient_of (Fract a b) = normalize (a, b)"
+proof -
+  have "Fract a b = Fract (fst (normalize (a, b))) (snd (normalize (a, b)))" (is ?Fract)
+    by (rule sym) (auto intro: normalize_eq)
+  moreover have "0 < snd (normalize (a, b))" (is ?denom_pos) 
+    by (cases "normalize (a, b)") (rule normalize_denom_pos, simp)
+  moreover have "coprime (fst (normalize (a, b))) (snd (normalize (a, b)))" (is ?coprime)
+    by (rule normalize_coprime) simp
+  ultimately have "?Fract \<and> ?denom_pos \<and> ?coprime" by blast
+  with quotient_of_unique have
+    "(THE p. Fract a b = Fract (fst p) (snd p) \<and> 0 < snd p \<and> coprime (fst p) (snd p)) = normalize (a, b)"
+    by (rule the1_equality)
+  then show ?thesis by (simp add: quotient_of_def)
+qed
+
+lemma quotient_of_number [simp]:
+  "quotient_of 0 = (0, 1)"
+  "quotient_of 1 = (1, 1)"
+  "quotient_of (number_of k) = (number_of k, 1)"
+  by (simp_all add: rat_number_expand quotient_of_Fract)
+
+lemma quotient_of_eq: "quotient_of (Fract a b) = (p, q) \<Longrightarrow> Fract p q = Fract a b"
+  by (simp add: quotient_of_Fract normalize_eq)
+
+lemma quotient_of_denom_pos: "quotient_of r = (p, q) \<Longrightarrow> q > 0"
+  by (cases r) (simp add: quotient_of_Fract normalize_denom_pos)
+
+lemma quotient_of_coprime: "quotient_of r = (p, q) \<Longrightarrow> coprime p q"
+  by (cases r) (simp add: quotient_of_Fract normalize_coprime)
+
+lemma quotient_of_inject:
+  assumes "quotient_of a = quotient_of b"
+  shows "a = b"
+proof -
+  obtain p q r s where a: "a = Fract p q"
+    and b: "b = Fract r s"
+    and "q > 0" and "s > 0" by (cases a, cases b)
+  with assms show ?thesis by (simp add: eq_rat quotient_of_Fract normalize_crossproduct)
+qed
+
+lemma quotient_of_inject_eq:
+  "quotient_of a = quotient_of b \<longleftrightarrow> a = b"
+  by (auto simp add: quotient_of_inject)
+
+
+subsubsection {* The field of rational numbers *}
+
+instantiation rat :: "{field, division_by_zero}"
+begin
+
+definition
+  inverse_rat_def:
+  "inverse q = Abs_Rat (\<Union>x \<in> Rep_Rat q.
+     ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
+
+lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a"
+proof -
+  have "(\<lambda>x. ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)}) respects ratrel"
+    by (auto simp add: congruent_def mult_commute)
+  then show ?thesis by (simp add: Fract_def inverse_rat_def UN_ratrel)
+qed
+
+definition
+  divide_rat_def: "q / r = q * inverse (r::rat)"
+
+lemma divide_rat [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
+  by (simp add: divide_rat_def)
+
+instance proof
+  show "inverse 0 = (0::rat)" by (simp add: rat_number_expand)
+    (simp add: rat_number_collapse)
+next
+  fix q :: rat
+  assume "q \<noteq> 0"
+  then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero)
+   (simp_all add: rat_number_expand eq_rat)
+next
+  fix q r :: rat
+  show "q / r = q * inverse r" by (simp add: divide_rat_def)
+qed
+
+end
+
+
+subsubsection {* Various *}
+
+lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1"
+  by (simp add: rat_number_expand)
+
+lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l"
+  by (simp add: Fract_of_int_eq [symmetric])
+
+lemma Fract_number_of_quotient:
+  "Fract (number_of k) (number_of l) = number_of k / number_of l"
+  unfolding Fract_of_int_quotient number_of_is_id number_of_eq ..
+
+lemma Fract_1_number_of:
+  "Fract 1 (number_of k) = 1 / number_of k"
+  unfolding Fract_of_int_quotient number_of_eq by simp
+
+subsubsection {* The ordered field of rational numbers *}
+
+instantiation rat :: linorder
+begin
+
+definition
+  le_rat_def:
+   "q \<le> r \<longleftrightarrow> contents (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
+      {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})"
+
+lemma le_rat [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
+proof -
+  have "(\<lambda>x y. {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})
+    respects2 ratrel"
+  proof (clarsimp simp add: congruent2_def)
+    fix a b a' b' c d c' d'::int
+    assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+    assume eq1: "a * b' = a' * b"
+    assume eq2: "c * d' = c' * d"
+
+    let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+    {
+      fix a b c d x :: int assume x: "x \<noteq> 0"
+      have "?le a b c d = ?le (a * x) (b * x) c d"
+      proof -
+        from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
+        hence "?le a b c d =
+            ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
+          by (simp add: mult_le_cancel_right)
+        also have "... = ?le (a * x) (b * x) c d"
+          by (simp add: mult_ac)
+        finally show ?thesis .
+      qed
+    } note le_factor = this
+
+    let ?D = "b * d" and ?D' = "b' * d'"
+    from neq have D: "?D \<noteq> 0" by simp
+    from neq have "?D' \<noteq> 0" by simp
+    hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
+      by (rule le_factor)
+    also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" 
+      by (simp add: mult_ac)
+    also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
+      by (simp only: eq1 eq2)
+    also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
+      by (simp add: mult_ac)
+    also from D have "... = ?le a' b' c' d'"
+      by (rule le_factor [symmetric])
+    finally show "?le a b c d = ?le a' b' c' d'" .
+  qed
+  with assms show ?thesis by (simp add: Fract_def le_rat_def UN_ratrel2)
+qed
+
+definition
+  less_rat_def: "z < (w::rat) \<longleftrightarrow> z \<le> w \<and> z \<noteq> w"
+
+lemma less_rat [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
+  using assms by (simp add: less_rat_def eq_rat order_less_le)
+
+instance proof
+  fix q r s :: rat
+  {
+    assume "q \<le> r" and "r \<le> s"
+    then show "q \<le> s" 
+    proof (induct q, induct r, induct s)
+      fix a b c d e f :: int
+      assume neq: "b > 0"  "d > 0"  "f > 0"
+      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
+      show "Fract a b \<le> Fract e f"
+      proof -
+        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
+          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
+        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
+        proof -
+          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+            by simp
+          with ff show ?thesis by (simp add: mult_le_cancel_right)
+        qed
+        also have "... = (c * f) * (d * f) * (b * b)" by algebra
+        also have "... \<le> (e * d) * (d * f) * (b * b)"
+        proof -
+          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
+            by simp
+          with bb show ?thesis by (simp add: mult_le_cancel_right)
+        qed
+        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
+          by (simp only: mult_ac)
+        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
+          by (simp add: mult_le_cancel_right)
+        with neq show ?thesis by simp
+      qed
+    qed
+  next
+    assume "q \<le> r" and "r \<le> q"
+    then show "q = r"
+    proof (induct q, induct r)
+      fix a b c d :: int
+      assume neq: "b > 0"  "d > 0"
+      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
+      show "Fract a b = Fract c d"
+      proof -
+        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+          by simp
+        also have "... \<le> (a * d) * (b * d)"
+        proof -
+          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
+            by simp
+          thus ?thesis by (simp only: mult_ac)
+        qed
+        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
+        moreover from neq have "b * d \<noteq> 0" by simp
+        ultimately have "a * d = c * b" by simp
+        with neq show ?thesis by (simp add: eq_rat)
+      qed
+    qed
+  next
+    show "q \<le> q"
+      by (induct q) simp
+    show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
+      by (induct q, induct r) (auto simp add: le_less mult_commute)
+    show "q \<le> r \<or> r \<le> q"
+      by (induct q, induct r)
+         (simp add: mult_commute, rule linorder_linear)
+  }
+qed
+
+end
+
+instantiation rat :: "{distrib_lattice, abs_if, sgn_if}"
+begin
+
+definition
+  abs_rat_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::rat))"
+
+lemma abs_rat [simp, code]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
+  by (auto simp add: abs_rat_def zabs_def Zero_rat_def not_less le_less eq_rat zero_less_mult_iff)
+
+definition
+  sgn_rat_def: "sgn (q::rat) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
+
+lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)"
+  unfolding Fract_of_int_eq
+  by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat)
+    (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff)
+
+definition
+  "(inf \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = min"
+
+definition
+  "(sup \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = max"
+
+instance by intro_classes
+  (auto simp add: abs_rat_def sgn_rat_def min_max.sup_inf_distrib1 inf_rat_def sup_rat_def)
+
+end
+
+instance rat :: linordered_field
+proof
+  fix q r s :: rat
+  show "q \<le> r ==> s + q \<le> s + r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: int
+    assume neq: "b > 0"  "d > 0"  "f > 0"
+    assume le: "Fract a b \<le> Fract c d"
+    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
+    proof -
+      let ?F = "f * f" from neq have F: "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by simp
+      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by (simp add: mult_ac int_distrib)
+    qed
+  qed
+  show "q < r ==> 0 < s ==> s * q < s * r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: int
+    assume neq: "b > 0"  "d > 0"  "f > 0"
+    assume le: "Fract a b < Fract c d"
+    assume gt: "0 < Fract e f"
+    show "Fract e f * Fract a b < Fract e f * Fract c d"
+    proof -
+      let ?E = "e * f" and ?F = "f * f"
+      from neq gt have "0 < ?E"
+        by (auto simp add: Zero_rat_def order_less_le eq_rat)
+      moreover from neq have "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
+        by simp
+      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
+        by (simp add: mult_less_cancel_right)
+      with neq show ?thesis
+        by (simp add: mult_ac)
+    qed
+  qed
+qed auto
+
+lemma Rat_induct_pos [case_names Fract, induct type: rat]:
+  assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
+  shows "P q"
+proof (cases q)
+  have step': "\<And>a b. b < 0 \<Longrightarrow> P (Fract a b)"
+  proof -
+    fix a::int and b::int
+    assume b: "b < 0"
+    hence "0 < -b" by simp
+    hence "P (Fract (-a) (-b))" by (rule step)
+    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
+  qed
+  case (Fract a b)
+  thus "P q" by (force simp add: linorder_neq_iff step step')
+qed
+
+lemma zero_less_Fract_iff:
+  "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
+  by (simp add: Zero_rat_def zero_less_mult_iff)
+
+lemma Fract_less_zero_iff:
+  "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
+  by (simp add: Zero_rat_def mult_less_0_iff)
+
+lemma zero_le_Fract_iff:
+  "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
+  by (simp add: Zero_rat_def zero_le_mult_iff)
+
+lemma Fract_le_zero_iff:
+  "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
+  by (simp add: Zero_rat_def mult_le_0_iff)
+
+lemma one_less_Fract_iff:
+  "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
+  by (simp add: One_rat_def mult_less_cancel_right_disj)
+
+lemma Fract_less_one_iff:
+  "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
+  by (simp add: One_rat_def mult_less_cancel_right_disj)
+
+lemma one_le_Fract_iff:
+  "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
+  by (simp add: One_rat_def mult_le_cancel_right)
+
+lemma Fract_le_one_iff:
+  "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
+  by (simp add: One_rat_def mult_le_cancel_right)
+
+
+subsubsection {* Rationals are an Archimedean field *}
+
+lemma rat_floor_lemma:
+  shows "of_int (a div b) \<le> Fract a b \<and> Fract a b < of_int (a div b + 1)"
+proof -
+  have "Fract a b = of_int (a div b) + Fract (a mod b) b"
+    by (cases "b = 0", simp, simp add: of_int_rat)
+  moreover have "0 \<le> Fract (a mod b) b \<and> Fract (a mod b) b < 1"
+    unfolding Fract_of_int_quotient
+    by (rule linorder_cases [of b 0])
+       (simp add: divide_nonpos_neg, simp, simp add: divide_nonneg_pos)
+  ultimately show ?thesis by simp
+qed
+
+instance rat :: archimedean_field
+proof
+  fix r :: rat
+  show "\<exists>z. r \<le> of_int z"
+  proof (induct r)
+    case (Fract a b)
+    have "Fract a b \<le> of_int (a div b + 1)"
+      using rat_floor_lemma [of a b] by simp
+    then show "\<exists>z. Fract a b \<le> of_int z" ..
+  qed
+qed
+
+lemma floor_Fract: "floor (Fract a b) = a div b"
+  using rat_floor_lemma [of a b]
+  by (simp add: floor_unique)
+
+
+subsection {* Linear arithmetic setup *}
+
+declaration {*
+  K (Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2]
+    (* not needed because x < (y::nat) can be rewritten as Suc x <= y: of_nat_less_iff RS iffD2 *)
+  #> Lin_Arith.add_inj_thms [@{thm of_int_le_iff} RS iffD2, @{thm of_int_eq_iff} RS iffD2]
+    (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *)
+  #> Lin_Arith.add_simps [@{thm neg_less_iff_less},
+      @{thm True_implies_equals},
+      read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib},
+      @{thm divide_1}, @{thm divide_zero_left},
+      @{thm times_divide_eq_right}, @{thm times_divide_eq_left},
+      @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym,
+      @{thm of_int_minus}, @{thm of_int_diff},
+      @{thm of_int_of_nat_eq}]
+  #> Lin_Arith.add_simprocs Numeral_Simprocs.field_cancel_numeral_factors
+  #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat => rat"})
+  #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int => rat"}))
+*}
+
+
+subsection {* Embedding from Rationals to other Fields *}
+
+class field_char_0 = field + ring_char_0
+
+subclass (in linordered_field) field_char_0 ..
+
+context field_char_0
+begin
+
+definition of_rat :: "rat \<Rightarrow> 'a" where
+  "of_rat q = contents (\<Union>(a,b) \<in> Rep_Rat q. {of_int a / of_int b})"
+
+end
+
+lemma of_rat_congruent:
+  "(\<lambda>(a, b). {of_int a / of_int b :: 'a::field_char_0}) respects ratrel"
+apply (rule congruent.intro)
+apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
+apply (simp only: of_int_mult [symmetric])
+done
+
+lemma of_rat_rat: "b \<noteq> 0 \<Longrightarrow> of_rat (Fract a b) = of_int a / of_int b"
+  unfolding Fract_def of_rat_def by (simp add: UN_ratrel of_rat_congruent)
+
+lemma of_rat_0 [simp]: "of_rat 0 = 0"
+by (simp add: Zero_rat_def of_rat_rat)
+
+lemma of_rat_1 [simp]: "of_rat 1 = 1"
+by (simp add: One_rat_def of_rat_rat)
+
+lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b"
+by (induct a, induct b, simp add: of_rat_rat add_frac_eq)
+
+lemma of_rat_minus: "of_rat (- a) = - of_rat a"
+by (induct a, simp add: of_rat_rat)
+
+lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
+by (simp only: diff_minus of_rat_add of_rat_minus)
+
+lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
+apply (induct a, induct b, simp add: of_rat_rat)
+apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac)
+done
+
+lemma nonzero_of_rat_inverse:
+  "a \<noteq> 0 \<Longrightarrow> of_rat (inverse a) = inverse (of_rat a)"
+apply (rule inverse_unique [symmetric])
+apply (simp add: of_rat_mult [symmetric])
+done
+
+lemma of_rat_inverse:
+  "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
+   inverse (of_rat a)"
+by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
+
+lemma nonzero_of_rat_divide:
+  "b \<noteq> 0 \<Longrightarrow> of_rat (a / b) = of_rat a / of_rat b"
+by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
+
+lemma of_rat_divide:
+  "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
+   = of_rat a / of_rat b"
+by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
+
+lemma of_rat_power:
+  "(of_rat (a ^ n)::'a::field_char_0) = of_rat a ^ n"
+by (induct n) (simp_all add: of_rat_mult)
+
+lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)"
+apply (induct a, induct b)
+apply (simp add: of_rat_rat eq_rat)
+apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
+apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
+done
+
+lemma of_rat_less:
+  "(of_rat r :: 'a::linordered_field) < of_rat s \<longleftrightarrow> r < s"
+proof (induct r, induct s)
+  fix a b c d :: int
+  assume not_zero: "b > 0" "d > 0"
+  then have "b * d > 0" by (rule mult_pos_pos)
+  have of_int_divide_less_eq:
+    "(of_int a :: 'a) / of_int b < of_int c / of_int d
+      \<longleftrightarrow> (of_int a :: 'a) * of_int d < of_int c * of_int b"
+    using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq)
+  show "(of_rat (Fract a b) :: 'a::linordered_field) < of_rat (Fract c d)
+    \<longleftrightarrow> Fract a b < Fract c d"
+    using not_zero `b * d > 0`
+    by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult)
+qed
+
+lemma of_rat_less_eq:
+  "(of_rat r :: 'a::linordered_field) \<le> of_rat s \<longleftrightarrow> r \<le> s"
+  unfolding le_less by (auto simp add: of_rat_less)
+
+lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified]
+
+lemma of_rat_eq_id [simp]: "of_rat = id"
+proof
+  fix a
+  show "of_rat a = id a"
+  by (induct a)
+     (simp add: of_rat_rat Fract_of_int_eq [symmetric])
+qed
+
+text{*Collapse nested embeddings*}
+lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
+by (induct n) (simp_all add: of_rat_add)
+
+lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z"
+by (cases z rule: int_diff_cases) (simp add: of_rat_diff)
+
+lemma of_rat_number_of_eq [simp]:
+  "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})"
+by (simp add: number_of_eq)
+
+lemmas zero_rat = Zero_rat_def
+lemmas one_rat = One_rat_def
+
+abbreviation
+  rat_of_nat :: "nat \<Rightarrow> rat"
+where
+  "rat_of_nat \<equiv> of_nat"
+
+abbreviation
+  rat_of_int :: "int \<Rightarrow> rat"
+where
+  "rat_of_int \<equiv> of_int"
+
+subsection {* The Set of Rational Numbers *}
+
+context field_char_0
+begin
+
+definition
+  Rats  :: "'a set" where
+  "Rats = range of_rat"
+
+notation (xsymbols)
+  Rats  ("\<rat>")
+
+end
+
+lemma Rats_of_rat [simp]: "of_rat r \<in> Rats"
+by (simp add: Rats_def)
+
+lemma Rats_of_int [simp]: "of_int z \<in> Rats"
+by (subst of_rat_of_int_eq [symmetric], rule Rats_of_rat)
+
+lemma Rats_of_nat [simp]: "of_nat n \<in> Rats"
+by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat)
+
+lemma Rats_number_of [simp]:
+  "(number_of w::'a::{number_ring,field_char_0}) \<in> Rats"
+by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat)
+
+lemma Rats_0 [simp]: "0 \<in> Rats"
+apply (unfold Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_0 [symmetric])
+done
+
+lemma Rats_1 [simp]: "1 \<in> Rats"
+apply (unfold Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_1 [symmetric])
+done
+
+lemma Rats_add [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a + b \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_add [symmetric])
+done
+
+lemma Rats_minus [simp]: "a \<in> Rats \<Longrightarrow> - a \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_minus [symmetric])
+done
+
+lemma Rats_diff [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a - b \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_diff [symmetric])
+done
+
+lemma Rats_mult [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a * b \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_mult [symmetric])
+done
+
+lemma nonzero_Rats_inverse:
+  fixes a :: "'a::field_char_0"
+  shows "\<lbrakk>a \<in> Rats; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_rat_inverse [symmetric])
+done
+
+lemma Rats_inverse [simp]:
+  fixes a :: "'a::{field_char_0,division_by_zero}"
+  shows "a \<in> Rats \<Longrightarrow> inverse a \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_inverse [symmetric])
+done
+
+lemma nonzero_Rats_divide:
+  fixes a b :: "'a::field_char_0"
+  shows "\<lbrakk>a \<in> Rats; b \<in> Rats; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (erule nonzero_of_rat_divide [symmetric])
+done
+
+lemma Rats_divide [simp]:
+  fixes a b :: "'a::{field_char_0,division_by_zero}"
+  shows "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_divide [symmetric])
+done
+
+lemma Rats_power [simp]:
+  fixes a :: "'a::field_char_0"
+  shows "a \<in> Rats \<Longrightarrow> a ^ n \<in> Rats"
+apply (auto simp add: Rats_def)
+apply (rule range_eqI)
+apply (rule of_rat_power [symmetric])
+done
+
+lemma Rats_cases [cases set: Rats]:
+  assumes "q \<in> \<rat>"
+  obtains (of_rat) r where "q = of_rat r"
+  unfolding Rats_def
+proof -
+  from `q \<in> \<rat>` have "q \<in> range of_rat" unfolding Rats_def .
+  then obtain r where "q = of_rat r" ..
+  then show thesis ..
+qed
+
+lemma Rats_induct [case_names of_rat, induct set: Rats]:
+  "q \<in> \<rat> \<Longrightarrow> (\<And>r. P (of_rat r)) \<Longrightarrow> P q"
+  by (rule Rats_cases) auto
+
+
+subsection {* Implementation of rational numbers as pairs of integers *}
+
+definition Frct :: "int \<times> int \<Rightarrow> rat" where
+  [simp]: "Frct p = Fract (fst p) (snd p)"
+
+code_abstype Frct quotient_of
+proof (rule eq_reflection)
+  fix r :: rat
+  show "Frct (quotient_of r) = r" by (cases r) (auto intro: quotient_of_eq)
+qed
+
+lemma Frct_code_post [code_post]:
+  "Frct (0, k) = 0"
+  "Frct (k, 0) = 0"
+  "Frct (1, 1) = 1"
+  "Frct (number_of k, 1) = number_of k"
+  "Frct (1, number_of k) = 1 / number_of k"
+  "Frct (number_of k, number_of l) = number_of k / number_of l"
+  by (simp_all add: rat_number_collapse Fract_number_of_quotient Fract_1_number_of)
+
+declare quotient_of_Fract [code abstract]
+
+lemma rat_zero_code [code abstract]:
+  "quotient_of 0 = (0, 1)"
+  by (simp add: Zero_rat_def quotient_of_Fract normalize_def)
+
+lemma rat_one_code [code abstract]:
+  "quotient_of 1 = (1, 1)"
+  by (simp add: One_rat_def quotient_of_Fract normalize_def)
+
+lemma rat_plus_code [code abstract]:
+  "quotient_of (p + q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q
+     in normalize (a * d + b * c, c * d))"
+  by (cases p, cases q) (simp add: quotient_of_Fract)
+
+lemma rat_uminus_code [code abstract]:
+  "quotient_of (- p) = (let (a, b) = quotient_of p in (- a, b))"
+  by (cases p) (simp add: quotient_of_Fract)
+
+lemma rat_minus_code [code abstract]:
+  "quotient_of (p - q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q
+     in normalize (a * d - b * c, c * d))"
+  by (cases p, cases q) (simp add: quotient_of_Fract)
+
+lemma rat_times_code [code abstract]:
+  "quotient_of (p * q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q
+     in normalize (a * b, c * d))"
+  by (cases p, cases q) (simp add: quotient_of_Fract)
+
+lemma rat_inverse_code [code abstract]:
+  "quotient_of (inverse p) = (let (a, b) = quotient_of p
+    in if a = 0 then (0, 1) else (sgn a * b, \<bar>a\<bar>))"
+proof (cases p)
+  case (Fract a b) then show ?thesis
+    by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract gcd_int.commute)
+qed
+
+lemma rat_divide_code [code abstract]:
+  "quotient_of (p / q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q
+     in normalize (a * d, c * b))"
+  by (cases p, cases q) (simp add: quotient_of_Fract)
+
+lemma rat_abs_code [code abstract]:
+  "quotient_of \<bar>p\<bar> = (let (a, b) = quotient_of p in (\<bar>a\<bar>, b))"
+  by (cases p) (simp add: quotient_of_Fract)
+
+lemma rat_sgn_code [code abstract]:
+  "quotient_of (sgn p) = (sgn (fst (quotient_of p)), 1)"
+proof (cases p)
+  case (Fract a b) then show ?thesis
+  by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract)
+qed
+
+instantiation rat :: eq
+begin
+
+definition [code]:
+  "eq_class.eq a b \<longleftrightarrow> quotient_of a = quotient_of b"
+
+instance proof
+qed (simp add: eq_rat_def quotient_of_inject_eq)
+
+lemma rat_eq_refl [code nbe]:
+  "eq_class.eq (r::rat) r \<longleftrightarrow> True"
+  by (rule HOL.eq_refl)
+
+end
+
+lemma rat_less_eq_code [code]:
+  "p \<le> q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d \<le> c * b)"
+  by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+
+lemma rat_less_code [code]:
+  "p < q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d < c * b)"
+  by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+
+lemma [code]:
+  "of_rat p = (let (a, b) = quotient_of p in of_int a / of_int b)"
+  by (cases p) (simp add: quotient_of_Fract of_rat_rat)
+
+definition (in term_syntax)
+  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\<cdot>} k {\<cdot>} l"
+
+notation fcomp (infixl "o>" 60)
+notation scomp (infixl "o\<rightarrow>" 60)
+
+instantiation rat :: random
+begin
+
+definition
+  "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
+     let j = Code_Numeral.int_of (denom + 1)
+     in valterm_fract num (j, \<lambda>u. Code_Evaluation.term_of j))))"
+
+instance ..
+
+end
+
+no_notation fcomp (infixl "o>" 60)
+no_notation scomp (infixl "o\<rightarrow>" 60)
+
+text {* Setup for SML code generator *}
+
+types_code
+  rat ("(int */ int)")
+attach (term_of) {*
+fun term_of_rat (p, q) =
+  let
+    val rT = Type ("Rat.rat", [])
+  in
+    if q = 1 orelse p = 0 then HOLogic.mk_number rT p
+    else @{term "op / \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"} $
+      HOLogic.mk_number rT p $ HOLogic.mk_number rT q
+  end;
+*}
+attach (test) {*
+fun gen_rat i =
+  let
+    val p = random_range 0 i;
+    val q = random_range 1 (i + 1);
+    val g = Integer.gcd p q;
+    val p' = p div g;
+    val q' = q div g;
+    val r = (if one_of [true, false] then p' else ~ p',
+      if p' = 0 then 1 else q')
+  in
+    (r, fn () => term_of_rat r)
+  end;
+*}
+
+consts_code
+  Fract ("(_,/ _)")
+
+consts_code
+  quotient_of ("{*normalize*}")
+
+consts_code
+  "of_int :: int \<Rightarrow> rat" ("\<module>rat'_of'_int")
+attach {*
+fun rat_of_int i = (i, 1);
+*}
+
+setup {*
+  Nitpick.register_frac_type @{type_name rat}
+   [(@{const_name zero_rat_inst.zero_rat}, @{const_name Nitpick.zero_frac}),
+    (@{const_name one_rat_inst.one_rat}, @{const_name Nitpick.one_frac}),
+    (@{const_name plus_rat_inst.plus_rat}, @{const_name Nitpick.plus_frac}),
+    (@{const_name times_rat_inst.times_rat}, @{const_name Nitpick.times_frac}),
+    (@{const_name uminus_rat_inst.uminus_rat}, @{const_name Nitpick.uminus_frac}),
+    (@{const_name number_rat_inst.number_of_rat}, @{const_name Nitpick.number_of_frac}),
+    (@{const_name inverse_rat_inst.inverse_rat}, @{const_name Nitpick.inverse_frac}),
+    (@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}),
+    (@{const_name field_char_0_class.of_rat}, @{const_name Nitpick.of_frac}),
+    (@{const_name field_char_0_class.Rats}, @{const_abbrev UNIV})]
+*}
+
+lemmas [nitpick_def] = inverse_rat_inst.inverse_rat
+  number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_eq_rat
+  plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat
+  zero_rat_inst.zero_rat
+
+subsection{* Float syntax *}
+
+syntax "_Float" :: "float_const \<Rightarrow> 'a"    ("_")
+
+use "Tools/float_syntax.ML"
+setup Float_Syntax.setup
+
+text{* Test: *}
+lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::rat)"
+by simp
+
+end
--- a/src/HOL/Rational.thy	Tue Mar 02 04:31:50 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1227 +0,0 @@
-(*  Title:  HOL/Rational.thy
-    Author: Markus Wenzel, TU Muenchen
-*)
-
-header {* Rational numbers *}
-
-theory Rational
-imports GCD Archimedean_Field
-uses ("Tools/float_syntax.ML")
-begin
-
-subsection {* Rational numbers as quotient *}
-
-subsubsection {* Construction of the type of rational numbers *}
-
-definition
-  ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" where
-  "ratrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
-
-lemma ratrel_iff [simp]:
-  "(x, y) \<in> ratrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
-  by (simp add: ratrel_def)
-
-lemma refl_on_ratrel: "refl_on {x. snd x \<noteq> 0} ratrel"
-  by (auto simp add: refl_on_def ratrel_def)
-
-lemma sym_ratrel: "sym ratrel"
-  by (simp add: ratrel_def sym_def)
-
-lemma trans_ratrel: "trans ratrel"
-proof (rule transI, unfold split_paired_all)
-  fix a b a' b' a'' b'' :: int
-  assume A: "((a, b), (a', b')) \<in> ratrel"
-  assume B: "((a', b'), (a'', b'')) \<in> ratrel"
-  have "b' * (a * b'') = b'' * (a * b')" by simp
-  also from A have "a * b' = a' * b" by auto
-  also have "b'' * (a' * b) = b * (a' * b'')" by simp
-  also from B have "a' * b'' = a'' * b'" by auto
-  also have "b * (a'' * b') = b' * (a'' * b)" by simp
-  finally have "b' * (a * b'') = b' * (a'' * b)" .
-  moreover from B have "b' \<noteq> 0" by auto
-  ultimately have "a * b'' = a'' * b" by simp
-  with A B show "((a, b), (a'', b'')) \<in> ratrel" by auto
-qed
-  
-lemma equiv_ratrel: "equiv {x. snd x \<noteq> 0} ratrel"
-  by (rule equiv.intro [OF refl_on_ratrel sym_ratrel trans_ratrel])
-
-lemmas UN_ratrel = UN_equiv_class [OF equiv_ratrel]
-lemmas UN_ratrel2 = UN_equiv_class2 [OF equiv_ratrel equiv_ratrel]
-
-lemma equiv_ratrel_iff [iff]: 
-  assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
-  shows "ratrel `` {x} = ratrel `` {y} \<longleftrightarrow> (x, y) \<in> ratrel"
-  by (rule eq_equiv_class_iff, rule equiv_ratrel) (auto simp add: assms)
-
-typedef (Rat) rat = "{x. snd x \<noteq> 0} // ratrel"
-proof
-  have "(0::int, 1::int) \<in> {x. snd x \<noteq> 0}" by simp
-  then show "ratrel `` {(0, 1)} \<in> {x. snd x \<noteq> 0} // ratrel" by (rule quotientI)
-qed
-
-lemma ratrel_in_Rat [simp]: "snd x \<noteq> 0 \<Longrightarrow> ratrel `` {x} \<in> Rat"
-  by (simp add: Rat_def quotientI)
-
-declare Abs_Rat_inject [simp] Abs_Rat_inverse [simp]
-
-
-subsubsection {* Representation and basic operations *}
-
-definition
-  Fract :: "int \<Rightarrow> int \<Rightarrow> rat" where
-  [code del]: "Fract a b = Abs_Rat (ratrel `` {if b = 0 then (0, 1) else (a, b)})"
-
-code_datatype Fract
-
-lemma Rat_cases [case_names Fract, cases type: rat]:
-  assumes "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> C"
-  shows C
-  using assms by (cases q) (clarsimp simp add: Fract_def Rat_def quotient_def)
-
-lemma Rat_induct [case_names Fract, induct type: rat]:
-  assumes "\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)"
-  shows "P q"
-  using assms by (cases q) simp
-
-lemma eq_rat:
-  shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
-  and "\<And>a. Fract a 0 = Fract 0 1"
-  and "\<And>a c. Fract 0 a = Fract 0 c"
-  by (simp_all add: Fract_def)
-
-instantiation rat :: comm_ring_1
-begin
-
-definition
-  Zero_rat_def [code, code_unfold]: "0 = Fract 0 1"
-
-definition
-  One_rat_def [code, code_unfold]: "1 = Fract 1 1"
-
-definition
-  add_rat_def [code del]:
-  "q + r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
-    ratrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
-
-lemma add_rat [simp]:
-  assumes "b \<noteq> 0" and "d \<noteq> 0"
-  shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
-proof -
-  have "(\<lambda>x y. ratrel``{(fst x * snd y + fst y * snd x, snd x * snd y)})
-    respects2 ratrel"
-  by (rule equiv_ratrel [THEN congruent2_commuteI]) (simp_all add: left_distrib)
-  with assms show ?thesis by (simp add: Fract_def add_rat_def UN_ratrel2)
-qed
-
-definition
-  minus_rat_def [code del]:
-  "- q = Abs_Rat (\<Union>x \<in> Rep_Rat q. ratrel `` {(- fst x, snd x)})"
-
-lemma minus_rat [simp, code]: "- Fract a b = Fract (- a) b"
-proof -
-  have "(\<lambda>x. ratrel `` {(- fst x, snd x)}) respects ratrel"
-    by (simp add: congruent_def)
-  then show ?thesis by (simp add: Fract_def minus_rat_def UN_ratrel)
-qed
-
-lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b"
-  by (cases "b = 0") (simp_all add: eq_rat)
-
-definition
-  diff_rat_def [code del]: "q - r = q + - (r::rat)"
-
-lemma diff_rat [simp]:
-  assumes "b \<noteq> 0" and "d \<noteq> 0"
-  shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
-  using assms by (simp add: diff_rat_def)
-
-definition
-  mult_rat_def [code del]:
-  "q * r = Abs_Rat (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
-    ratrel``{(fst x * fst y, snd x * snd y)})"
-
-lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)"
-proof -
-  have "(\<lambda>x y. ratrel `` {(fst x * fst y, snd x * snd y)}) respects2 ratrel"
-    by (rule equiv_ratrel [THEN congruent2_commuteI]) simp_all
-  then show ?thesis by (simp add: Fract_def mult_rat_def UN_ratrel2)
-qed
-
-lemma mult_rat_cancel:
-  assumes "c \<noteq> 0"
-  shows "Fract (c * a) (c * b) = Fract a b"
-proof -
-  from assms have "Fract c c = Fract 1 1" by (simp add: Fract_def)
-  then show ?thesis by (simp add: mult_rat [symmetric])
-qed
-
-instance proof
-  fix q r s :: rat show "(q * r) * s = q * (r * s)" 
-    by (cases q, cases r, cases s) (simp add: eq_rat)
-next
-  fix q r :: rat show "q * r = r * q"
-    by (cases q, cases r) (simp add: eq_rat)
-next
-  fix q :: rat show "1 * q = q"
-    by (cases q) (simp add: One_rat_def eq_rat)
-next
-  fix q r s :: rat show "(q + r) + s = q + (r + s)"
-    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
-next
-  fix q r :: rat show "q + r = r + q"
-    by (cases q, cases r) (simp add: eq_rat)
-next
-  fix q :: rat show "0 + q = q"
-    by (cases q) (simp add: Zero_rat_def eq_rat)
-next
-  fix q :: rat show "- q + q = 0"
-    by (cases q) (simp add: Zero_rat_def eq_rat)
-next
-  fix q r :: rat show "q - r = q + - r"
-    by (cases q, cases r) (simp add: eq_rat)
-next
-  fix q r s :: rat show "(q + r) * s = q * s + r * s"
-    by (cases q, cases r, cases s) (simp add: eq_rat algebra_simps)
-next
-  show "(0::rat) \<noteq> 1" by (simp add: Zero_rat_def One_rat_def eq_rat)
-qed
-
-end
-
-lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1"
-  by (induct k) (simp_all add: Zero_rat_def One_rat_def)
-
-lemma of_int_rat: "of_int k = Fract k 1"
-  by (cases k rule: int_diff_cases) (simp add: of_nat_rat)
-
-lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
-  by (rule of_nat_rat [symmetric])
-
-lemma Fract_of_int_eq: "Fract k 1 = of_int k"
-  by (rule of_int_rat [symmetric])
-
-instantiation rat :: number_ring
-begin
-
-definition
-  rat_number_of_def [code del]: "number_of w = Fract w 1"
-
-instance proof
-qed (simp add: rat_number_of_def of_int_rat)
-
-end
-
-lemma rat_number_collapse [code_post]:
-  "Fract 0 k = 0"
-  "Fract 1 1 = 1"
-  "Fract (number_of k) 1 = number_of k"
-  "Fract k 0 = 0"
-  by (cases "k = 0")
-    (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def)
-
-lemma rat_number_expand [code_unfold]:
-  "0 = Fract 0 1"
-  "1 = Fract 1 1"
-  "number_of k = Fract (number_of k) 1"
-  by (simp_all add: rat_number_collapse)
-
-lemma iszero_rat [simp]:
-  "iszero (number_of k :: rat) \<longleftrightarrow> iszero (number_of k :: int)"
-  by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat)
-
-lemma Rat_cases_nonzero [case_names Fract 0]:
-  assumes Fract: "\<And>a b. q = Fract a b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> C"
-  assumes 0: "q = 0 \<Longrightarrow> C"
-  shows C
-proof (cases "q = 0")
-  case True then show C using 0 by auto
-next
-  case False
-  then obtain a b where "q = Fract a b" and "b \<noteq> 0" by (cases q) auto
-  moreover with False have "0 \<noteq> Fract a b" by simp
-  with `b \<noteq> 0` have "a \<noteq> 0" by (simp add: Zero_rat_def eq_rat)
-  with Fract `q = Fract a b` `b \<noteq> 0` show C by auto
-qed
-
-subsubsection {* Function @{text normalize} *}
-
-text{*
-Decompose a fraction into normalized, i.e. coprime numerator and denominator:
-*}
-
-definition normalize :: "rat \<Rightarrow> int \<times> int" where
-"normalize x \<equiv> THE pair. x = Fract (fst pair) (snd pair) &
-                   snd pair > 0 & gcd (fst pair) (snd pair) = 1"
-
-declare normalize_def[code del]
-
-lemma Fract_norm: "Fract (a div gcd a b) (b div gcd a b) = Fract a b"
-proof (cases "a = 0 | b = 0")
-  case True then show ?thesis by (auto simp add: eq_rat)
-next
-  let ?c = "gcd a b"
-  case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
-  then have "?c \<noteq> 0" by simp
-  then have "Fract ?c ?c = Fract 1 1" by (simp add: eq_rat)
-  moreover have "Fract (a div ?c * ?c + a mod ?c) (b div ?c * ?c + b mod ?c) = Fract a b"
-    by (simp add: semiring_div_class.mod_div_equality)
-  moreover have "a mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric])
-  moreover have "b mod ?c = 0" by (simp add: dvd_eq_mod_eq_0 [symmetric])
-  ultimately show ?thesis
-    by (simp add: mult_rat [symmetric])
-qed
-
-text{* Proof by Ren\'e Thiemann: *}
-lemma normalize_code[code]:
-"normalize (Fract a b) =
- (if b > 0 then (let g = gcd a b in (a div g, b div g))
-  else if b = 0 then (0,1)
-  else (let g = - gcd a b in (a div g, b div g)))"
-proof -
-  let ?cond = "% r p. r = Fract (fst p) (snd p) & snd p > 0 &
-                 gcd (fst p) (snd p) = 1"
-  show ?thesis
-  proof (cases "b = 0")
-    case True
-    thus ?thesis
-    proof (simp add: normalize_def)
-      show "(THE pair. ?cond (Fract a 0) pair) = (0,1)"
-      proof
-        show "?cond (Fract a 0) (0,1)"
-          by (simp add: rat_number_collapse)
-      next
-        fix pair
-        assume cond: "?cond (Fract a 0) pair"
-        show "pair = (0,1)"
-        proof (cases pair)
-          case (Pair den num)
-          with cond have num: "num > 0" by auto
-          with Pair cond have den: "den = 0" by (simp add: eq_rat)
-          show ?thesis
-          proof (cases "num = 1", simp add: Pair den)
-            case False
-            with num have gr: "num > 1" by auto
-            with den have "gcd den num = num" by auto
-            with Pair cond False gr show ?thesis by auto
-          qed
-        qed
-      qed
-    qed
-  next
-    { fix a b :: int assume b: "b > 0"
-      hence b0: "b \<noteq> 0" and "b >= 0" by auto
-      let ?g = "gcd a b"
-      from b0 have g0: "?g \<noteq> 0" by auto
-      then have gp: "?g > 0" by simp
-      then have gs: "?g <= b" by (metis b gcd_le2_int)
-      from gcd_dvd1_int[of a b] obtain a' where a': "a = ?g * a'"
-        unfolding dvd_def by auto
-      from gcd_dvd2_int[of a b] obtain b' where b': "b = ?g * b'"
-        unfolding dvd_def by auto
-      hence b'2: "b' * ?g = b" by (simp add: ring_simps)
-      with b0 have b'0: "b' \<noteq> 0" by auto
-      from b b' zero_less_mult_iff[of ?g b'] gp have b'p: "b' > 0" by arith
-      have "normalize (Fract a b) = (a div ?g, b div ?g)"
-      proof (simp add: normalize_def)
-        show "(THE pair. ?cond (Fract a b) pair) = (a div ?g, b div ?g)"
-        proof
-          have "1 = b div b" using b0 by auto
-          also have "\<dots> <= b div ?g" by (rule zdiv_mono2[OF `b >= 0` gp gs])
-          finally have div0: "b div ?g > 0" by simp
-          show "?cond (Fract a b) (a div ?g, b div ?g)"
-            by (simp add: b0 Fract_norm div_gcd_coprime_int div0)
-        next
-          fix pair assume cond: "?cond (Fract a b) pair"
-          show "pair = (a div ?g, b div ?g)"
-          proof (cases pair)
-            case (Pair den num)
-            with cond
-            have num: "num > 0" and num0: "num \<noteq> 0" and gcd: "gcd den num = 1"
-              by auto
-            obtain g where g: "g = ?g" by auto
-            with gp have gg0: "g > 0" by auto
-            from cond Pair eq_rat(1)[OF b0 num0]
-            have eq: "a * num = den * b" by auto
-            hence "a' * g * num = den * g * b'"
-              using a'[simplified g[symmetric]] b'[simplified g[symmetric]]
-              by simp
-            hence "a' * num * g = b' * den * g" by (simp add: algebra_simps)
-            hence eq2: "a' * num = b' * den" using gg0 by auto
-            have "a div ?g = ?g * a' div ?g" using a' by force
-            hence adiv: "a div ?g = a'" using g0 by auto
-            have "b div ?g = ?g * b' div ?g" using b' by force
-            hence bdiv: "b div ?g = b'" using g0 by auto
-            from div_gcd_coprime_int[of a b] b0
-            have "gcd (a div ?g) (b div ?g) = 1" by auto
-            with adiv bdiv have gcd2: "gcd a' b' = 1" by auto
-            from gcd have gcd3: "gcd num den = 1"
-              by (simp add: gcd_commute_int[of den num])
-            from gcd2 have gcd4: "gcd b' a' = 1"
-              by (simp add: gcd_commute_int[of a' b'])
-            have one: "num dvd b'"
-              by (metis coprime_dvd_mult_int[OF gcd3] dvd_triv_right eq2)
-            have two: "b' dvd num"
-              by (metis coprime_dvd_mult_int[OF gcd4] dvd_triv_left eq2 zmult_commute)
-            from zdvd_antisym_abs[OF one two] b'p num
-            have numb': "num = b'" by auto
-            with eq2 b'0 have "a' = den" by auto
-            with numb' adiv bdiv Pair show ?thesis by simp
-          qed
-        qed
-      qed
-    }
-    note main = this
-    assume "b \<noteq> 0"
-    hence "b > 0 | b < 0" by arith
-    thus ?thesis
-    proof
-      assume b: "b > 0" thus ?thesis by (simp add: Let_def main[OF b])
-    next
-      assume b: "b < 0"
-      thus ?thesis
-        by(simp add:main Let_def minus_rat_cancel[of a b, symmetric]
-                    zdiv_zminus2 del:minus_rat_cancel)
-    qed
-  qed
-qed
-
-lemma normalize_id: "normalize (Fract a b) = (p,q) \<Longrightarrow> Fract p q = Fract a b"
-by(auto simp add: normalize_code Let_def Fract_norm dvd_div_neg rat_number_collapse
-        split:split_if_asm)
-
-lemma normalize_denom_pos: "normalize (Fract a b) = (p,q) \<Longrightarrow> q > 0"
-by(auto simp add: normalize_code Let_def dvd_div_neg pos_imp_zdiv_neg_iff nonneg1_imp_zdiv_pos_iff
-        split:split_if_asm)
-
-lemma normalize_coprime: "normalize (Fract a b) = (p,q) \<Longrightarrow> coprime p q"
-by(auto simp add: normalize_code Let_def dvd_div_neg div_gcd_coprime_int
-        split:split_if_asm)
-
-
-subsubsection {* The field of rational numbers *}
-
-instantiation rat :: "{field, division_by_zero}"
-begin
-
-definition
-  inverse_rat_def [code del]:
-  "inverse q = Abs_Rat (\<Union>x \<in> Rep_Rat q.
-     ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
-
-lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a"
-proof -
-  have "(\<lambda>x. ratrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)}) respects ratrel"
-    by (auto simp add: congruent_def mult_commute)
-  then show ?thesis by (simp add: Fract_def inverse_rat_def UN_ratrel)
-qed
-
-definition
-  divide_rat_def [code del]: "q / r = q * inverse (r::rat)"
-
-lemma divide_rat [simp]: "Fract a b / Fract c d = Fract (a * d) (b * c)"
-  by (simp add: divide_rat_def)
-
-instance proof
-  show "inverse 0 = (0::rat)" by (simp add: rat_number_expand)
-    (simp add: rat_number_collapse)
-next
-  fix q :: rat
-  assume "q \<noteq> 0"
-  then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero)
-   (simp_all add: rat_number_expand eq_rat)
-next
-  fix q r :: rat
-  show "q / r = q * inverse r" by (simp add: divide_rat_def)
-qed
-
-end
-
-
-subsubsection {* Various *}
-
-lemma Fract_add_one: "n \<noteq> 0 ==> Fract (m + n) n = Fract m n + 1"
-  by (simp add: rat_number_expand)
-
-lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l"
-  by (simp add: Fract_of_int_eq [symmetric])
-
-lemma Fract_number_of_quotient [code_post]:
-  "Fract (number_of k) (number_of l) = number_of k / number_of l"
-  unfolding Fract_of_int_quotient number_of_is_id number_of_eq ..
-
-lemma Fract_1_number_of [code_post]:
-  "Fract 1 (number_of k) = 1 / number_of k"
-  unfolding Fract_of_int_quotient number_of_eq by simp
-
-subsubsection {* The ordered field of rational numbers *}
-
-instantiation rat :: linorder
-begin
-
-definition
-  le_rat_def [code del]:
-   "q \<le> r \<longleftrightarrow> contents (\<Union>x \<in> Rep_Rat q. \<Union>y \<in> Rep_Rat r.
-      {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})"
-
-lemma le_rat [simp]:
-  assumes "b \<noteq> 0" and "d \<noteq> 0"
-  shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
-proof -
-  have "(\<lambda>x y. {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})
-    respects2 ratrel"
-  proof (clarsimp simp add: congruent2_def)
-    fix a b a' b' c d c' d'::int
-    assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-    assume eq1: "a * b' = a' * b"
-    assume eq2: "c * d' = c' * d"
-
-    let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
-    {
-      fix a b c d x :: int assume x: "x \<noteq> 0"
-      have "?le a b c d = ?le (a * x) (b * x) c d"
-      proof -
-        from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
-        hence "?le a b c d =
-            ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
-          by (simp add: mult_le_cancel_right)
-        also have "... = ?le (a * x) (b * x) c d"
-          by (simp add: mult_ac)
-        finally show ?thesis .
-      qed
-    } note le_factor = this
-
-    let ?D = "b * d" and ?D' = "b' * d'"
-    from neq have D: "?D \<noteq> 0" by simp
-    from neq have "?D' \<noteq> 0" by simp
-    hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
-      by (rule le_factor)
-    also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')" 
-      by (simp add: mult_ac)
-    also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
-      by (simp only: eq1 eq2)
-    also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
-      by (simp add: mult_ac)
-    also from D have "... = ?le a' b' c' d'"
-      by (rule le_factor [symmetric])
-    finally show "?le a b c d = ?le a' b' c' d'" .
-  qed
-  with assms show ?thesis by (simp add: Fract_def le_rat_def UN_ratrel2)
-qed
-
-definition
-  less_rat_def [code del]: "z < (w::rat) \<longleftrightarrow> z \<le> w \<and> z \<noteq> w"
-
-lemma less_rat [simp]:
-  assumes "b \<noteq> 0" and "d \<noteq> 0"
-  shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
-  using assms by (simp add: less_rat_def eq_rat order_less_le)
-
-instance proof
-  fix q r s :: rat
-  {
-    assume "q \<le> r" and "r \<le> s"
-    show "q \<le> s"
-    proof (insert prems, induct q, induct r, induct s)
-      fix a b c d e f :: int
-      assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
-      show "Fract a b \<le> Fract e f"
-      proof -
-        from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
-          by (auto simp add: zero_less_mult_iff linorder_neq_iff)
-        have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
-        proof -
-          from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-            by simp
-          with ff show ?thesis by (simp add: mult_le_cancel_right)
-        qed
-        also have "... = (c * f) * (d * f) * (b * b)" by algebra
-        also have "... \<le> (e * d) * (d * f) * (b * b)"
-        proof -
-          from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
-            by simp
-          with bb show ?thesis by (simp add: mult_le_cancel_right)
-        qed
-        finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
-          by (simp only: mult_ac)
-        with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
-          by (simp add: mult_le_cancel_right)
-        with neq show ?thesis by simp
-      qed
-    qed
-  next
-    assume "q \<le> r" and "r \<le> q"
-    show "q = r"
-    proof (insert prems, induct q, induct r)
-      fix a b c d :: int
-      assume neq: "b \<noteq> 0"  "d \<noteq> 0"
-      assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
-      show "Fract a b = Fract c d"
-      proof -
-        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-          by simp
-        also have "... \<le> (a * d) * (b * d)"
-        proof -
-          from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
-            by simp
-          thus ?thesis by (simp only: mult_ac)
-        qed
-        finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
-        moreover from neq have "b * d \<noteq> 0" by simp
-        ultimately have "a * d = c * b" by simp
-        with neq show ?thesis by (simp add: eq_rat)
-      qed
-    qed
-  next
-    show "q \<le> q"
-      by (induct q) simp
-    show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
-      by (induct q, induct r) (auto simp add: le_less mult_commute)
-    show "q \<le> r \<or> r \<le> q"
-      by (induct q, induct r)
-         (simp add: mult_commute, rule linorder_linear)
-  }
-qed
-
-end
-
-instantiation rat :: "{distrib_lattice, abs_if, sgn_if}"
-begin
-
-definition
-  abs_rat_def [code del]: "\<bar>q\<bar> = (if q < 0 then -q else (q::rat))"
-
-lemma abs_rat [simp, code]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
-  by (auto simp add: abs_rat_def zabs_def Zero_rat_def not_less le_less eq_rat zero_less_mult_iff)
-
-definition
-  sgn_rat_def [code del]: "sgn (q::rat) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
-
-lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)"
-  unfolding Fract_of_int_eq
-  by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat)
-    (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff)
-
-definition
-  "(inf \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = min"
-
-definition
-  "(sup \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat) = max"
-
-instance by intro_classes
-  (auto simp add: abs_rat_def sgn_rat_def min_max.sup_inf_distrib1 inf_rat_def sup_rat_def)
-
-end
-
-instance rat :: linordered_field
-proof
-  fix q r s :: rat
-  show "q \<le> r ==> s + q \<le> s + r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: int
-    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-    assume le: "Fract a b \<le> Fract c d"
-    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
-    proof -
-      let ?F = "f * f" from neq have F: "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
-        by simp
-      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
-        by (simp add: mult_le_cancel_right)
-      with neq show ?thesis by (simp add: mult_ac int_distrib)
-    qed
-  qed
-  show "q < r ==> 0 < s ==> s * q < s * r"
-  proof (induct q, induct r, induct s)
-    fix a b c d e f :: int
-    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
-    assume le: "Fract a b < Fract c d"
-    assume gt: "0 < Fract e f"
-    show "Fract e f * Fract a b < Fract e f * Fract c d"
-    proof -
-      let ?E = "e * f" and ?F = "f * f"
-      from neq gt have "0 < ?E"
-        by (auto simp add: Zero_rat_def order_less_le eq_rat)
-      moreover from neq have "0 < ?F"
-        by (auto simp add: zero_less_mult_iff)
-      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
-        by simp
-      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
-        by (simp add: mult_less_cancel_right)
-      with neq show ?thesis
-        by (simp add: mult_ac)
-    qed
-  qed
-qed auto
-
-lemma Rat_induct_pos [case_names Fract, induct type: rat]:
-  assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
-  shows "P q"
-proof (cases q)
-  have step': "\<And>a b. b < 0 \<Longrightarrow> P (Fract a b)"
-  proof -
-    fix a::int and b::int
-    assume b: "b < 0"
-    hence "0 < -b" by simp
-    hence "P (Fract (-a) (-b))" by (rule step)
-    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
-  qed
-  case (Fract a b)
-  thus "P q" by (force simp add: linorder_neq_iff step step')
-qed
-
-lemma zero_less_Fract_iff:
-  "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
-  by (simp add: Zero_rat_def zero_less_mult_iff)
-
-lemma Fract_less_zero_iff:
-  "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
-  by (simp add: Zero_rat_def mult_less_0_iff)
-
-lemma zero_le_Fract_iff:
-  "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
-  by (simp add: Zero_rat_def zero_le_mult_iff)
-
-lemma Fract_le_zero_iff:
-  "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
-  by (simp add: Zero_rat_def mult_le_0_iff)
-
-lemma one_less_Fract_iff:
-  "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
-  by (simp add: One_rat_def mult_less_cancel_right_disj)
-
-lemma Fract_less_one_iff:
-  "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
-  by (simp add: One_rat_def mult_less_cancel_right_disj)
-
-lemma one_le_Fract_iff:
-  "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
-  by (simp add: One_rat_def mult_le_cancel_right)
-
-lemma Fract_le_one_iff:
-  "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
-  by (simp add: One_rat_def mult_le_cancel_right)
-
-
-subsubsection {* Rationals are an Archimedean field *}
-
-lemma rat_floor_lemma:
-  shows "of_int (a div b) \<le> Fract a b \<and> Fract a b < of_int (a div b + 1)"
-proof -
-  have "Fract a b = of_int (a div b) + Fract (a mod b) b"
-    by (cases "b = 0", simp, simp add: of_int_rat)
-  moreover have "0 \<le> Fract (a mod b) b \<and> Fract (a mod b) b < 1"
-    unfolding Fract_of_int_quotient
-    by (rule linorder_cases [of b 0])
-       (simp add: divide_nonpos_neg, simp, simp add: divide_nonneg_pos)
-  ultimately show ?thesis by simp
-qed
-
-instance rat :: archimedean_field
-proof
-  fix r :: rat
-  show "\<exists>z. r \<le> of_int z"
-  proof (induct r)
-    case (Fract a b)
-    have "Fract a b \<le> of_int (a div b + 1)"
-      using rat_floor_lemma [of a b] by simp
-    then show "\<exists>z. Fract a b \<le> of_int z" ..
-  qed
-qed
-
-lemma floor_Fract: "floor (Fract a b) = a div b"
-  using rat_floor_lemma [of a b]
-  by (simp add: floor_unique)
-
-
-subsection {* Linear arithmetic setup *}
-
-declaration {*
-  K (Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2]
-    (* not needed because x < (y::nat) can be rewritten as Suc x <= y: of_nat_less_iff RS iffD2 *)
-  #> Lin_Arith.add_inj_thms [@{thm of_int_le_iff} RS iffD2, @{thm of_int_eq_iff} RS iffD2]
-    (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *)
-  #> Lin_Arith.add_simps [@{thm neg_less_iff_less},
-      @{thm True_implies_equals},
-      read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib},
-      @{thm divide_1}, @{thm divide_zero_left},
-      @{thm times_divide_eq_right}, @{thm times_divide_eq_left},
-      @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym,
-      @{thm of_int_minus}, @{thm of_int_diff},
-      @{thm of_int_of_nat_eq}]
-  #> Lin_Arith.add_simprocs Numeral_Simprocs.field_cancel_numeral_factors
-  #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat => rat"})
-  #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int => rat"}))
-*}
-
-
-subsection {* Embedding from Rationals to other Fields *}
-
-class field_char_0 = field + ring_char_0
-
-subclass (in linordered_field) field_char_0 ..
-
-context field_char_0
-begin
-
-definition of_rat :: "rat \<Rightarrow> 'a" where
-  [code del]: "of_rat q = contents (\<Union>(a,b) \<in> Rep_Rat q. {of_int a / of_int b})"
-
-end
-
-lemma of_rat_congruent:
-  "(\<lambda>(a, b). {of_int a / of_int b :: 'a::field_char_0}) respects ratrel"
-apply (rule congruent.intro)
-apply (clarsimp simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
-apply (simp only: of_int_mult [symmetric])
-done
-
-lemma of_rat_rat: "b \<noteq> 0 \<Longrightarrow> of_rat (Fract a b) = of_int a / of_int b"
-  unfolding Fract_def of_rat_def by (simp add: UN_ratrel of_rat_congruent)
-
-lemma of_rat_0 [simp]: "of_rat 0 = 0"
-by (simp add: Zero_rat_def of_rat_rat)
-
-lemma of_rat_1 [simp]: "of_rat 1 = 1"
-by (simp add: One_rat_def of_rat_rat)
-
-lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b"
-by (induct a, induct b, simp add: of_rat_rat add_frac_eq)
-
-lemma of_rat_minus: "of_rat (- a) = - of_rat a"
-by (induct a, simp add: of_rat_rat)
-
-lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
-by (simp only: diff_minus of_rat_add of_rat_minus)
-
-lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
-apply (induct a, induct b, simp add: of_rat_rat)
-apply (simp add: divide_inverse nonzero_inverse_mult_distrib mult_ac)
-done
-
-lemma nonzero_of_rat_inverse:
-  "a \<noteq> 0 \<Longrightarrow> of_rat (inverse a) = inverse (of_rat a)"
-apply (rule inverse_unique [symmetric])
-apply (simp add: of_rat_mult [symmetric])
-done
-
-lemma of_rat_inverse:
-  "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
-   inverse (of_rat a)"
-by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
-
-lemma nonzero_of_rat_divide:
-  "b \<noteq> 0 \<Longrightarrow> of_rat (a / b) = of_rat a / of_rat b"
-by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
-
-lemma of_rat_divide:
-  "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
-   = of_rat a / of_rat b"
-by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
-
-lemma of_rat_power:
-  "(of_rat (a ^ n)::'a::field_char_0) = of_rat a ^ n"
-by (induct n) (simp_all add: of_rat_mult)
-
-lemma of_rat_eq_iff [simp]: "(of_rat a = of_rat b) = (a = b)"
-apply (induct a, induct b)
-apply (simp add: of_rat_rat eq_rat)
-apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq)
-apply (simp only: of_int_mult [symmetric] of_int_eq_iff)
-done
-
-lemma of_rat_less:
-  "(of_rat r :: 'a::linordered_field) < of_rat s \<longleftrightarrow> r < s"
-proof (induct r, induct s)
-  fix a b c d :: int
-  assume not_zero: "b > 0" "d > 0"
-  then have "b * d > 0" by (rule mult_pos_pos)
-  have of_int_divide_less_eq:
-    "(of_int a :: 'a) / of_int b < of_int c / of_int d
-      \<longleftrightarrow> (of_int a :: 'a) * of_int d < of_int c * of_int b"
-    using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq)
-  show "(of_rat (Fract a b) :: 'a::linordered_field) < of_rat (Fract c d)
-    \<longleftrightarrow> Fract a b < Fract c d"
-    using not_zero `b * d > 0`
-    by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult)
-qed
-
-lemma of_rat_less_eq:
-  "(of_rat r :: 'a::linordered_field) \<le> of_rat s \<longleftrightarrow> r \<le> s"
-  unfolding le_less by (auto simp add: of_rat_less)
-
-lemmas of_rat_eq_0_iff [simp] = of_rat_eq_iff [of _ 0, simplified]
-
-lemma of_rat_eq_id [simp]: "of_rat = id"
-proof
-  fix a
-  show "of_rat a = id a"
-  by (induct a)
-     (simp add: of_rat_rat Fract_of_int_eq [symmetric])
-qed
-
-text{*Collapse nested embeddings*}
-lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
-by (induct n) (simp_all add: of_rat_add)
-
-lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z"
-by (cases z rule: int_diff_cases) (simp add: of_rat_diff)
-
-lemma of_rat_number_of_eq [simp]:
-  "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})"
-by (simp add: number_of_eq)
-
-lemmas zero_rat = Zero_rat_def
-lemmas one_rat = One_rat_def
-
-abbreviation
-  rat_of_nat :: "nat \<Rightarrow> rat"
-where
-  "rat_of_nat \<equiv> of_nat"
-
-abbreviation
-  rat_of_int :: "int \<Rightarrow> rat"
-where
-  "rat_of_int \<equiv> of_int"
-
-subsection {* The Set of Rational Numbers *}
-
-context field_char_0
-begin
-
-definition
-  Rats  :: "'a set" where
-  [code del]: "Rats = range of_rat"
-
-notation (xsymbols)
-  Rats  ("\<rat>")
-
-end
-
-lemma Rats_of_rat [simp]: "of_rat r \<in> Rats"
-by (simp add: Rats_def)
-
-lemma Rats_of_int [simp]: "of_int z \<in> Rats"
-by (subst of_rat_of_int_eq [symmetric], rule Rats_of_rat)
-
-lemma Rats_of_nat [simp]: "of_nat n \<in> Rats"
-by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat)
-
-lemma Rats_number_of [simp]:
-  "(number_of w::'a::{number_ring,field_char_0}) \<in> Rats"
-by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat)
-
-lemma Rats_0 [simp]: "0 \<in> Rats"
-apply (unfold Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_0 [symmetric])
-done
-
-lemma Rats_1 [simp]: "1 \<in> Rats"
-apply (unfold Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_1 [symmetric])
-done
-
-lemma Rats_add [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a + b \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_add [symmetric])
-done
-
-lemma Rats_minus [simp]: "a \<in> Rats \<Longrightarrow> - a \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_minus [symmetric])
-done
-
-lemma Rats_diff [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a - b \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_diff [symmetric])
-done
-
-lemma Rats_mult [simp]: "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a * b \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_mult [symmetric])
-done
-
-lemma nonzero_Rats_inverse:
-  fixes a :: "'a::field_char_0"
-  shows "\<lbrakk>a \<in> Rats; a \<noteq> 0\<rbrakk> \<Longrightarrow> inverse a \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_rat_inverse [symmetric])
-done
-
-lemma Rats_inverse [simp]:
-  fixes a :: "'a::{field_char_0,division_by_zero}"
-  shows "a \<in> Rats \<Longrightarrow> inverse a \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_inverse [symmetric])
-done
-
-lemma nonzero_Rats_divide:
-  fixes a b :: "'a::field_char_0"
-  shows "\<lbrakk>a \<in> Rats; b \<in> Rats; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (erule nonzero_of_rat_divide [symmetric])
-done
-
-lemma Rats_divide [simp]:
-  fixes a b :: "'a::{field_char_0,division_by_zero}"
-  shows "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_divide [symmetric])
-done
-
-lemma Rats_power [simp]:
-  fixes a :: "'a::field_char_0"
-  shows "a \<in> Rats \<Longrightarrow> a ^ n \<in> Rats"
-apply (auto simp add: Rats_def)
-apply (rule range_eqI)
-apply (rule of_rat_power [symmetric])
-done
-
-lemma Rats_cases [cases set: Rats]:
-  assumes "q \<in> \<rat>"
-  obtains (of_rat) r where "q = of_rat r"
-  unfolding Rats_def
-proof -
-  from `q \<in> \<rat>` have "q \<in> range of_rat" unfolding Rats_def .
-  then obtain r where "q = of_rat r" ..
-  then show thesis ..
-qed
-
-lemma Rats_induct [case_names of_rat, induct set: Rats]:
-  "q \<in> \<rat> \<Longrightarrow> (\<And>r. P (of_rat r)) \<Longrightarrow> P q"
-  by (rule Rats_cases) auto
-
-
-subsection {* Implementation of rational numbers as pairs of integers *}
-
-definition Fract_norm :: "int \<Rightarrow> int \<Rightarrow> rat" where
-  [simp, code del]: "Fract_norm a b = Fract a b"
-
-lemma Fract_norm_code [code]: "Fract_norm a b = (if a = 0 \<or> b = 0 then 0 else let c = gcd a b in
-  if b > 0 then Fract (a div c) (b div c) else Fract (- (a div c)) (- (b div c)))"
-  by (simp add: eq_rat Zero_rat_def Let_def Fract_norm)
-
-lemma [code]:
-  "of_rat (Fract a b) = (if b \<noteq> 0 then of_int a / of_int b else 0)"
-  by (cases "b = 0") (simp_all add: rat_number_collapse of_rat_rat)
-
-instantiation rat :: eq
-begin
-
-definition [code del]: "eq_class.eq (a\<Colon>rat) b \<longleftrightarrow> a - b = 0"
-
-instance by default (simp add: eq_rat_def)
-
-lemma rat_eq_code [code]:
-  "eq_class.eq (Fract a b) (Fract c d) \<longleftrightarrow> (if b = 0
-       then c = 0 \<or> d = 0
-     else if d = 0
-       then a = 0 \<or> b = 0
-     else a * d = b * c)"
-  by (auto simp add: eq eq_rat)
-
-lemma rat_eq_refl [code nbe]:
-  "eq_class.eq (r::rat) r \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
-
-end
-
-lemma le_rat':
-  assumes "b \<noteq> 0"
-    and "d \<noteq> 0"
-  shows "Fract a b \<le> Fract c d \<longleftrightarrow> a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d"
-proof -
-  have abs_sgn: "\<And>k::int. \<bar>k\<bar> = k * sgn k" unfolding abs_if sgn_if by simp
-  have "a * d * (b * d) \<le> c * b * (b * d) \<longleftrightarrow> a * d * (sgn b * sgn d) \<le> c * b * (sgn b * sgn d)"
-  proof (cases "b * d > 0")
-    case True
-    moreover from True have "sgn b * sgn d = 1"
-      by (simp add: sgn_times [symmetric] sgn_1_pos)
-    ultimately show ?thesis by (simp add: mult_le_cancel_right)
-  next
-    case False with assms have "b * d < 0" by (simp add: less_le)
-    moreover from this have "sgn b * sgn d = - 1"
-      by (simp only: sgn_times [symmetric] sgn_1_neg)
-    ultimately show ?thesis by (simp add: mult_le_cancel_right)
-  qed
-  also have "\<dots> \<longleftrightarrow> a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d"
-    by (simp add: abs_sgn mult_ac)
-  finally show ?thesis using assms by simp
-qed
-
-lemma less_rat': 
-  assumes "b \<noteq> 0"
-    and "d \<noteq> 0"
-  shows "Fract a b < Fract c d \<longleftrightarrow> a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d"
-proof -
-  have abs_sgn: "\<And>k::int. \<bar>k\<bar> = k * sgn k" unfolding abs_if sgn_if by simp
-  have "a * d * (b * d) < c * b * (b * d) \<longleftrightarrow> a * d * (sgn b * sgn d) < c * b * (sgn b * sgn d)"
-  proof (cases "b * d > 0")
-    case True
-    moreover from True have "sgn b * sgn d = 1"
-      by (simp add: sgn_times [symmetric] sgn_1_pos)
-    ultimately show ?thesis by (simp add: mult_less_cancel_right)
-  next
-    case False with assms have "b * d < 0" by (simp add: less_le)
-    moreover from this have "sgn b * sgn d = - 1"
-      by (simp only: sgn_times [symmetric] sgn_1_neg)
-    ultimately show ?thesis by (simp add: mult_less_cancel_right)
-  qed
-  also have "\<dots> \<longleftrightarrow> a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d"
-    by (simp add: abs_sgn mult_ac)
-  finally show ?thesis using assms by simp
-qed
-
-lemma rat_le_eq_code [code]:
-  "Fract a b < Fract c d \<longleftrightarrow> (if b = 0
-       then sgn c * sgn d > 0
-     else if d = 0
-       then sgn a * sgn b < 0
-     else a * \<bar>d\<bar> * sgn b < c * \<bar>b\<bar> * sgn d)"
-  by (auto simp add: sgn_times mult_less_0_iff zero_less_mult_iff less_rat' eq_rat simp del: less_rat)
-
-lemma rat_less_eq_code [code]:
-  "Fract a b \<le> Fract c d \<longleftrightarrow> (if b = 0
-       then sgn c * sgn d \<ge> 0
-     else if d = 0
-       then sgn a * sgn b \<le> 0
-     else a * \<bar>d\<bar> * sgn b \<le> c * \<bar>b\<bar> * sgn d)"
-  by (auto simp add: sgn_times mult_le_0_iff zero_le_mult_iff le_rat' eq_rat simp del: le_rat)
-    (auto simp add: le_less not_less sgn_0_0)
-
-
-lemma rat_plus_code [code]:
-  "Fract a b + Fract c d = (if b = 0
-     then Fract c d
-   else if d = 0
-     then Fract a b
-   else Fract_norm (a * d + c * b) (b * d))"
-  by (simp add: eq_rat, simp add: Zero_rat_def)
-
-lemma rat_times_code [code]:
-  "Fract a b * Fract c d = Fract_norm (a * c) (b * d)"
-  by simp
-
-lemma rat_minus_code [code]:
-  "Fract a b - Fract c d = (if b = 0
-     then Fract (- c) d
-   else if d = 0
-     then Fract a b
-   else Fract_norm (a * d - c * b) (b * d))"
-  by (simp add: eq_rat, simp add: Zero_rat_def)
-
-lemma rat_inverse_code [code]:
-  "inverse (Fract a b) = (if b = 0 then Fract 1 0
-    else if a < 0 then Fract (- b) (- a)
-    else Fract b a)"
-  by (simp add: eq_rat)
-
-lemma rat_divide_code [code]:
-  "Fract a b / Fract c d = Fract_norm (a * d) (b * c)"
-  by simp
-
-definition (in term_syntax)
-  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
-  [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\<cdot>} k {\<cdot>} l"
-
-notation fcomp (infixl "o>" 60)
-notation scomp (infixl "o\<rightarrow>" 60)
-
-instantiation rat :: random
-begin
-
-definition
-  "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
-     let j = Code_Numeral.int_of (denom + 1)
-     in valterm_fract num (j, \<lambda>u. Code_Evaluation.term_of j))))"
-
-instance ..
-
-end
-
-no_notation fcomp (infixl "o>" 60)
-no_notation scomp (infixl "o\<rightarrow>" 60)
-
-hide (open) const Fract_norm
-
-text {* Setup for SML code generator *}
-
-types_code
-  rat ("(int */ int)")
-attach (term_of) {*
-fun term_of_rat (p, q) =
-  let
-    val rT = Type ("Rational.rat", [])
-  in
-    if q = 1 orelse p = 0 then HOLogic.mk_number rT p
-    else @{term "op / \<Colon> rat \<Rightarrow> rat \<Rightarrow> rat"} $
-      HOLogic.mk_number rT p $ HOLogic.mk_number rT q
-  end;
-*}
-attach (test) {*
-fun gen_rat i =
-  let
-    val p = random_range 0 i;
-    val q = random_range 1 (i + 1);
-    val g = Integer.gcd p q;
-    val p' = p div g;
-    val q' = q div g;
-    val r = (if one_of [true, false] then p' else ~ p',
-      if p' = 0 then 1 else q')
-  in
-    (r, fn () => term_of_rat r)
-  end;
-*}
-
-consts_code
-  Fract ("(_,/ _)")
-
-consts_code
-  "of_int :: int \<Rightarrow> rat" ("\<module>rat'_of'_int")
-attach {*
-fun rat_of_int i = (i, 1);
-*}
-
-setup {*
-  Nitpick.register_frac_type @{type_name rat}
-   [(@{const_name zero_rat_inst.zero_rat}, @{const_name Nitpick.zero_frac}),
-    (@{const_name one_rat_inst.one_rat}, @{const_name Nitpick.one_frac}),
-    (@{const_name plus_rat_inst.plus_rat}, @{const_name Nitpick.plus_frac}),
-    (@{const_name times_rat_inst.times_rat}, @{const_name Nitpick.times_frac}),
-    (@{const_name uminus_rat_inst.uminus_rat}, @{const_name Nitpick.uminus_frac}),
-    (@{const_name number_rat_inst.number_of_rat}, @{const_name Nitpick.number_of_frac}),
-    (@{const_name inverse_rat_inst.inverse_rat}, @{const_name Nitpick.inverse_frac}),
-    (@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}),
-    (@{const_name field_char_0_class.of_rat}, @{const_name Nitpick.of_frac}),
-    (@{const_name field_char_0_class.Rats}, @{const_name UNIV})]
-*}
-
-lemmas [nitpick_def] = inverse_rat_inst.inverse_rat
-  number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_eq_rat
-  plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat
-  zero_rat_inst.zero_rat
-
-subsection{* Float syntax *}
-
-syntax "_Float" :: "float_const \<Rightarrow> 'a"    ("_")
-
-use "Tools/float_syntax.ML"
-setup Float_Syntax.setup
-
-text{* Test: *}
-lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::rat)"
-by simp
-
-end
--- a/src/HOL/Recdef.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Recdef.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -27,15 +27,17 @@
   wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
             wfrec_rel R F x (F g x)"
 
-constdefs
-  cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
+definition
+  cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" where
   "cut f r x == (%y. if (y,x):r then f y else undefined)"
 
-  adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
+definition
+  adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" where
   "adm_wf R F == ALL f g x.
      (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
 
-  wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
+definition
+  wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" where
   [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
 
 subsection{*Well-Founded Recursion*}
--- a/src/HOL/SET_Protocol/Message_SET.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/SET_Protocol/Message_SET.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -48,8 +48,7 @@
 text{*The inverse of a symmetric key is itself; that of a public key
       is the private key and vice versa*}
 
-constdefs
-  symKeys :: "key set"
+definition symKeys :: "key set" where
   "symKeys == {K. invKey K = K}"
 
 text{*Agents. We allow any number of certification authorities, cardholders
@@ -81,8 +80,7 @@
   "{|x, y|}"      == "CONST MPair x y"
 
 
-constdefs
-  nat_of_agent :: "agent => nat"
+definition nat_of_agent :: "agent => nat" where
    "nat_of_agent == agent_case (curry nat2_to_nat 0)
                                (curry nat2_to_nat 1)
                                (curry nat2_to_nat 2)
--- a/src/HOL/SET_Protocol/Public_SET.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/SET_Protocol/Public_SET.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -161,21 +161,19 @@
 
 subsection{*Encryption Primitives*}
 
-constdefs
-
-  EXcrypt :: "[key,key,msg,msg] => msg"
+definition EXcrypt :: "[key,key,msg,msg] => msg" where
   --{*Extra Encryption*}
     (*K: the symmetric key   EK: the public encryption key*)
     "EXcrypt K EK M m ==
        {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m|}|}"
 
-  EXHcrypt :: "[key,key,msg,msg] => msg"
+definition EXHcrypt :: "[key,key,msg,msg] => msg" where
   --{*Extra Encryption with Hashing*}
     (*K: the symmetric key   EK: the public encryption key*)
     "EXHcrypt K EK M m ==
        {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m, Hash M|}|}"
 
-  Enc :: "[key,key,key,msg] => msg"
+definition Enc :: "[key,key,key,msg] => msg" where
   --{*Simple Encapsulation with SIGNATURE*}
     (*SK: the sender's signing key
       K: the symmetric key
@@ -183,7 +181,7 @@
     "Enc SK K EK M ==
        {|Crypt K (sign SK M), Crypt EK (Key K)|}"
 
-  EncB :: "[key,key,key,msg,msg] => msg"
+definition EncB :: "[key,key,key,msg,msg] => msg" where
   --{*Encapsulation with Baggage.  Keys as above, and baggage b.*}
     "EncB SK K EK M b == 
        {|Enc SK K EK {|M, Hash b|}, b|}"
--- a/src/HOL/SMT/Tools/smt_translate.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/SMT/Tools/smt_translate.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -253,9 +253,9 @@
       | atoms_ord (SNum (i, _), SNum (j, _)) = int_ord (i, j)
       | atoms_ord _ = sys_error "atoms_ord"
 
-    fun types_ord (SConst (_, T), SConst (_, U)) = TermOrd.typ_ord (T, U)
-      | types_ord (SFree (_, T), SFree (_, U)) = TermOrd.typ_ord (T, U)
-      | types_ord (SNum (_, T), SNum (_, U)) = TermOrd.typ_ord (T, U)
+    fun types_ord (SConst (_, T), SConst (_, U)) = Term_Ord.typ_ord (T, U)
+      | types_ord (SFree (_, T), SFree (_, U)) = Term_Ord.typ_ord (T, U)
+      | types_ord (SNum (_, T), SNum (_, U)) = Term_Ord.typ_ord (T, U)
       | types_ord _ = sys_error "types_ord"
 
     fun fast_sym_ord tu =
--- a/src/HOL/Set.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Set.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1586,8 +1586,7 @@
 
 subsubsection {* Inverse image of a function *}
 
-constdefs
-  vimage :: "('a => 'b) => 'b set => 'a set"    (infixr "-`" 90)
+definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where
   [code del]: "f -` B == {x. f x : B}"
 
 lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)"
--- a/src/HOL/Statespace/DistinctTreeProver.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Statespace/DistinctTreeProver.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -645,9 +645,9 @@
 val const_decls = map (fn i => Syntax.no_syn 
                                  ("const" ^ string_of_int i,Type ("nat",[]))) nums
 
-val consts = sort TermOrd.fast_term_ord 
+val consts = sort Term_Ord.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums)
-val consts' = sort TermOrd.fast_term_ord 
+val consts' = sort Term_Ord.fast_term_ord 
               (map (fn i => Const ("DistinctTreeProver.const"^string_of_int i,Type ("nat",[]))) nums')
 
 val t = DistinctTreeProver.mk_tree I (Type ("nat",[])) consts
--- a/src/HOL/Statespace/StateFun.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Statespace/StateFun.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -22,7 +22,7 @@
 better compositionality, especially if you think of nested state
 spaces.  *} 
 
-constdefs K_statefun:: "'a \<Rightarrow> 'b \<Rightarrow> 'a" "K_statefun c x \<equiv> c"
+definition K_statefun :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where "K_statefun c x \<equiv> c"
 
 lemma K_statefun_apply [simp]: "K_statefun c x = c"
   by (simp add: K_statefun_def)
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -79,7 +79,7 @@
   | bin_find_tree order e t = raise TERM ("find_tree: input not a tree",[t])
 
 fun find_tree e t =
-  (case bin_find_tree TermOrd.fast_term_ord e t of
+  (case bin_find_tree Term_Ord.fast_term_ord e t of
      NONE => lin_find_tree e t
    | x => x);
 
--- a/src/HOL/Statespace/state_fun.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Statespace/state_fun.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -76,7 +76,7 @@
 
 val string_eq_simp_tac = simp_tac (HOL_basic_ss 
   addsimps (@{thms list.inject} @ @{thms char.inject}
-    @ @{thms list.distinct} @ @{thms char.distinct} @ simp_thms)
+    @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms})
   addsimprocs [lazy_conj_simproc]
   addcongs [@{thm block_conj_cong}])
 
@@ -84,7 +84,7 @@
 
 val lookup_ss = (HOL_basic_ss 
   addsimps (@{thms list.inject} @ @{thms char.inject}
-    @ @{thms list.distinct} @ @{thms char.distinct} @ simp_thms
+    @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms}
     @ [@{thm StateFun.lookup_update_id_same}, @{thm StateFun.id_id_cancel},
       @{thm StateFun.lookup_update_same}, @{thm StateFun.lookup_update_other}])
   addsimprocs [lazy_conj_simproc]
--- a/src/HOL/TLA/Action.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/TLA/Action.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -33,7 +33,7 @@
 
 syntax
   (* Syntax for writing action expressions in arbitrary contexts *)
-  "ACT"         :: "lift => 'a"                      ("(ACT _)")
+  "_ACT"        :: "lift => 'a"                      ("(ACT _)")
 
   "_before"     :: "lift => lift"                    ("($_)"  [100] 99)
   "_after"      :: "lift => lift"                    ("(_$)"  [100] 99)
--- a/src/HOL/TLA/Intensional.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/TLA/Intensional.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -51,7 +51,7 @@
   "_holdsAt"    :: "['a, lift] => bool"                  ("(_ |= _)" [100,10] 10)
 
   (* Syntax for lifted expressions outside the scope of |- or |= *)
-  "LIFT"        :: "lift => 'a"                          ("LIFT _")
+  "_LIFT"       :: "lift => 'a"                          ("LIFT _")
 
   (* generic syntax for lifted constants and functions *)
   "_const"      :: "'a => lift"                          ("(#_)" [1000] 999)
--- a/src/HOL/TLA/Stfun.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/TLA/Stfun.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -35,7 +35,7 @@
   stvars    :: "'a stfun => bool"
 
 syntax
-  "PRED"    :: "lift => 'a"                          ("PRED _")
+  "_PRED"   :: "lift => 'a"                          ("PRED _")
   "_stvars" :: "lift => bool"                        ("basevars _")
 
 translations
--- a/src/HOL/Tools/Datatype/datatype.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -75,7 +75,7 @@
     val leafTs' = get_nonrec_types descr' sorts;
     val branchTs = get_branching_types descr' sorts;
     val branchT = if null branchTs then HOLogic.unitT
-      else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
+      else Balanced_Tree.make (fn (T, U) => Type (@{type_name "+"}, [T, U])) branchTs;
     val arities = remove (op =) 0 (get_arities descr');
     val unneeded_vars =
       subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars);
@@ -83,7 +83,7 @@
     val recTs = get_rec_types descr' sorts;
     val (newTs, oldTs) = chop (length (hd descr)) recTs;
     val sumT = if null leafTs then HOLogic.unitT
-      else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
+      else Balanced_Tree.make (fn (T, U) => Type (@{type_name "+"}, [T, U])) leafTs;
     val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
     val UnivT = HOLogic.mk_setT Univ_elT;
     val UnivT' = Univ_elT --> HOLogic.boolT;
@@ -104,9 +104,9 @@
               val Type (_, [T1, T2]) = T
           in
             if i <= n2 then
-              Const (@{const_name "Sum_Type.Inl"}, T1 --> T) $ (mk_inj' T1 n2 i)
+              Const (@{const_name Inl}, T1 --> T) $ (mk_inj' T1 n2 i)
             else
-              Const (@{const_name "Sum_Type.Inr"}, T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
+              Const (@{const_name Inr}, T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
           end
       in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs)
       end;
@@ -465,7 +465,7 @@
           (if i < length newTs then HOLogic.true_const
            else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
              Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
-               Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
+               Const (iso_name, isoT) $ Const (@{const_abbrev UNIV}, HOLogic.mk_setT T)))
       end;
 
     val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
@@ -481,7 +481,7 @@
         (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY
            [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
             REPEAT (rtac TrueI 1),
-            rewrite_goals_tac (mk_meta_eq choice_eq ::
+            rewrite_goals_tac (mk_meta_eq @{thm choice_eq} ::
               symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
             rewrite_goals_tac (map symmetric range_eqs),
             REPEAT (EVERY
@@ -731,7 +731,7 @@
   in (names, specs) end;
 
 val parse_datatype_decl =
-  (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
+  (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_mixfix --
     (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix)));
 
 val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls;
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -53,7 +53,7 @@
     fun prove_casedist_thm (i, (T, t)) =
       let
         val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
-          Abs ("z", T', Const ("True", T''))) induct_Ps;
+          Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
         val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
           Var (("P", 0), HOLogic.boolT))
         val insts = take i dummyPs @ (P::(drop (i + 1) dummyPs));
@@ -200,7 +200,7 @@
     val rec_unique_thms =
       let
         val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
-          Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
+          Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
             absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
               (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
         val cert = cterm_of thy1
@@ -210,13 +210,13 @@
           (map cert insts)) induct;
         val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
            (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
-              THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs));
+              THEN rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs));
 
       in split_conj_thm (Skip_Proof.prove_global thy1 [] []
         (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
       end;
 
-    val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
+    val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
 
     (* define primrec combinators *)
 
@@ -236,7 +236,7 @@
           (reccomb_names ~~ recTs ~~ rec_result_Ts))
       |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
           (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
-           Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+           Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
              set $ Free ("x", T) $ Free ("y", T'))))))
                (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
       ||> Sign.parent_path
@@ -250,7 +250,7 @@
     val rec_thms = map (fn t => Skip_Proof.prove_global thy2 [] [] t
       (fn _ => EVERY
         [rewrite_goals_tac reccomb_defs,
-         rtac the1_equality 1,
+         rtac @{thm the1_equality} 1,
          resolve_tac rec_unique_thms 1,
          resolve_tac rec_intrs 1,
          REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
@@ -416,7 +416,7 @@
     fun prove_case_cong ((t, nchotomy), case_rewrites) =
       let
         val (Const ("==>", _) $ tm $ _) = t;
-        val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
+        val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ Ma)) = tm;
         val cert = cterm_of thy;
         val nchotomy' = nchotomy RS spec;
         val [v] = Term.add_vars (concl_of nchotomy') [];
--- a/src/HOL/Tools/Datatype/datatype_aux.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -120,8 +120,8 @@
 fun split_conj_thm th =
   ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
 
-val mk_conj = foldr1 (HOLogic.mk_binop "op &");
-val mk_disj = foldr1 (HOLogic.mk_binop "op |");
+val mk_conj = foldr1 (HOLogic.mk_binop @{const_name "op &"});
+val mk_disj = foldr1 (HOLogic.mk_binop @{const_name "op |"});
 
 fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
 
--- a/src/HOL/Tools/Datatype/datatype_case.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_case.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -76,7 +76,7 @@
 
 fun mk_fun_constrain tT t =
   Syntax.const @{syntax_const "_constrain"} $ t $
-    (Syntax.free "fun" $ tT $ Syntax.free "dummy");    (* FIXME @{type_syntax} (!?) *)
+    (Syntax.const @{type_syntax fun} $ tT $ Syntax.const @{type_syntax dummy});
 
 
 (*---------------------------------------------------------------------------
@@ -381,7 +381,7 @@
         fun count_cases (_, _, true) = I
           | count_cases (c, (_, body), false) =
               AList.map_default op aconv (body, []) (cons c);
-        val is_undefined = name_of #> equal (SOME @{const_syntax undefined});
+        val is_undefined = name_of #> equal (SOME @{const_name undefined});
         fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
       in case ty_info tab cname of
           SOME {constructors, case_name} =>
@@ -399,7 +399,7 @@
                   (fold_rev count_cases cases []);
                 val R = type_of t;
                 val dummy =
-                  if d then Const (@{const_syntax dummy_pattern}, R)
+                  if d then Const (@{const_name dummy_pattern}, R)
                   else Free (Name.variant used "x", R)
               in
                 SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
--- a/src/HOL/Tools/Datatype/datatype_data.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_data.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -236,7 +236,7 @@
 
 (** document antiquotation **)
 
-val _ = ThyOutput.antiquotation "datatype" Args.tyname
+val _ = ThyOutput.antiquotation "datatype" (Args.type_name true)
   (fn {source = src, context = ctxt, ...} => fn dtco =>
     let
       val thy = ProofContext.theory_of ctxt;
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -70,7 +70,7 @@
           val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
         in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
           (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
-           foldr1 (HOLogic.mk_binop "op &")
+           foldr1 (HOLogic.mk_binop @{const_name "op &"})
              (map HOLogic.mk_eq (frees ~~ frees')))))
         end;
   in
@@ -149,7 +149,7 @@
     val prems = maps (fn ((i, (_, _, constrs)), T) =>
       map (make_ind_prem i T) constrs) (descr' ~~ recTs);
     val tnames = make_tnames recTs;
-    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
       (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
         (descr' ~~ recTs ~~ tnames)))
 
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -16,10 +16,11 @@
 
 open Datatype_Aux;
 
-fun subsets i j = if i <= j then
-       let val is = subsets (i+1) j
-       in map (fn ks => i::ks) is @ is end
-     else [[]];
+fun subsets i j =
+  if i <= j then
+    let val is = subsets (i+1) j
+    in map (fn ks => i::ks) is @ is end
+  else [[]];
 
 fun forall_intr_prf (t, prf) =
   let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
@@ -102,7 +103,7 @@
         if i mem is then SOME
           (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
         else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
-    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+    val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
       (map (fn ((((i, _), T), U), tname) =>
         make_pred i U T (mk_proj i is r) (Free (tname, T)))
           (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
@@ -129,8 +130,8 @@
 
     val ivs = rev (Term.add_vars (Logic.varify (Datatype_Prop.make_ind [descr] sorts)) []);
     val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
-    val ivs1 = map Var (filter_out (fn (_, T) =>
-      tname_of (body_type T) mem ["set", "bool"]) ivs);
+    val ivs1 = map Var (filter_out (fn (_, T) =>  (* FIXME set (!??) *)
+      tname_of (body_type T) mem [@{type_abbrev set}, @{type_name bool}]) ivs);
     val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs;
 
     val prf =
--- a/src/HOL/Tools/Function/context_tree.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Function/context_tree.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -64,7 +64,7 @@
 val cong_del = Thm.declaration_attribute (map_function_congs o Thm.del_thm o safe_mk_meta_eq);
 
 
-type depgraph = int IntGraph.T
+type depgraph = int Int_Graph.T
 
 datatype ctx_tree =
   Leaf of term
@@ -90,11 +90,11 @@
   let
     val num_branches = map_index (apsnd branch_vars) (prems_of crule)
   in
-    IntGraph.empty
-    |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches
+    Int_Graph.empty
+    |> fold (fn (i,_)=> Int_Graph.new_node (i,i)) num_branches
     |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) =>
          if i = j orelse null (inter (op =) c1 t2)
-         then I else IntGraph.add_edge_acyclic (i,j))
+         then I else Int_Graph.add_edge_acyclic (i,j))
        num_branches num_branches
     end
 
@@ -204,13 +204,13 @@
         SOME _ => (T, x)
       | NONE =>
         let
-          val (T', x') = fold fill_table (IntGraph.imm_succs G i) (T, x)
+          val (T', x') = fold fill_table (Int_Graph.imm_succs G i) (T, x)
           val (v, x'') = f (the o Inttab.lookup T') i x'
         in
           (Inttab.update (i, v) T', x'')
         end
 
-    val (T, x) = fold fill_table (IntGraph.keys G) (Inttab.empty, x)
+    val (T, x) = fold fill_table (Int_Graph.keys G) (Inttab.empty, x)
   in
     (Inttab.fold (cons o snd) T [], x)
   end
@@ -225,7 +225,7 @@
         fun sub_step lu i x =
           let
             val (ctx', subtree) = nth branches i
-            val used = fold_rev (append o lu) (IntGraph.imm_succs deps i) u
+            val used = fold_rev (append o lu) (Int_Graph.imm_succs deps i) u
             val (subs, x') = traverse_help (compose ctx ctx') subtree used x
             val exported_subs = map (apfst (compose ctx')) subs (* FIXME: Right order of composition? *)
           in
@@ -263,7 +263,7 @@
           fun sub_step lu i x =
             let
               val ((fixes, assumes), st) = nth branches i
-              val used = map lu (IntGraph.imm_succs deps i)
+              val used = map lu (Int_Graph.imm_succs deps i)
                 |> map (fn u_eq => (u_eq RS sym) RS eq_reflection)
                 |> filter_out Thm.is_reflexive
 
--- a/src/HOL/Tools/Function/function.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Function/function.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -146,7 +146,7 @@
         let
           val totality = Thm.close_derivation totality
           val remove_domain_condition =
-            full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
+            full_simplify (HOL_basic_ss addsimps [totality, @{thm True_implies_equals}])
           val tsimps = map remove_domain_condition psimps
           val tinduct = map remove_domain_condition pinducts
 
--- a/src/HOL/Tools/Function/function_lib.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Function/function_lib.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -168,7 +168,7 @@
 
 (* instance for unions *)
 val regroup_union_conv =
-  regroup_conv @{const_name Set.empty} @{const_name Lattices.sup}
+  regroup_conv @{const_abbrev Set.empty} @{const_name Lattices.sup}
     (map (fn t => t RS eq_reflection)
       (@{thms Un_ac} @ @{thms Un_empty_right} @ @{thms Un_empty_left}))
 
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -24,7 +24,7 @@
   let
     val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))
     val mlexT = (domT --> HOLogic.natT) --> relT --> relT
-    fun mk_ms [] = Const (@{const_name Set.empty}, relT)
+    fun mk_ms [] = Const (@{const_abbrev Set.empty}, relT)
       | mk_ms (f::fs) =
         Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs
   in
--- a/src/HOL/Tools/Function/termination.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Function/termination.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -51,9 +51,10 @@
 
 open Function_Lib
 
-val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
+val term2_ord = prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord
 structure Term2tab = Table(type key = term * term val ord = term2_ord);
-structure Term3tab = Table(type key = term * (term * term) val ord = prod_ord TermOrd.fast_term_ord term2_ord);
+structure Term3tab =
+  Table(type key = term * (term * term) val ord = prod_ord Term_Ord.fast_term_ord term2_ord);
 
 (** Analyzing binary trees **)
 
@@ -286,7 +287,7 @@
 
     val relation =
       if null ineqs
-      then Const (@{const_name Set.empty}, fastype_of rel)
+      then Const (@{const_abbrev Set.empty}, fastype_of rel)
       else map mk_compr ineqs
         |> foldr1 (HOLogic.mk_binop @{const_name Lattices.sup})
 
@@ -314,14 +315,11 @@
 
 (*** DEPENDENCY GRAPHS ***)
 
-structure TermGraph = Graph(type key = term val ord = TermOrd.fast_term_ord);
-
-
 fun prove_chain thy chain_tac c1 c2 =
   let
     val goal =
       HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.rel_comp} (c1, c2),
-        Const (@{const_name Set.empty}, fastype_of c1))
+        Const (@{const_abbrev Set.empty}, fastype_of c1))
       |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
   in
     case Function_Lib.try_proof (cterm_of thy goal) chain_tac of
@@ -342,11 +340,11 @@
 
 
 fun mk_dgraph D cs =
-  TermGraph.empty
-  |> fold (fn c => TermGraph.new_node (c,())) cs
+  Term_Graph.empty
+  |> fold (fn c => Term_Graph.new_node (c, ())) cs
   |> fold_product (fn c1 => fn c2 =>
      if is_none (get_chain D c1 c2 |> the_default NONE)
-     then TermGraph.add_edge (c1, c2) else I)
+     then Term_Graph.add_edge (c1, c2) else I)
      cs cs
 
 fun ucomp_empty_tac T =
@@ -373,7 +371,7 @@
 fun decompose_tac' cont err_cont D = CALLS (fn (cs, i) =>
   let
     val G = mk_dgraph D cs
-    val sccs = TermGraph.strong_conn G
+    val sccs = Term_Graph.strong_conn G
 
     fun split [SCC] i = (solve_trivial_tac D i ORELSE cont D i)
       | split (SCC::rest) i =
--- a/src/HOL/Tools/Groebner_Basis/groebner.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Groebner_Basis/groebner.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -447,7 +447,8 @@
 val initial_conv =
     Simplifier.rewrite
      (HOL_basic_ss addsimps nnf_simps
-     addsimps [not_all, not_ex] addsimps map (fn th => th RS sym) (ex_simps @ all_simps));
+       addsimps [not_all, not_ex]
+       addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps}));
 
 val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
 
@@ -689,7 +690,7 @@
   val cjs = conjs tm
   val  rawvars = fold_rev (fn eq => fn a =>
                                        grobvars (dest_arg1 eq) (grobvars (dest_arg eq) a)) cjs []
-  val vars = sort (fn (x, y) => TermOrd.term_ord(term_of x,term_of y))
+  val vars = sort (fn (x, y) => Term_Ord.term_ord(term_of x,term_of y))
                   (distinct (op aconvc) rawvars)
  in (vars,map (grobify_equation vars) cjs)
  end;
@@ -820,7 +821,7 @@
    in make_simproc {lhss = [Thm.lhs_of idl_sub], 
                 name = "poly_eq_simproc", proc = proc, identifier = []}
    end;
-  val poly_eq_ss = HOL_basic_ss addsimps simp_thms 
+  val poly_eq_ss = HOL_basic_ss addsimps @{thms simp_thms}
                         addsimprocs [poly_eq_simproc]
 
  local
@@ -898,7 +899,7 @@
   val vars = filter (fn v => free_in v eq) evs
   val (qs,p) = isolate_monomials vars eq
   val rs = ideal (qs @ ps) p 
-              (fn (s,t) => TermOrd.term_ord (term_of s, term_of t))
+              (fn (s,t) => Term_Ord.term_ord (term_of s, term_of t))
  in (eq, take (length qs) rs ~~ vars)
  end;
  fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p));
--- a/src/HOL/Tools/Groebner_Basis/normalizer.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -60,7 +60,7 @@
   (Simplifier.rewrite 
     (HOL_basic_ss 
        addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps}
-             @ [if_False, if_True, @{thm Nat.add_0}, @{thm add_Suc},
+             @ [@{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc},
                  @{thm add_number_of_left}, @{thm Suc_eq_plus1}]
              @ map (fn th => th RS sym) @{thms numerals}));
 
@@ -632,11 +632,11 @@
  end
 end;
 
-val nat_arith = @{thms "nat_arith"};
-val nat_exp_ss = HOL_basic_ss addsimps (@{thms nat_number} @ nat_arith @ @{thms arith_simps} @ @{thms rel_simps})
-                              addsimps [Let_def, if_False, if_True, @{thm Nat.add_0}, @{thm add_Suc}];
+val nat_exp_ss =
+  HOL_basic_ss addsimps (@{thms nat_number} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps})
+    addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}];
 
-fun simple_cterm_ord t u = TermOrd.term_ord (term_of t, term_of u) = LESS;
+fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
 
 fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, 
                                      {conv, dest_const, mk_const, is_const}) ord =
--- a/src/HOL/Tools/Nitpick/HISTORY	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/HISTORY	Tue Mar 02 04:35:44 2010 -0800
@@ -4,7 +4,11 @@
   * Added "std" option and implemented support for nonstandard models
   * Added support for quotient types
   * Added support for local definitions
-  * Optimized "Multiset.multiset"
+  * Disabled "fast_descrs" option by default
+  * Optimized "Multiset.multiset" and "FinFun.finfun"
+  * Improved efficiency of "destroy_constrs" optimization
+  * Improved precision of infinite datatypes whose constructors don't appear
+    in the formula to falsify based on a monotonicity analysis
   * Fixed soundness bugs related to "destroy_constrs" optimization and record
     getters
   * Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and "SAT4JLight" to
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -131,9 +131,7 @@
   nonsel_names: nut list,
   rel_table: nut NameTable.table,
   unsound: bool,
-  scope: scope,
-  core: KK.formula,
-  defs: KK.formula list}
+  scope: scope}
 
 type rich_problem = KK.problem * problem_extension
 
@@ -195,7 +193,8 @@
     val timer = Timer.startRealTimer ()
     val thy = Proof.theory_of state
     val ctxt = Proof.context_of state
-(* FIXME: reintroduce code before new release
+(* FIXME: reintroduce code before new release:
+
     val nitpick_thy = ThyInfo.get_theory "Nitpick"
     val _ = Theory.subthy (nitpick_thy, thy) orelse
             error "You must import the theory \"Nitpick\" to use Nitpick"
@@ -289,8 +288,8 @@
     val frees = Term.add_frees assms_t []
     val _ = null (Term.add_tvars assms_t []) orelse
             raise NOT_SUPPORTED "schematic type variables"
-    val (((def_ts, nondef_ts), (got_all_mono_user_axioms, no_poly_user_axioms)),
-         core_t, binarize) = preprocess_term hol_ctxt assms_t
+    val (nondef_ts, def_ts, got_all_mono_user_axioms, no_poly_user_axioms,
+         binarize) = preprocess_term hol_ctxt assms_t
     val got_all_user_axioms =
       got_all_mono_user_axioms andalso no_poly_user_axioms
 
@@ -307,54 +306,66 @@
                    \unroll it.")))
     val _ = if verbose then List.app print_wf (!wf_cache) else ()
     val _ =
-      pprint_d (fn () =>
-          Pretty.chunks
-              (pretties_for_formulas ctxt "Preprocessed formula" [core_t] @
-               pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @
-               pretties_for_formulas ctxt "Relevant nondefinitional axiom"
-                                     nondef_ts))
-    val _ = List.app (ignore o Term.type_of) (core_t :: def_ts @ nondef_ts)
+      pprint_d (fn () => Pretty.chunks
+          (pretties_for_formulas ctxt "Preprocessed formula" [hd nondef_ts] @
+           pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @
+           pretties_for_formulas ctxt "Relevant nondefinitional axiom"
+                                 (tl nondef_ts)))
+    val _ = List.app (ignore o Term.type_of) (nondef_ts @ def_ts)
             handle TYPE (_, Ts, ts) =>
                    raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts)
 
-    val core_u = nut_from_term hol_ctxt Eq core_t
+    val nondef_us = map (nut_from_term hol_ctxt Eq) nondef_ts
     val def_us = map (nut_from_term hol_ctxt DefEq) def_ts
-    val nondef_us = map (nut_from_term hol_ctxt Eq) nondef_ts
     val (free_names, const_names) =
-      fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], [])
+      fold add_free_and_const_names (nondef_us @ def_us) ([], [])
     val (sel_names, nonsel_names) =
       List.partition (is_sel o nickname_of) const_names
     val genuine_means_genuine = got_all_user_axioms andalso none_true wfs
     val standard = forall snd stds
 (*
-    val _ = List.app (priority o string_for_nut ctxt)
-                     (core_u :: def_us @ nondef_us)
+    val _ = List.app (priority o string_for_nut ctxt) (nondef_us @ def_us)
 *)
 
     val unique_scope = forall (curry (op =) 1 o length o snd) cards_assigns
+    (* typ list -> string -> string *)
+    fun monotonicity_message Ts extra =
+      let val ss = map (quote o string_for_type ctxt) Ts in
+        "The type" ^ plural_s_for_list ss ^ " " ^
+        space_implode " " (serial_commas "and" ss) ^ " " ^
+        (if none_true monos then
+           "passed the monotonicity test"
+         else
+           (if length ss = 1 then "is" else "are") ^ " considered monotonic") ^
+        ". " ^ extra
+      end
     (* typ -> bool *)
     fun is_type_always_monotonic T =
       (is_datatype thy stds T andalso not (is_quot_type thy T) andalso
        (not (is_pure_typedef thy T) orelse is_univ_typedef thy T)) orelse
       is_number_type thy T orelse is_bit_type T
+    fun is_type_actually_monotonic T =
+      formulas_monotonic hol_ctxt binarize T (nondef_ts, def_ts)
     fun is_type_monotonic T =
       unique_scope orelse
       case triple_lookup (type_match thy) monos T of
         SOME (SOME b) => b
-      | _ => is_type_always_monotonic T orelse
-             formulas_monotonic hol_ctxt binarize T Plus def_ts nondef_ts core_t
-    fun is_deep_datatype T =
-      is_datatype thy stds T andalso
-      (not standard orelse T = nat_T orelse is_word_type T orelse
-       exists (curry (op =) T o domain_type o type_of) sel_names)
-    val all_Ts = ground_types_in_terms hol_ctxt binarize
-                                       (core_t :: def_ts @ nondef_ts)
-                 |> sort TermOrd.typ_ord
+      | _ => is_type_always_monotonic T orelse is_type_actually_monotonic T
+    fun is_type_finitizable T =
+      case triple_lookup (type_match thy) monos T of
+        SOME (SOME false) => false
+      | _ => is_type_actually_monotonic T
+    fun is_datatype_deep T =
+      not standard orelse T = nat_T orelse is_word_type T orelse
+      exists (curry (op =) T o domain_type o type_of) sel_names
+    val all_Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)
+                 |> sort Term_Ord.typ_ord
     val _ = if verbose andalso binary_ints = SOME true andalso
                exists (member (op =) [nat_T, int_T]) all_Ts then
               print_v (K "The option \"binary_ints\" will be ignored because \
                          \of the presence of rationals, reals, \"Suc\", \
-                         \\"gcd\", or \"lcm\" in the problem.")
+                         \\"gcd\", or \"lcm\" in the problem or because of the \
+                         \\"non_std\" option.")
             else
               ()
     val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts
@@ -363,23 +374,21 @@
         case filter_out is_type_always_monotonic mono_Ts of
           [] => ()
         | interesting_mono_Ts =>
-          print_v (fn () =>
-                      let
-                        val ss = map (quote o string_for_type ctxt)
-                                     interesting_mono_Ts
-                      in
-                        "The type" ^ plural_s_for_list ss ^ " " ^
-                        space_implode " " (serial_commas "and" ss) ^ " " ^
-                        (if none_true monos then
-                           "passed the monotonicity test"
-                         else
-                           (if length ss = 1 then "is" else "are") ^
-                           " considered monotonic") ^
-                        ". Nitpick might be able to skip some scopes."
-                      end)
+          print_v (K (monotonicity_message interesting_mono_Ts
+                         "Nitpick might be able to skip some scopes."))
       else
         ()
-    val deep_dataTs = filter is_deep_datatype all_Ts
+    val (deep_dataTs, shallow_dataTs) =
+      all_Ts |> filter (is_datatype thy stds) |> List.partition is_datatype_deep
+    val finitizable_dataTs =
+      shallow_dataTs |> filter_out (is_finite_type hol_ctxt)
+                     |> filter is_type_finitizable
+    val _ =
+      if verbose andalso not (null finitizable_dataTs) then
+        print_v (K (monotonicity_message finitizable_dataTs
+                        "Nitpick can use a more precise finite encoding."))
+      else
+        ()
     (* This detection code is an ugly hack. Fortunately, it is used only to
        provide a hint to the user. *)
     (* string * (Rule_Cases.T * bool) -> bool *)
@@ -446,7 +455,7 @@
                                        string_of_int j0))
                          (Typtab.dest ofs)
 *)
-        val all_exact = forall (is_exact_type datatypes) all_Ts
+        val all_exact = forall (is_exact_type datatypes true) all_Ts
         (* nut list -> rep NameTable.table -> nut list * rep NameTable.table *)
         val repify_consts = choose_reps_for_consts scope all_exact
         val main_j0 = offset_of_type ofs bool_T
@@ -467,7 +476,6 @@
                          (univ_card nat_card int_card main_j0 [] KK.True)
         val _ = check_arity min_univ_card min_highest_arity
 
-        val core_u = choose_reps_in_nut scope unsound rep_table false core_u
         val def_us = map (choose_reps_in_nut scope unsound rep_table true)
                          def_us
         val nondef_us = map (choose_reps_in_nut scope unsound rep_table false)
@@ -475,7 +483,7 @@
 (*
         val _ = List.app (priority o string_for_nut ctxt)
                          (free_names @ sel_names @ nonsel_names @
-                          core_u :: def_us @ nondef_us)
+                          nondef_us @ def_us)
 *)
         val (free_rels, pool, rel_table) =
           rename_free_vars free_names initial_pool NameTable.empty
@@ -483,13 +491,11 @@
           rename_free_vars sel_names pool rel_table
         val (other_rels, pool, rel_table) =
           rename_free_vars nonsel_names pool rel_table
-        val core_u = rename_vars_in_nut pool rel_table core_u
+        val nondef_us = map (rename_vars_in_nut pool rel_table) nondef_us
         val def_us = map (rename_vars_in_nut pool rel_table) def_us
-        val nondef_us = map (rename_vars_in_nut pool rel_table) nondef_us
-        val core_f = kodkod_formula_from_nut ofs kk core_u
+        val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us
         val def_fs = map (kodkod_formula_from_nut ofs kk) def_us
-        val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us
-        val formula = fold (fold s_and) [def_fs, nondef_fs] core_f
+        val formula = fold (fold s_and) [def_fs, nondef_fs] KK.True
         val comment = (if unsound then "unsound" else "sound") ^ "\n" ^
                       PrintMode.setmp [] multiline_string_for_scope scope
         val kodkod_sat_solver =
@@ -537,8 +543,7 @@
                expr_assigns = [], formula = formula},
               {free_names = free_names, sel_names = sel_names,
                nonsel_names = nonsel_names, rel_table = rel_table,
-               unsound = unsound, scope = scope, core = core_f,
-               defs = nondef_fs @ def_fs @ declarative_axioms})
+               unsound = unsound, scope = scope})
       end
       handle TOO_LARGE (loc, msg) =>
              if loc = "Nitpick_Kodkod.check_arity" andalso
@@ -834,8 +839,8 @@
                         sound_problems then
                 print_m (fn () =>
                     "Warning: The conjecture either trivially holds for the \
-                    \given scopes or (more likely) lies outside Nitpick's \
-                    \supported fragment." ^
+                    \given scopes or lies outside Nitpick's supported \
+                    \fragment." ^
                     (if exists (not o KK.is_problem_trivially_false o fst)
                                unsound_problems then
                        " Only potential counterexamples may be found."
@@ -907,6 +912,7 @@
     val (skipped, the_scopes) =
       all_scopes hol_ctxt binarize sym_break cards_assigns maxes_assigns
                  iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
+                 finitizable_dataTs
     val _ = if skipped > 0 then
               print_m (fn () => "Too many scopes. Skipping " ^
                                 string_of_int skipped ^ " scope" ^
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -151,9 +151,10 @@
   val card_of_type : (typ * int) list -> typ -> int
   val bounded_card_of_type : int -> int -> (typ * int) list -> typ -> int
   val bounded_exact_card_of_type :
-    hol_context -> int -> int -> (typ * int) list -> typ -> int
+    hol_context -> typ list -> int -> int -> (typ * int) list -> typ -> int
   val is_finite_type : hol_context -> typ -> bool
   val special_bounds : term list -> (indexname * typ) list
+  val abs_var : indexname * typ -> term -> term
   val is_funky_typedef : theory -> typ -> bool
   val all_axioms_of :
     theory -> (term * term) list -> term list * term list * term list
@@ -631,8 +632,13 @@
        in
          case t_opt of
            SOME (Const (@{const_name top}, _)) => true
+           (* "Multiset.multiset" *)
          | SOME (Const (@{const_name Collect}, _)
                  $ Abs (_, _, Const (@{const_name finite}, _) $ _)) => true
+           (* "FinFun.finfun" *)
+         | SOME (Const (@{const_name Collect}, _) $ Abs (_, _,
+                     Const (@{const_name Ex}, _) $ Abs (_, _,
+                         Const (@{const_name finite}, _) $ _))) => true
          | _ => false
        end
      | NONE => false)
@@ -1047,13 +1053,16 @@
                     card_of_type assigns T
                     handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
                            default_card)
-(* hol_context -> int -> (typ * int) list -> typ -> int *)
-fun bounded_exact_card_of_type hol_ctxt max default_card assigns T =
+(* hol_context -> typ list -> int -> (typ * int) list -> typ -> int *)
+fun bounded_exact_card_of_type hol_ctxt finitizable_dataTs max default_card
+                               assigns T =
   let
     (* typ list -> typ -> int *)
     fun aux avoid T =
       (if member (op =) avoid T then
          0
+       else if member (op =) finitizable_dataTs T then
+         raise SAME ()
        else case T of
          Type ("fun", [T1, T2]) =>
          let
@@ -1096,8 +1105,8 @@
   in Int.min (max, aux [] T) end
 
 (* hol_context -> typ -> bool *)
-fun is_finite_type hol_ctxt =
-  not_equal 0 o bounded_exact_card_of_type hol_ctxt 1 2 []
+fun is_finite_type hol_ctxt T =
+  bounded_exact_card_of_type hol_ctxt [] 1 2 [] T <> 0
 
 (* term -> bool *)
 fun is_ground_term (t1 $ t2) = is_ground_term t1 andalso is_ground_term t2
@@ -1113,7 +1122,7 @@
 
 (* term list -> (indexname * typ) list *)
 fun special_bounds ts =
-  fold Term.add_vars ts [] |> sort (TermOrd.fast_indexname_ord o pairself fst)
+  fold Term.add_vars ts [] |> sort (Term_Ord.fast_indexname_ord o pairself fst)
 
 (* indexname * typ -> term -> term *)
 fun abs_var ((s, j), T) body = Abs (s, T, abstract_over (Var ((s, j), T), body))
@@ -1442,9 +1451,10 @@
            @{const Not} $ (is_unknown_t $ normal_y),
            equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
            Logic.mk_equals (normal_x, normal_y)),
-     @{const "==>"}
-         $ (HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)))
-         $ (HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
+     Logic.list_implies
+         ([HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)),
+           HOLogic.mk_Trueprop (@{const Not} $ HOLogic.mk_eq (normal_x, x_var))],
+          HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
   end
 
 (* theory -> (typ option * bool) list -> int * styp -> term *)
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -55,7 +55,7 @@
 type nfa_entry = typ * nfa_transition list
 type nfa_table = nfa_entry list
 
-structure NfaGraph = Graph(type key = typ val ord = TermOrd.typ_ord)
+structure NfaGraph = Typ_Graph
 
 (* int -> KK.int_expr list *)
 fun flip_nums n = index_seq 1 n @ [0] |> map KK.Num
@@ -255,10 +255,10 @@
      ("int_divide", tabulate_int_op2 debug univ_card (int_card, j0) isa_div)
    else if x = nat_less_rel then
      ("nat_less", tabulate_nat_op2 debug univ_card (nat_card, j0)
-                                   (int_for_bool o op <))
+                                   (int_from_bool o op <))
    else if x = int_less_rel then
      ("int_less", tabulate_int_op2 debug univ_card (int_card, j0)
-                                   (int_for_bool o op <))
+                                   (int_from_bool o op <))
    else if x = gcd_rel then
      ("gcd", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_gcd)
    else if x = lcm_rel then
@@ -444,7 +444,7 @@
   single_rel_rel_let kk (fn r1 => basic_rel_rel_let 1 kk (f r1) r2) r1
 (* kodkod_constrs -> (KK.rel_expr -> KK.rel_expr -> KK.rel_expr -> KK.rel_expr)
    -> KK.rel_expr -> KK.rel_expr -> KK.rel_expr -> KK.rel_expr *)
-fun tripl_rel_rel_let kk f r1 r2 r3 =
+fun triple_rel_rel_let kk f r1 r2 r3 =
   double_rel_rel_let kk
       (fn r1 => fn r2 => basic_rel_rel_let 2 kk (f r1 r2) r3) r1 r2
 
@@ -1615,7 +1615,7 @@
         kk_rel_let [to_expr_assign u1 u2] (to_rep R u3)
       | Op3 (If, _, R, u1, u2, u3) =>
         if is_opt_rep (rep_of u1) then
-          tripl_rel_rel_let kk
+          triple_rel_rel_let kk
               (fn r1 => fn r2 => fn r3 =>
                   let val empty_r = empty_rel_for_rep R in
                     fold1 kk_union
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -103,7 +103,7 @@
              (case nice_term_ord (t11, t21) of
                 EQUAL => nice_term_ord (t12, t22)
               | ord => ord)
-           | _ => TermOrd.fast_term_ord tp
+           | _ => Term_Ord.fast_term_ord tp
 
 (* nut NameTable.table -> KK.raw_bound list -> nut -> int list list *)
 fun tuple_list_for_name rel_table bounds name =
@@ -281,38 +281,45 @@
              js 0
       end
     (* bool -> typ -> typ -> (term * term) list -> term *)
-    fun make_set maybe_opt T1 T2 =
+    fun make_set maybe_opt T1 T2 tps =
       let
-        val empty_const = Const (@{const_name Set.empty}, T1 --> T2)
+        val empty_const = Const (@{const_abbrev Set.empty}, T1 --> T2)
         val insert_const = Const (@{const_name insert},
                                   T1 --> (T1 --> T2) --> T1 --> T2)
         (* (term * term) list -> term *)
         fun aux [] =
-            if maybe_opt andalso not (is_complete_type datatypes T1) then
+            if maybe_opt andalso not (is_complete_type datatypes false T1) then
               insert_const $ Const (unrep, T1) $ empty_const
             else
               empty_const
           | aux ((t1, t2) :: zs) =
-            aux zs |> t2 <> @{const False}
-                      ? curry (op $) (insert_const
-                                      $ (t1 |> t2 <> @{const True}
-                                               ? curry (op $)
-                                                       (Const (maybe_name,
-                                                               T1 --> T1))))
-      in aux end
+            aux zs
+            |> t2 <> @{const False}
+               ? curry (op $)
+                       (insert_const
+                        $ (t1 |> t2 <> @{const True}
+                                 ? curry (op $)
+                                         (Const (maybe_name, T1 --> T1))))
+      in
+        if forall (fn (_, t) => t <> @{const True} andalso t <> @{const False})
+                  tps then
+          Const (unknown, T1 --> T2)
+        else
+          aux tps
+      end
     (* typ -> typ -> typ -> (term * term) list -> term *)
     fun make_map T1 T2 T2' =
       let
         val update_const = Const (@{const_name fun_upd},
                                   (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
         (* (term * term) list -> term *)
-        fun aux' [] = Const (@{const_name Map.empty}, T1 --> T2)
+        fun aux' [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
           | aux' ((t1, t2) :: ps) =
             (case t2 of
                Const (@{const_name None}, _) => aux' ps
              | _ => update_const $ aux' ps $ t1 $ t2)
         fun aux ps =
-          if not (is_complete_type datatypes T1) then
+          if not (is_complete_type datatypes false T1) then
             update_const $ aux' ps $ Const (unrep, T1)
             $ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2'))
           else
@@ -537,7 +544,7 @@
           val ts1 = map (term_for_rep seen T1 T1 R1 o single) jss1
           val ts2 =
             map (fn js => term_for_rep seen T2 T2 (Atom (2, 0))
-                                       [[int_for_bool (member (op =) jss js)]])
+                                       [[int_from_bool (member (op =) jss js)]])
                 jss1
         in make_fun false T1 T2 T' ts1 ts2 end
       | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, R2)) jss =
@@ -707,12 +714,13 @@
            Pretty.str "=",
            Pretty.enum "," "{" "}"
                (map (Syntax.pretty_term ctxt) (all_values_of_type card typ) @
-                (if complete then [] else [Pretty.str unrep]))])
+                (if fun_from_pair complete false then []
+                 else [Pretty.str unrep]))])
     (* typ -> dtype_spec list *)
     fun integer_datatype T =
       [{typ = T, card = card_of_type card_assigns T, co = false,
-        standard = true, complete = false, concrete = true, deep = true,
-        constrs = []}]
+        standard = true, complete = (false, false), concrete = (true, true),
+        deep = true, constrs = []}]
       handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
     val (codatatypes, datatypes) =
       datatypes |> filter #deep |> List.partition #co
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -2,16 +2,15 @@
     Author:     Jasmin Blanchette, TU Muenchen
     Copyright   2009, 2010
 
-Monotonicity predicate for higher-order logic.
+Monotonicity inference for higher-order logic.
 *)
 
 signature NITPICK_MONO =
 sig
-  datatype sign = Plus | Minus
   type hol_context = Nitpick_HOL.hol_context
 
   val formulas_monotonic :
-    hol_context -> bool -> typ -> sign -> term list -> term list -> term -> bool
+    hol_context -> bool -> typ -> term list * term list -> bool
 end;
 
 structure Nitpick_Mono : NITPICK_MONO =
@@ -27,22 +26,27 @@
 
 type literal = var * sign
 
-datatype ctype =
-  CAlpha |
-  CFun of ctype * sign_atom * ctype |
-  CPair of ctype * ctype |
-  CType of string * ctype list |
-  CRec of string * typ list
+datatype mtyp =
+  MAlpha |
+  MFun of mtyp * sign_atom * mtyp |
+  MPair of mtyp * mtyp |
+  MType of string * mtyp list |
+  MRec of string * typ list
 
-type cdata =
+datatype mterm =
+  MRaw of term * mtyp |
+  MAbs of string * typ * mtyp * sign_atom * mterm |
+  MApp of mterm * mterm
+
+type mdata =
   {hol_ctxt: hol_context,
    binarize: bool,
    alpha_T: typ,
    max_fresh: int Unsynchronized.ref,
-   datatype_cache: ((string * typ list) * ctype) list Unsynchronized.ref,
-   constr_cache: (styp * ctype) list Unsynchronized.ref}
+   datatype_cache: ((string * typ list) * mtyp) list Unsynchronized.ref,
+   constr_cache: (styp * mtyp) list Unsynchronized.ref}
 
-exception CTYPE of string * ctype list
+exception MTYPE of string * mtyp list
 
 (* string -> unit *)
 fun print_g (_ : string) = ()
@@ -71,55 +75,95 @@
 (* literal -> string *)
 fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn
 
-val bool_C = CType (@{type_name bool}, [])
+val bool_M = MType (@{type_name bool}, [])
+val dummy_M = MType (nitpick_prefix ^ "dummy", [])
 
-(* ctype -> bool *)
-fun is_CRec (CRec _) = true
-  | is_CRec _ = false
+(* mtyp -> bool *)
+fun is_MRec (MRec _) = true
+  | is_MRec _ = false
+(* mtyp -> mtyp * sign_atom * mtyp *)
+fun dest_MFun (MFun z) = z
+  | dest_MFun M = raise MTYPE ("Nitpick_Mono.dest_MFun", [M])
 
 val no_prec = 100
-val prec_CFun = 1
-val prec_CPair = 2
 
-(* tuple_set -> int *)
-fun precedence_of_ctype (CFun _) = prec_CFun
-  | precedence_of_ctype (CPair _) = prec_CPair
-  | precedence_of_ctype _ = no_prec
+(* mtyp -> int *)
+fun precedence_of_mtype (MFun _) = 1
+  | precedence_of_mtype (MPair _) = 2
+  | precedence_of_mtype _ = no_prec
 
-(* ctype -> string *)
-val string_for_ctype =
+(* mtyp -> string *)
+val string_for_mtype =
   let
-    (* int -> ctype -> string *)
-    fun aux outer_prec C =
+    (* int -> mtyp -> string *)
+    fun aux outer_prec M =
       let
-        val prec = precedence_of_ctype C
+        val prec = precedence_of_mtype M
         val need_parens = (prec < outer_prec)
       in
         (if need_parens then "(" else "") ^
-        (case C of
-           CAlpha => "\<alpha>"
-         | CFun (C1, a, C2) =>
-           aux (prec + 1) C1 ^ " \<Rightarrow>\<^bsup>" ^
-           string_for_sign_atom a ^ "\<^esup> " ^ aux prec C2
-         | CPair (C1, C2) => aux (prec + 1) C1 ^ " \<times> " ^ aux prec C2
-         | CType (s, []) =>
-           if s = @{type_name prop} orelse s = @{type_name bool} then "o" else s
-         | CType (s, Cs) => "(" ^ commas (map (aux 0) Cs) ^ ") " ^ s
-         | CRec (s, _) => "[" ^ s ^ "]") ^
+        (if M = dummy_M then
+           "_"
+         else case M of
+             MAlpha => "\<alpha>"
+           | MFun (M1, a, M2) =>
+             aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^
+             string_for_sign_atom a ^ "\<^esup> " ^ aux prec M2
+           | MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2
+           | MType (s, []) =>
+             if s = @{type_name prop} orelse s = @{type_name bool} then "o"
+             else s
+           | MType (s, Ms) => "(" ^ commas (map (aux 0) Ms) ^ ") " ^ s
+           | MRec (s, _) => "[" ^ s ^ "]") ^
         (if need_parens then ")" else "")
       end
   in aux 0 end
 
-(* ctype -> ctype list *)
-fun flatten_ctype (CPair (C1, C2)) = maps flatten_ctype [C1, C2]
-  | flatten_ctype (CType (_, Cs)) = maps flatten_ctype Cs
-  | flatten_ctype C = [C]
+(* mtyp -> mtyp list *)
+fun flatten_mtype (MPair (M1, M2)) = maps flatten_mtype [M1, M2]
+  | flatten_mtype (MType (_, Ms)) = maps flatten_mtype Ms
+  | flatten_mtype M = [M]
+
+(* mterm -> bool *)
+fun precedence_of_mterm (MRaw _) = no_prec
+  | precedence_of_mterm (MAbs _) = 1
+  | precedence_of_mterm (MApp _) = 2
 
-(* hol_context -> bool -> typ -> cdata *)
-fun initial_cdata hol_ctxt binarize alpha_T =
+(* Proof.context -> mterm -> string *)
+fun string_for_mterm ctxt =
+  let
+    (* mtype -> string *)
+    fun mtype_annotation M = "\<^bsup>" ^ string_for_mtype M ^ "\<^esup>"
+    (* int -> mterm -> string *)
+    fun aux outer_prec m =
+      let
+        val prec = precedence_of_mterm m
+        val need_parens = (prec < outer_prec)
+      in
+        (if need_parens then "(" else "") ^
+        (case m of
+           MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M
+         | MAbs (s, _, M, a, m) =>
+           "\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^
+           string_for_sign_atom a ^ "\<^esup> " ^ aux prec m
+         | MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^
+        (if need_parens then ")" else "")
+      end
+  in aux 0 end
+
+(* mterm -> mtyp *)
+fun mtype_of_mterm (MRaw (_, M)) = M
+  | mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m)
+  | mtype_of_mterm (MApp (m1, _)) =
+    case mtype_of_mterm m1 of
+      MFun (_, _, M12) => M12
+    | M1 => raise MTYPE ("Nitpick_Mono.mtype_of_mterm", [M1])
+
+(* hol_context -> bool -> typ -> mdata *)
+fun initial_mdata hol_ctxt binarize alpha_T =
   ({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T,
     max_fresh = Unsynchronized.ref 0, datatype_cache = Unsynchronized.ref [],
-    constr_cache = Unsynchronized.ref []} : cdata)
+    constr_cache = Unsynchronized.ref []} : mdata)
 
 (* typ -> typ -> bool *)
 fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) =
@@ -127,172 +171,179 @@
                         exists (could_exist_alpha_subtype alpha_T) Ts)
   | could_exist_alpha_subtype alpha_T T = (T = alpha_T)
 (* theory -> typ -> typ -> bool *)
-fun could_exist_alpha_sub_ctype _ (alpha_T as TFree _) T =
+fun could_exist_alpha_sub_mtype _ (alpha_T as TFree _) T =
     could_exist_alpha_subtype alpha_T T
-  | could_exist_alpha_sub_ctype thy alpha_T T =
+  | could_exist_alpha_sub_mtype thy alpha_T T =
     (T = alpha_T orelse is_datatype thy [(NONE, true)] T)
 
-(* ctype -> bool *)
-fun exists_alpha_sub_ctype CAlpha = true
-  | exists_alpha_sub_ctype (CFun (C1, _, C2)) =
-    exists exists_alpha_sub_ctype [C1, C2]
-  | exists_alpha_sub_ctype (CPair (C1, C2)) =
-    exists exists_alpha_sub_ctype [C1, C2]
-  | exists_alpha_sub_ctype (CType (_, Cs)) = exists exists_alpha_sub_ctype Cs
-  | exists_alpha_sub_ctype (CRec _) = true
+(* mtyp -> bool *)
+fun exists_alpha_sub_mtype MAlpha = true
+  | exists_alpha_sub_mtype (MFun (M1, _, M2)) =
+    exists exists_alpha_sub_mtype [M1, M2]
+  | exists_alpha_sub_mtype (MPair (M1, M2)) =
+    exists exists_alpha_sub_mtype [M1, M2]
+  | exists_alpha_sub_mtype (MType (_, Ms)) = exists exists_alpha_sub_mtype Ms
+  | exists_alpha_sub_mtype (MRec _) = true
 
-(* ctype -> bool *)
-fun exists_alpha_sub_ctype_fresh CAlpha = true
-  | exists_alpha_sub_ctype_fresh (CFun (_, V _, _)) = true
-  | exists_alpha_sub_ctype_fresh (CFun (_, _, C2)) =
-    exists_alpha_sub_ctype_fresh C2
-  | exists_alpha_sub_ctype_fresh (CPair (C1, C2)) =
-    exists exists_alpha_sub_ctype_fresh [C1, C2]
-  | exists_alpha_sub_ctype_fresh (CType (_, Cs)) =
-    exists exists_alpha_sub_ctype_fresh Cs
-  | exists_alpha_sub_ctype_fresh (CRec _) = true
+(* mtyp -> bool *)
+fun exists_alpha_sub_mtype_fresh MAlpha = true
+  | exists_alpha_sub_mtype_fresh (MFun (_, V _, _)) = true
+  | exists_alpha_sub_mtype_fresh (MFun (_, _, M2)) =
+    exists_alpha_sub_mtype_fresh M2
+  | exists_alpha_sub_mtype_fresh (MPair (M1, M2)) =
+    exists exists_alpha_sub_mtype_fresh [M1, M2]
+  | exists_alpha_sub_mtype_fresh (MType (_, Ms)) =
+    exists exists_alpha_sub_mtype_fresh Ms
+  | exists_alpha_sub_mtype_fresh (MRec _) = true
 
-(* string * typ list -> ctype list -> ctype *)
-fun constr_ctype_for_binders z Cs =
-  fold_rev (fn C => curry3 CFun C (S Minus)) Cs (CRec z)
+(* string * typ list -> mtyp list -> mtyp *)
+fun constr_mtype_for_binders z Ms =
+  fold_rev (fn M => curry3 MFun M (S Minus)) Ms (MRec z)
 
-(* ((string * typ list) * ctype) list -> ctype list -> ctype -> ctype *)
-fun repair_ctype _ _ CAlpha = CAlpha
-  | repair_ctype cache seen (CFun (C1, a, C2)) =
-    CFun (repair_ctype cache seen C1, a, repair_ctype cache seen C2)
-  | repair_ctype cache seen (CPair Cp) =
-    CPair (pairself (repair_ctype cache seen) Cp)
-  | repair_ctype cache seen (CType (s, Cs)) =
-    CType (s, maps (flatten_ctype o repair_ctype cache seen) Cs)
-  | repair_ctype cache seen (CRec (z as (s, _))) =
+(* ((string * typ list) * mtyp) list -> mtyp list -> mtyp -> mtyp *)
+fun repair_mtype _ _ MAlpha = MAlpha
+  | repair_mtype cache seen (MFun (M1, a, M2)) =
+    MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2)
+  | repair_mtype cache seen (MPair Mp) =
+    MPair (pairself (repair_mtype cache seen) Mp)
+  | repair_mtype cache seen (MType (s, Ms)) =
+    MType (s, maps (flatten_mtype o repair_mtype cache seen) Ms)
+  | repair_mtype cache seen (MRec (z as (s, _))) =
     case AList.lookup (op =) cache z |> the of
-      CRec _ => CType (s, [])
-    | C => if member (op =) seen C then CType (s, [])
-           else repair_ctype cache (C :: seen) C
+      MRec _ => MType (s, [])
+    | M => if member (op =) seen M then MType (s, [])
+           else repair_mtype cache (M :: seen) M
 
-(* ((string * typ list) * ctype) list Unsynchronized.ref -> unit *)
+(* ((string * typ list) * mtyp) list Unsynchronized.ref -> unit *)
 fun repair_datatype_cache cache =
   let
-    (* (string * typ list) * ctype -> unit *)
-    fun repair_one (z, C) =
+    (* (string * typ list) * mtyp -> unit *)
+    fun repair_one (z, M) =
       Unsynchronized.change cache
-          (AList.update (op =) (z, repair_ctype (!cache) [] C))
+          (AList.update (op =) (z, repair_mtype (!cache) [] M))
   in List.app repair_one (rev (!cache)) end
 
-(* (typ * ctype) list -> (styp * ctype) list Unsynchronized.ref -> unit *)
+(* (typ * mtyp) list -> (styp * mtyp) list Unsynchronized.ref -> unit *)
 fun repair_constr_cache dtype_cache constr_cache =
   let
-    (* styp * ctype -> unit *)
-    fun repair_one (x, C) =
+    (* styp * mtyp -> unit *)
+    fun repair_one (x, M) =
       Unsynchronized.change constr_cache
-          (AList.update (op =) (x, repair_ctype dtype_cache [] C))
+          (AList.update (op =) (x, repair_mtype dtype_cache [] M))
   in List.app repair_one (!constr_cache) end
 
-(* cdata -> typ -> ctype *)
-fun fresh_ctype_for_type ({hol_ctxt as {thy, ...}, binarize, alpha_T, max_fresh,
-                           datatype_cache, constr_cache, ...} : cdata) =
+(* mdata -> typ -> typ -> mtyp * sign_atom * mtyp *)
+fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) T1 T2 =
   let
-    (* typ -> typ -> ctype *)
-    fun do_fun T1 T2 =
-      let
-        val C1 = do_type T1
-        val C2 = do_type T2
-        val a = if is_boolean_type (body_type T2) andalso
-                   exists_alpha_sub_ctype_fresh C1 then
-                  V (Unsynchronized.inc max_fresh)
-                else
-                  S Minus
-      in CFun (C1, a, C2) end
-    (* typ -> ctype *)
-    and do_type T =
+    val M1 = fresh_mtype_for_type mdata T1
+    val M2 = fresh_mtype_for_type mdata T2
+    val a = if is_boolean_type (body_type T2) andalso
+               exists_alpha_sub_mtype_fresh M1 then
+              V (Unsynchronized.inc max_fresh)
+            else
+              S Minus
+  in (M1, a, M2) end
+(* mdata -> typ -> mtyp *)
+and fresh_mtype_for_type (mdata as {hol_ctxt as {thy, ...}, binarize, alpha_T,
+                                    datatype_cache, constr_cache, ...}) =
+  let
+    (* typ -> typ -> mtyp *)
+    val do_fun = MFun oo fresh_mfun_for_fun_type mdata
+    (* typ -> mtyp *)
+    fun do_type T =
       if T = alpha_T then
-        CAlpha
+        MAlpha
       else case T of
         Type ("fun", [T1, T2]) => do_fun T1 T2
       | Type (@{type_name fun_box}, [T1, T2]) => do_fun T1 T2
-      | Type ("*", [T1, T2]) => CPair (pairself do_type (T1, T2))
+      | Type ("*", [T1, T2]) => MPair (pairself do_type (T1, T2))
       | Type (z as (s, _)) =>
-        if could_exist_alpha_sub_ctype thy alpha_T T then
+        if could_exist_alpha_sub_mtype thy alpha_T T then
           case AList.lookup (op =) (!datatype_cache) z of
-            SOME C => C
+            SOME M => M
           | NONE =>
             let
-              val _ = Unsynchronized.change datatype_cache (cons (z, CRec z))
+              val _ = Unsynchronized.change datatype_cache (cons (z, MRec z))
               val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
-              val (all_Cs, constr_Cs) =
-                fold_rev (fn (_, T') => fn (all_Cs, constr_Cs) =>
+              val (all_Ms, constr_Ms) =
+                fold_rev (fn (_, T') => fn (all_Ms, constr_Ms) =>
                              let
-                               val binder_Cs = map do_type (binder_types T')
-                               val new_Cs = filter exists_alpha_sub_ctype_fresh
-                                                   binder_Cs
-                               val constr_C = constr_ctype_for_binders z
-                                                                       binder_Cs
+                               val binder_Ms = map do_type (binder_types T')
+                               val new_Ms = filter exists_alpha_sub_mtype_fresh
+                                                   binder_Ms
+                               val constr_M = constr_mtype_for_binders z
+                                                                       binder_Ms
                              in
-                               (union (op =) new_Cs all_Cs,
-                                constr_C :: constr_Cs)
+                               (union (op =) new_Ms all_Ms,
+                                constr_M :: constr_Ms)
                              end)
                          xs ([], [])
-              val C = CType (s, all_Cs)
+              val M = MType (s, all_Ms)
               val _ = Unsynchronized.change datatype_cache
-                          (AList.update (op =) (z, C))
+                          (AList.update (op =) (z, M))
               val _ = Unsynchronized.change constr_cache
-                          (append (xs ~~ constr_Cs))
+                          (append (xs ~~ constr_Ms))
             in
-              if forall (not o is_CRec o snd) (!datatype_cache) then
+              if forall (not o is_MRec o snd) (!datatype_cache) then
                 (repair_datatype_cache datatype_cache;
                  repair_constr_cache (!datatype_cache) constr_cache;
                  AList.lookup (op =) (!datatype_cache) z |> the)
               else
-                C
+                M
             end
         else
-          CType (s, [])
-      | _ => CType (Refute.string_of_typ T, [])
+          MType (s, [])
+      | _ => MType (Refute.string_of_typ T, [])
   in do_type end
 
-(* ctype -> ctype list *)
-fun prodC_factors (CPair (C1, C2)) = maps prodC_factors [C1, C2]
-  | prodC_factors C = [C]
-(* ctype -> ctype list * ctype *)
-fun curried_strip_ctype (CFun (C1, S Minus, C2)) =
-    curried_strip_ctype C2 |>> append (prodC_factors C1)
-  | curried_strip_ctype C = ([], C)
-(* string -> ctype -> ctype *)
-fun sel_ctype_from_constr_ctype s C =
-  let val (arg_Cs, dataC) = curried_strip_ctype C in
-    CFun (dataC, S Minus,
-          case sel_no_from_name s of ~1 => bool_C | n => nth arg_Cs n)
+(* mtyp -> mtyp list *)
+fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2]
+  | prodM_factors M = [M]
+(* mtyp -> mtyp list * mtyp *)
+fun curried_strip_mtype (MFun (M1, S Minus, M2)) =
+    curried_strip_mtype M2 |>> append (prodM_factors M1)
+  | curried_strip_mtype M = ([], M)
+(* string -> mtyp -> mtyp *)
+fun sel_mtype_from_constr_mtype s M =
+  let val (arg_Ms, dataM) = curried_strip_mtype M in
+    MFun (dataM, S Minus,
+          case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n)
   end
 
-(* cdata -> styp -> ctype *)
-fun ctype_for_constr (cdata as {hol_ctxt = {thy, ...}, alpha_T, constr_cache,
+(* mdata -> styp -> mtyp *)
+fun mtype_for_constr (mdata as {hol_ctxt = {thy, ...}, alpha_T, constr_cache,
                                 ...}) (x as (_, T)) =
-  if could_exist_alpha_sub_ctype thy alpha_T T then
+  if could_exist_alpha_sub_mtype thy alpha_T T then
     case AList.lookup (op =) (!constr_cache) x of
-      SOME C => C
-    | NONE => (fresh_ctype_for_type cdata (body_type T);
-               AList.lookup (op =) (!constr_cache) x |> the)
+      SOME M => M
+    | NONE => if T = alpha_T then
+                let val M = fresh_mtype_for_type mdata T in
+                  (Unsynchronized.change constr_cache (cons (x, M)); M)
+                end
+              else
+                (fresh_mtype_for_type mdata (body_type T);
+                 AList.lookup (op =) (!constr_cache) x |> the)
   else
-    fresh_ctype_for_type cdata T
-fun ctype_for_sel (cdata as {hol_ctxt, binarize, ...}) (x as (s, _)) =
+    fresh_mtype_for_type mdata T
+fun mtype_for_sel (mdata as {hol_ctxt, binarize, ...}) (x as (s, _)) =
   x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
-    |> ctype_for_constr cdata |> sel_ctype_from_constr_ctype s
+    |> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
 
-(* literal list -> ctype -> ctype *)
-fun instantiate_ctype lits =
+(* literal list -> mtyp -> mtyp *)
+fun instantiate_mtype lits =
   let
-    (* ctype -> ctype *)
-    fun aux CAlpha = CAlpha
-      | aux (CFun (C1, V x, C2)) =
+    (* mtyp -> mtyp *)
+    fun aux MAlpha = MAlpha
+      | aux (MFun (M1, V x, M2)) =
         let
           val a = case AList.lookup (op =) lits x of
                     SOME sn => S sn
                   | NONE => V x
-        in CFun (aux C1, a, aux C2) end
-      | aux (CFun (C1, a, C2)) = CFun (aux C1, a, aux C2)
-      | aux (CPair Cp) = CPair (pairself aux Cp)
-      | aux (CType (s, Cs)) = CType (s, map aux Cs)
-      | aux (CRec z) = CRec z
+        in MFun (aux M1, a, aux M2) end
+      | aux (MFun (M1, a, M2)) = MFun (aux M1, a, aux M2)
+      | aux (MPair Mp) = MPair (pairself aux Mp)
+      | aux (MType (s, Ms)) = MType (s, map aux Ms)
+      | aux (MRec z) = MRec z
   in aux end
 
 datatype comp_op = Eq | Leq
@@ -342,105 +393,106 @@
   | do_sign_atom_comp cmp xs a1 a2 (lits, comps) =
     SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
 
-(* comp -> var list -> ctype -> ctype -> (literal list * comp list) option
+(* comp -> var list -> mtyp -> mtyp -> (literal list * comp list) option
    -> (literal list * comp list) option *)
-fun do_ctype_comp _ _ _ _ NONE = NONE
-  | do_ctype_comp _ _ CAlpha CAlpha accum = accum
-  | do_ctype_comp Eq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22))
+fun do_mtype_comp _ _ _ _ NONE = NONE
+  | do_mtype_comp _ _ MAlpha MAlpha accum = accum
+  | do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
                   (SOME accum) =
-     accum |> do_sign_atom_comp Eq xs a1 a2 |> do_ctype_comp Eq xs C11 C21
-           |> do_ctype_comp Eq xs C12 C22
-  | do_ctype_comp Leq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22))
+     accum |> do_sign_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21
+           |> do_mtype_comp Eq xs M12 M22
+  | do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
                   (SOME accum) =
-    (if exists_alpha_sub_ctype C11 then
+    (if exists_alpha_sub_mtype M11 then
        accum |> do_sign_atom_comp Leq xs a1 a2
-             |> do_ctype_comp Leq xs C21 C11
+             |> do_mtype_comp Leq xs M21 M11
              |> (case a2 of
                    S Minus => I
-                 | S Plus => do_ctype_comp Leq xs C11 C21
-                 | V x => do_ctype_comp Leq (x :: xs) C11 C21)
+                 | S Plus => do_mtype_comp Leq xs M11 M21
+                 | V x => do_mtype_comp Leq (x :: xs) M11 M21)
      else
        SOME accum)
-    |> do_ctype_comp Leq xs C12 C22
-  | do_ctype_comp cmp xs (C1 as CPair (C11, C12)) (C2 as CPair (C21, C22))
+    |> do_mtype_comp Leq xs M12 M22
+  | do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22))
                   accum =
-    (accum |> fold (uncurry (do_ctype_comp cmp xs)) [(C11, C21), (C12, C22)]
+    (accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
      handle Library.UnequalLengths =>
-            raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2]))
-  | do_ctype_comp _ _ (CType _) (CType _) accum =
+            raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2]))
+  | do_mtype_comp _ _ (MType _) (MType _) accum =
     accum (* no need to compare them thanks to the cache *)
-  | do_ctype_comp _ _ C1 C2 _ =
-    raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2])
+  | do_mtype_comp _ _ M1 M2 _ =
+    raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2])
 
-(* comp_op -> ctype -> ctype -> constraint_set -> constraint_set *)
-fun add_ctype_comp _ _ _ UnsolvableCSet = UnsolvableCSet
-  | add_ctype_comp cmp C1 C2 (CSet (lits, comps, sexps)) =
-    (print_g ("*** Add " ^ string_for_ctype C1 ^ " " ^ string_for_comp_op cmp ^
-              " " ^ string_for_ctype C2);
-     case do_ctype_comp cmp [] C1 C2 (SOME (lits, comps)) of
+(* comp_op -> mtyp -> mtyp -> constraint_set -> constraint_set *)
+fun add_mtype_comp _ _ _ UnsolvableCSet = UnsolvableCSet
+  | add_mtype_comp cmp M1 M2 (CSet (lits, comps, sexps)) =
+    (print_g ("*** Add " ^ string_for_mtype M1 ^ " " ^ string_for_comp_op cmp ^
+              " " ^ string_for_mtype M2);
+     case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
        NONE => (print_g "**** Unsolvable"; UnsolvableCSet)
      | SOME (lits, comps) => CSet (lits, comps, sexps))
 
-(* ctype -> ctype -> constraint_set -> constraint_set *)
-val add_ctypes_equal = add_ctype_comp Eq
-val add_is_sub_ctype = add_ctype_comp Leq
+(* mtyp -> mtyp -> constraint_set -> constraint_set *)
+val add_mtypes_equal = add_mtype_comp Eq
+val add_is_sub_mtype = add_mtype_comp Leq
 
-(* sign -> sign_expr -> ctype -> (literal list * sign_expr list) option
+(* sign -> sign_expr -> mtyp -> (literal list * sign_expr list) option
    -> (literal list * sign_expr list) option *)
-fun do_notin_ctype_fv _ _ _ NONE = NONE
-  | do_notin_ctype_fv Minus _ CAlpha accum = accum
-  | do_notin_ctype_fv Plus [] CAlpha _ = NONE
-  | do_notin_ctype_fv Plus [(x, sn)] CAlpha (SOME (lits, sexps)) =
+fun do_notin_mtype_fv _ _ _ NONE = NONE
+  | do_notin_mtype_fv Minus _ MAlpha accum = accum
+  | do_notin_mtype_fv Plus [] MAlpha _ = NONE
+  | do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) =
     SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
-  | do_notin_ctype_fv Plus sexp CAlpha (SOME (lits, sexps)) =
+  | do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) =
     SOME (lits, insert (op =) sexp sexps)
-  | do_notin_ctype_fv sn sexp (CFun (C1, S sn', C2)) accum =
+  | do_notin_mtype_fv sn sexp (MFun (M1, S sn', M2)) accum =
     accum |> (if sn' = Plus andalso sn = Plus then
-                do_notin_ctype_fv Plus sexp C1
+                do_notin_mtype_fv Plus sexp M1
               else
                 I)
           |> (if sn' = Minus orelse sn = Plus then
-                do_notin_ctype_fv Minus sexp C1
+                do_notin_mtype_fv Minus sexp M1
               else
                 I)
-          |> do_notin_ctype_fv sn sexp C2
-  | do_notin_ctype_fv Plus sexp (CFun (C1, V x, C2)) accum =
+          |> do_notin_mtype_fv sn sexp M2
+  | do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum =
     accum |> (case do_literal (x, Minus) (SOME sexp) of
                 NONE => I
-              | SOME sexp' => do_notin_ctype_fv Plus sexp' C1)
-          |> do_notin_ctype_fv Minus sexp C1
-          |> do_notin_ctype_fv Plus sexp C2
-  | do_notin_ctype_fv Minus sexp (CFun (C1, V x, C2)) accum =
+              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
+          |> do_notin_mtype_fv Minus sexp M1
+          |> do_notin_mtype_fv Plus sexp M2
+  | do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum =
     accum |> (case do_literal (x, Plus) (SOME sexp) of
                 NONE => I
-              | SOME sexp' => do_notin_ctype_fv Plus sexp' C1)
-          |> do_notin_ctype_fv Minus sexp C2
-  | do_notin_ctype_fv sn sexp (CPair (C1, C2)) accum =
-    accum |> fold (do_notin_ctype_fv sn sexp) [C1, C2]
-  | do_notin_ctype_fv sn sexp (CType (_, Cs)) accum =
-    accum |> fold (do_notin_ctype_fv sn sexp) Cs
-  | do_notin_ctype_fv _ _ C _ =
-    raise CTYPE ("Nitpick_Mono.do_notin_ctype_fv", [C])
+              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
+          |> do_notin_mtype_fv Minus sexp M2
+  | do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum =
+    accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2]
+  | do_notin_mtype_fv sn sexp (MType (_, Ms)) accum =
+    accum |> fold (do_notin_mtype_fv sn sexp) Ms
+  | do_notin_mtype_fv _ _ M _ =
+    raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M])
 
-(* sign -> ctype -> constraint_set -> constraint_set *)
-fun add_notin_ctype_fv _ _ UnsolvableCSet = UnsolvableCSet
-  | add_notin_ctype_fv sn C (CSet (lits, comps, sexps)) =
-    (print_g ("*** Add " ^ string_for_ctype C ^ " is right-" ^
+(* sign -> mtyp -> constraint_set -> constraint_set *)
+fun add_notin_mtype_fv _ _ UnsolvableCSet = UnsolvableCSet
+  | add_notin_mtype_fv sn M (CSet (lits, comps, sexps)) =
+    (print_g ("*** Add " ^ string_for_mtype M ^ " is right-" ^
               (case sn of Minus => "unique" | Plus => "total") ^ ".");
-     case do_notin_ctype_fv sn [] C (SOME (lits, sexps)) of
+     case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
        NONE => (print_g "**** Unsolvable"; UnsolvableCSet)
      | SOME (lits, sexps) => CSet (lits, comps, sexps))
 
-(* ctype -> constraint_set -> constraint_set *)
-val add_ctype_is_right_unique = add_notin_ctype_fv Minus
-val add_ctype_is_right_total = add_notin_ctype_fv Plus
+(* mtyp -> constraint_set -> constraint_set *)
+val add_mtype_is_right_unique = add_notin_mtype_fv Minus
+val add_mtype_is_right_total = add_notin_mtype_fv Plus
+
+val bool_from_minus = true
 
 (* sign -> bool *)
-fun bool_from_sign Plus = false
-  | bool_from_sign Minus = true
+fun bool_from_sign Plus = not bool_from_minus
+  | bool_from_sign Minus = bool_from_minus
 (* bool -> sign *)
-fun sign_from_bool false = Plus
-  | sign_from_bool true = Minus
+fun sign_from_bool b = if b = bool_from_minus then Minus else Plus
 
 (* literal -> PropLogic.prop_formula *)
 fun prop_for_literal (x, sn) =
@@ -492,228 +544,264 @@
              "-: " ^ commas (map (string_for_var o fst) neg))
   end
 
-(* var -> constraint_set -> literal list list option *)
+(* var -> constraint_set -> literal list option *)
 fun solve _ UnsolvableCSet = (print_g "*** Problem: Unsolvable"; NONE)
   | solve max_var (CSet (lits, comps, sexps)) =
     let
+      (* (int -> bool option) -> literal list option *)
+      fun do_assigns assigns =
+        SOME (literals_from_assignments max_var assigns lits
+              |> tap print_solution)
       val _ = print_problem lits comps sexps
       val prop = PropLogic.all (map prop_for_literal lits @
                                 map prop_for_comp comps @
                                 map prop_for_sign_expr sexps)
-      (* use the first ML solver (to avoid startup overhead) *)
-      val solvers = !SatSolver.solvers
-                    |> filter (member (op =) ["dptsat", "dpll"] o fst)
+      val default_val = bool_from_sign Minus
     in
-      case snd (hd solvers) prop of
-        SatSolver.SATISFIABLE assigns =>
-        SOME (literals_from_assignments max_var assigns lits
-              |> tap print_solution)
-      | _ => NONE
+      if PropLogic.eval (K default_val) prop then
+        do_assigns (K (SOME default_val))
+      else
+        let
+          (* use the first ML solver (to avoid startup overhead) *)
+          val solvers = !SatSolver.solvers
+                        |> filter (member (op =) ["dptsat", "dpll"] o fst)
+        in
+          case snd (hd solvers) prop of
+            SatSolver.SATISFIABLE assigns => do_assigns assigns
+          | _ => NONE
+        end
     end
 
-type ctype_schema = ctype * constraint_set
-type ctype_context =
-  {bounds: ctype list,
-   frees: (styp * ctype) list,
-   consts: (styp * ctype) list}
+type mtype_schema = mtyp * constraint_set
+type mtype_context =
+  {bounds: mtyp list,
+   frees: (styp * mtyp) list,
+   consts: (styp * mtyp) list}
 
-type accumulator = ctype_context * constraint_set
+type accumulator = mtype_context * constraint_set
 
 val initial_gamma = {bounds = [], frees = [], consts = []}
 val unsolvable_accum = (initial_gamma, UnsolvableCSet)
 
-(* ctype -> ctype_context -> ctype_context *)
-fun push_bound C {bounds, frees, consts} =
-  {bounds = C :: bounds, frees = frees, consts = consts}
-(* ctype_context -> ctype_context *)
+(* mtyp -> mtype_context -> mtype_context *)
+fun push_bound M {bounds, frees, consts} =
+  {bounds = M :: bounds, frees = frees, consts = consts}
+(* mtype_context -> mtype_context *)
 fun pop_bound {bounds, frees, consts} =
   {bounds = tl bounds, frees = frees, consts = consts}
   handle List.Empty => initial_gamma
 
-(* cdata -> term -> accumulator -> ctype * accumulator *)
-fun consider_term (cdata as {hol_ctxt = {thy, ctxt, stds, fast_descrs,
+(* mdata -> term -> accumulator -> mterm * accumulator *)
+fun consider_term (mdata as {hol_ctxt as {thy, ctxt, stds, fast_descrs,
                                          def_table, ...},
                              alpha_T, max_fresh, ...}) =
   let
-    (* typ -> ctype *)
-    val ctype_for = fresh_ctype_for_type cdata
-    (* ctype -> ctype *)
-    fun pos_set_ctype_for_dom C =
-      CFun (C, S (if exists_alpha_sub_ctype C then Plus else Minus), bool_C)
-    (* typ -> accumulator -> ctype * accumulator *)
-    fun do_quantifier T (gamma, cset) =
+    (* typ -> typ -> mtyp * sign_atom * mtyp *)
+    val mfun_for = fresh_mfun_for_fun_type mdata
+    (* typ -> mtyp *)
+    val mtype_for = fresh_mtype_for_type mdata
+    (* mtyp -> mtyp *)
+    fun pos_set_mtype_for_dom M =
+      MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
+    (* typ -> accumulator -> mterm * accumulator *)
+    fun do_all T (gamma, cset) =
       let
-        val abs_C = ctype_for (domain_type (domain_type T))
-        val body_C = ctype_for (range_type T)
+        val abs_M = mtype_for (domain_type (domain_type T))
+        val body_M = mtype_for (body_type T)
       in
-        (CFun (CFun (abs_C, S Minus, body_C), S Minus, body_C),
-         (gamma, cset |> add_ctype_is_right_total abs_C))
+        (MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M),
+         (gamma, cset |> add_mtype_is_right_total abs_M))
       end
     fun do_equals T (gamma, cset) =
-      let val C = ctype_for (domain_type T) in
-        (CFun (C, S Minus, CFun (C, V (Unsynchronized.inc max_fresh),
-                                 ctype_for (nth_range_type 2 T))),
-         (gamma, cset |> add_ctype_is_right_unique C))
+      let val M = mtype_for (domain_type T) in
+        (MFun (M, S Minus, MFun (M, V (Unsynchronized.inc max_fresh),
+                                 mtype_for (nth_range_type 2 T))),
+         (gamma, cset |> add_mtype_is_right_unique M))
       end
     fun do_robust_set_operation T (gamma, cset) =
       let
         val set_T = domain_type T
-        val C1 = ctype_for set_T
-        val C2 = ctype_for set_T
-        val C3 = ctype_for set_T
+        val M1 = mtype_for set_T
+        val M2 = mtype_for set_T
+        val M3 = mtype_for set_T
       in
-        (CFun (C1, S Minus, CFun (C2, S Minus, C3)),
-         (gamma, cset |> add_is_sub_ctype C1 C3 |> add_is_sub_ctype C2 C3))
+        (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
+         (gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3))
       end
     fun do_fragile_set_operation T (gamma, cset) =
       let
         val set_T = domain_type T
-        val set_C = ctype_for set_T
-        (* typ -> ctype *)
-        fun custom_ctype_for (T as Type ("fun", [T1, T2])) =
-            if T = set_T then set_C
-            else CFun (custom_ctype_for T1, S Minus, custom_ctype_for T2)
-          | custom_ctype_for T = ctype_for T
+        val set_M = mtype_for set_T
+        (* typ -> mtyp *)
+        fun custom_mtype_for (T as Type ("fun", [T1, T2])) =
+            if T = set_T then set_M
+            else MFun (custom_mtype_for T1, S Minus, custom_mtype_for T2)
+          | custom_mtype_for T = mtype_for T
       in
-        (custom_ctype_for T, (gamma, cset |> add_ctype_is_right_unique set_C))
+        (custom_mtype_for T, (gamma, cset |> add_mtype_is_right_unique set_M))
       end
-    (* typ -> accumulator -> ctype * accumulator *)
+    (* typ -> accumulator -> mtyp * accumulator *)
     fun do_pair_constr T accum =
-      case ctype_for (nth_range_type 2 T) of
-        C as CPair (a_C, b_C) =>
-        (CFun (a_C, S Minus, CFun (b_C, S Minus, C)), accum)
-      | C => raise CTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [C])
-    (* int -> typ -> accumulator -> ctype * accumulator *)
+      case mtype_for (nth_range_type 2 T) of
+        M as MPair (a_M, b_M) =>
+        (MFun (a_M, S Minus, MFun (b_M, S Minus, M)), accum)
+      | M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M])
+    (* int -> typ -> accumulator -> mtyp * accumulator *)
     fun do_nth_pair_sel n T =
-      case ctype_for (domain_type T) of
-        C as CPair (a_C, b_C) =>
-        pair (CFun (C, S Minus, if n = 0 then a_C else b_C))
-      | C => raise CTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [C])
-    val unsolvable = (CType ("unsolvable", []), unsolvable_accum)
-    (* typ -> term -> accumulator -> ctype * accumulator *)
-    fun do_bounded_quantifier abs_T bound_t body_t accum =
+      case mtype_for (domain_type T) of
+        M as MPair (a_M, b_M) =>
+        pair (MFun (M, S Minus, if n = 0 then a_M else b_M))
+      | M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M])
+    (* mtyp * accumulator *)
+    val mtype_unsolvable = (dummy_M, unsolvable_accum)
+    (* term -> mterm * accumulator *)
+    fun mterm_unsolvable t = (MRaw (t, dummy_M), unsolvable_accum)
+    (* term -> string -> typ -> term -> term -> term -> accumulator
+       -> mterm * accumulator *)
+    fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum =
       let
-        val abs_C = ctype_for abs_T
-        val (bound_C, accum) = accum |>> push_bound abs_C |> do_term bound_t
-        val expected_bound_C = pos_set_ctype_for_dom abs_C
+        val abs_M = mtype_for abs_T
+        val (bound_m, accum) = accum |>> push_bound abs_M |> do_term bound_t
+        val expected_bound_M = pos_set_mtype_for_dom abs_M
+        val (body_m, accum) =
+          accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
+                |> do_term body_t ||> apfst pop_bound
+        val bound_M = mtype_of_mterm bound_m
+        val (M1, a, M2) = dest_MFun bound_M
       in
-        accum ||> add_ctypes_equal expected_bound_C bound_C |> do_term body_t
-              ||> apfst pop_bound
+        (MApp (MRaw (t0, MFun (bound_M, S Minus, bool_M)),
+               MAbs (abs_s, abs_T, M1, a,
+                     MApp (MApp (MRaw (connective_t,
+                                       mtype_for (fastype_of connective_t)),
+                                 MApp (bound_m, MRaw (Bound 0, M1))),
+                           body_m))), accum)
       end
-    (* term -> accumulator -> ctype * accumulator *)
-    and do_term _ (_, UnsolvableCSet) = unsolvable
+    (* term -> accumulator -> mterm * accumulator *)
+    and do_term t (_, UnsolvableCSet) = mterm_unsolvable t
       | do_term t (accum as (gamma as {bounds, frees, consts}, cset)) =
         (case t of
            Const (x as (s, T)) =>
            (case AList.lookup (op =) consts x of
-              SOME C => (C, accum)
+              SOME M => (M, accum)
             | NONE =>
               if not (could_exist_alpha_subtype alpha_T T) then
-                (ctype_for T, accum)
+                (mtype_for T, accum)
               else case s of
-                @{const_name all} => do_quantifier T accum
+                @{const_name all} => do_all T accum
               | @{const_name "=="} => do_equals T accum
-              | @{const_name All} => do_quantifier T accum
-              | @{const_name Ex} => do_quantifier T accum
+              | @{const_name All} => do_all T accum
+              | @{const_name Ex} =>
+                let val set_T = domain_type T in
+                  do_term (Abs (Name.uu, set_T,
+                                @{const Not} $ (HOLogic.mk_eq
+                                    (Abs (Name.uu, domain_type set_T,
+                                          @{const False}),
+                                     Bound 0)))) accum
+                  |>> mtype_of_mterm
+                end
               | @{const_name "op ="} => do_equals T accum
-              | @{const_name The} => (print_g "*** The"; unsolvable)
-              | @{const_name Eps} => (print_g "*** Eps"; unsolvable)
+              | @{const_name The} => (print_g "*** The"; mtype_unsolvable)
+              | @{const_name Eps} => (print_g "*** Eps"; mtype_unsolvable)
               | @{const_name If} =>
                 do_robust_set_operation (range_type T) accum
-                |>> curry3 CFun bool_C (S Minus)
+                |>> curry3 MFun bool_M (S Minus)
               | @{const_name Pair} => do_pair_constr T accum
               | @{const_name fst} => do_nth_pair_sel 0 T accum
               | @{const_name snd} => do_nth_pair_sel 1 T accum 
               | @{const_name Id} =>
-                (CFun (ctype_for (domain_type T), S Minus, bool_C), accum)
+                (MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
               | @{const_name insert} =>
                 let
                   val set_T = domain_type (range_type T)
-                  val C1 = ctype_for (domain_type set_T)
-                  val C1' = pos_set_ctype_for_dom C1
-                  val C2 = ctype_for set_T
-                  val C3 = ctype_for set_T
+                  val M1 = mtype_for (domain_type set_T)
+                  val M1' = pos_set_mtype_for_dom M1
+                  val M2 = mtype_for set_T
+                  val M3 = mtype_for set_T
                 in
-                  (CFun (C1, S Minus, CFun (C2, S Minus, C3)),
-                   (gamma, cset |> add_ctype_is_right_unique C1
-                                |> add_is_sub_ctype C1' C3
-                                |> add_is_sub_ctype C2 C3))
+                  (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
+                   (gamma, cset |> add_mtype_is_right_unique M1
+                                |> add_is_sub_mtype M1' M3
+                                |> add_is_sub_mtype M2 M3))
                 end
               | @{const_name converse} =>
                 let
                   val x = Unsynchronized.inc max_fresh
-                  (* typ -> ctype *)
-                  fun ctype_for_set T =
-                    CFun (ctype_for (domain_type T), V x, bool_C)
-                  val ab_set_C = domain_type T |> ctype_for_set
-                  val ba_set_C = range_type T |> ctype_for_set
-                in (CFun (ab_set_C, S Minus, ba_set_C), accum) end
+                  (* typ -> mtyp *)
+                  fun mtype_for_set T =
+                    MFun (mtype_for (domain_type T), V x, bool_M)
+                  val ab_set_M = domain_type T |> mtype_for_set
+                  val ba_set_M = range_type T |> mtype_for_set
+                in (MFun (ab_set_M, S Minus, ba_set_M), accum) end
               | @{const_name trancl} => do_fragile_set_operation T accum
-              | @{const_name rtrancl} => (print_g "*** rtrancl"; unsolvable)
+              | @{const_name rtrancl} =>
+                (print_g "*** rtrancl"; mtype_unsolvable)
               | @{const_name finite} =>
-                let val C1 = ctype_for (domain_type (domain_type T)) in
-                  (CFun (pos_set_ctype_for_dom C1, S Minus, bool_C), accum)
-                end
+                if is_finite_type hol_ctxt T then
+                  let val M1 = mtype_for (domain_type (domain_type T)) in
+                    (MFun (pos_set_mtype_for_dom M1, S Minus, bool_M), accum)
+                  end
+                else
+                  (print_g "*** finite"; mtype_unsolvable)
               | @{const_name rel_comp} =>
                 let
                   val x = Unsynchronized.inc max_fresh
-                  (* typ -> ctype *)
-                  fun ctype_for_set T =
-                    CFun (ctype_for (domain_type T), V x, bool_C)
-                  val bc_set_C = domain_type T |> ctype_for_set
-                  val ab_set_C = domain_type (range_type T) |> ctype_for_set
-                  val ac_set_C = nth_range_type 2 T |> ctype_for_set
+                  (* typ -> mtyp *)
+                  fun mtype_for_set T =
+                    MFun (mtype_for (domain_type T), V x, bool_M)
+                  val bc_set_M = domain_type T |> mtype_for_set
+                  val ab_set_M = domain_type (range_type T) |> mtype_for_set
+                  val ac_set_M = nth_range_type 2 T |> mtype_for_set
                 in
-                  (CFun (bc_set_C, S Minus, CFun (ab_set_C, S Minus, ac_set_C)),
+                  (MFun (bc_set_M, S Minus, MFun (ab_set_M, S Minus, ac_set_M)),
                    accum)
                 end
               | @{const_name image} =>
                 let
-                  val a_C = ctype_for (domain_type (domain_type T))
-                  val b_C = ctype_for (range_type (domain_type T))
+                  val a_M = mtype_for (domain_type (domain_type T))
+                  val b_M = mtype_for (range_type (domain_type T))
                 in
-                  (CFun (CFun (a_C, S Minus, b_C), S Minus,
-                         CFun (pos_set_ctype_for_dom a_C, S Minus,
-                               pos_set_ctype_for_dom b_C)), accum)
+                  (MFun (MFun (a_M, S Minus, b_M), S Minus,
+                         MFun (pos_set_mtype_for_dom a_M, S Minus,
+                               pos_set_mtype_for_dom b_M)), accum)
                 end
               | @{const_name Sigma} =>
                 let
                   val x = Unsynchronized.inc max_fresh
-                  (* typ -> ctype *)
-                  fun ctype_for_set T =
-                    CFun (ctype_for (domain_type T), V x, bool_C)
+                  (* typ -> mtyp *)
+                  fun mtype_for_set T =
+                    MFun (mtype_for (domain_type T), V x, bool_M)
                   val a_set_T = domain_type T
-                  val a_C = ctype_for (domain_type a_set_T)
-                  val b_set_C = ctype_for_set (range_type (domain_type
+                  val a_M = mtype_for (domain_type a_set_T)
+                  val b_set_M = mtype_for_set (range_type (domain_type
                                                                (range_type T)))
-                  val a_set_C = ctype_for_set a_set_T
-                  val a_to_b_set_C = CFun (a_C, S Minus, b_set_C)
-                  val ab_set_C = ctype_for_set (nth_range_type 2 T)
+                  val a_set_M = mtype_for_set a_set_T
+                  val a_to_b_set_M = MFun (a_M, S Minus, b_set_M)
+                  val ab_set_M = mtype_for_set (nth_range_type 2 T)
                 in
-                  (CFun (a_set_C, S Minus,
-                         CFun (a_to_b_set_C, S Minus, ab_set_C)), accum)
+                  (MFun (a_set_M, S Minus,
+                         MFun (a_to_b_set_M, S Minus, ab_set_M)), accum)
                 end
               | @{const_name Tha} =>
                 let
-                  val a_C = ctype_for (domain_type (domain_type T))
-                  val a_set_C = pos_set_ctype_for_dom a_C
-                in (CFun (a_set_C, S Minus, a_C), accum) end
+                  val a_M = mtype_for (domain_type (domain_type T))
+                  val a_set_M = pos_set_mtype_for_dom a_M
+                in (MFun (a_set_M, S Minus, a_M), accum) end
               | @{const_name FunBox} =>
-                let val dom_C = ctype_for (domain_type T) in
-                  (CFun (dom_C, S Minus, dom_C), accum)
+                let val dom_M = mtype_for (domain_type T) in
+                  (MFun (dom_M, S Minus, dom_M), accum)
                 end
               | _ =>
                 if s = @{const_name minus_class.minus} andalso
                    is_set_type (domain_type T) then
                   let
                     val set_T = domain_type T
-                    val left_set_C = ctype_for set_T
-                    val right_set_C = ctype_for set_T
+                    val left_set_M = mtype_for set_T
+                    val right_set_M = mtype_for set_T
                   in
-                    (CFun (left_set_C, S Minus,
-                           CFun (right_set_C, S Minus, left_set_C)),
-                     (gamma, cset |> add_ctype_is_right_unique right_set_C
-                                  |> add_is_sub_ctype right_set_C left_set_C))
+                    (MFun (left_set_M, S Minus,
+                           MFun (right_set_M, S Minus, left_set_M)),
+                     (gamma, cset |> add_mtype_is_right_unique right_set_M
+                                  |> add_is_sub_mtype right_set_M left_set_M))
                   end
                 else if s = @{const_name ord_class.less_eq} andalso
                         is_set_type (domain_type T) then
@@ -724,34 +812,37 @@
                   do_robust_set_operation T accum
                 else if is_sel s then
                   if constr_name_for_sel_like s = @{const_name FunBox} then
-                    let val dom_C = ctype_for (domain_type T) in
-                      (CFun (dom_C, S Minus, dom_C), accum)
+                    let val dom_M = mtype_for (domain_type T) in
+                      (MFun (dom_M, S Minus, dom_M), accum)
                     end
                   else
-                    (ctype_for_sel cdata x, accum)
+                    (mtype_for_sel mdata x, accum)
                 else if is_constr thy stds x then
-                  (ctype_for_constr cdata x, accum)
+                  (mtype_for_constr mdata x, accum)
                 else if is_built_in_const thy stds fast_descrs x then
                   case def_of_const thy def_table x of
-                    SOME t' => do_term t' accum
-                  | NONE => (print_g ("*** built-in " ^ s); unsolvable)
+                    SOME t' => do_term t' accum |>> mtype_of_mterm
+                  | NONE => (print_g ("*** built-in " ^ s); mtype_unsolvable)
                 else
-                  let val C = ctype_for T in
-                    (C, ({bounds = bounds, frees = frees,
-                          consts = (x, C) :: consts}, cset))
-                  end)
+                  let val M = mtype_for T in
+                    (M, ({bounds = bounds, frees = frees,
+                          consts = (x, M) :: consts}, cset))
+                  end) |>> curry MRaw t
          | Free (x as (_, T)) =>
            (case AList.lookup (op =) frees x of
-              SOME C => (C, accum)
+              SOME M => (M, accum)
             | NONE =>
-              let val C = ctype_for T in
-                (C, ({bounds = bounds, frees = (x, C) :: frees,
+              let val M = mtype_for T in
+                (M, ({bounds = bounds, frees = (x, M) :: frees,
                       consts = consts}, cset))
-              end)
-         | Var _ => (print_g "*** Var"; unsolvable)
-         | Bound j => (nth bounds j, accum)
-         | Abs (_, T, @{const False}) => (ctype_for (T --> bool_T), accum)
-         | Abs (_, T, t') =>
+              end) |>> curry MRaw t
+         | Var _ => (print_g "*** Var"; mterm_unsolvable t)
+         | Bound j => (MRaw (t, nth bounds j), accum)
+         | Abs (s, T, t' as @{const False}) =>
+           let val (M1, a, M2) = mfun_for T bool_T in
+             (MAbs (s, T, M1, a, MRaw (t', M2)), accum)
+           end
+         | Abs (s, T, t') =>
            ((case t' of
                t1' $ Bound 0 =>
                if not (loose_bvar1 (t1', 0)) then
@@ -761,107 +852,127 @@
              | _ => raise SAME ())
             handle SAME () =>
                    let
-                     val C = ctype_for T
-                     val (C', accum) = do_term t' (accum |>> push_bound C)
-                   in (CFun (C, S Minus, C'), accum |>> pop_bound) end)
-         | Const (@{const_name All}, _)
-           $ Abs (_, T', @{const "op -->"} $ (t1 $ Bound 0) $ t2) =>
-           do_bounded_quantifier T' t1 t2 accum
-         | Const (@{const_name Ex}, _)
-           $ Abs (_, T', @{const "op &"} $ (t1 $ Bound 0) $ t2) =>
-           do_bounded_quantifier T' t1 t2 accum
+                     val M = mtype_for T
+                     val (m', accum) = do_term t' (accum |>> push_bound M)
+                   in (MAbs (s, T, M, S Minus, m'), accum |>> pop_bound) end)
+         | (t0 as Const (@{const_name All}, _))
+           $ Abs (s', T', (t10 as @{const "op -->"}) $ (t11 $ Bound 0) $ t12) =>
+           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
+         | (t0 as Const (@{const_name Ex}, _))
+           $ Abs (s', T', (t10 as @{const "op &"}) $ (t11 $ Bound 0) $ t12) =>
+           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
          | Const (@{const_name Let}, _) $ t1 $ t2 =>
            do_term (betapply (t2, t1)) accum
          | t1 $ t2 =>
            let
-             val (C1, accum) = do_term t1 accum
-             val (C2, accum) = do_term t2 accum
+             val (m1, accum) = do_term t1 accum
+             val (m2, accum) = do_term t2 accum
            in
              case accum of
-               (_, UnsolvableCSet) => unsolvable
-             | _ => case C1 of
-                      CFun (C11, _, C12) =>
-                      (C12, accum ||> add_is_sub_ctype C2 C11)
-                    | _ => raise CTYPE ("Nitpick_Mono.consider_term.do_term \
-                                        \(op $)", [C1])
+               (_, UnsolvableCSet) => mterm_unsolvable t
+             | _ =>
+               let
+                 val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
+                 val M2 = mtype_of_mterm m2
+               in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
            end)
-        |> tap (fn (C, _) =>
-                   print_g ("  \<Gamma> \<turnstile> " ^
-                            Syntax.string_of_term ctxt t ^ " : " ^
-                            string_for_ctype C))
+        |> tap (fn (m, _) => print_g ("  \<Gamma> \<turnstile> " ^
+                                      string_for_mterm ctxt m))
   in do_term end
 
-(* cdata -> sign -> term -> accumulator -> accumulator *)
-fun consider_general_formula (cdata as {hol_ctxt = {ctxt, ...}, ...}) =
+(* mdata -> styp -> term -> term -> mterm * accumulator *)
+fun consider_general_equals mdata (x as (_, T)) t1 t2 accum =
   let
-    (* typ -> ctype *)
-    val ctype_for = fresh_ctype_for_type cdata
-    (* term -> accumulator -> ctype * accumulator *)
-    val do_term = consider_term cdata
-    (* sign -> term -> accumulator -> accumulator *)
-    fun do_formula _ _ (_, UnsolvableCSet) = unsolvable_accum
-      | do_formula sn t (accum as (gamma, cset)) =
+    val (m1, accum) = consider_term mdata t1 accum
+    val (m2, accum) = consider_term mdata t2 accum
+    val M1 = mtype_of_mterm m1
+    val M2 = mtype_of_mterm m2
+    val body_M = fresh_mtype_for_type mdata (nth_range_type 2 T)
+  in
+    (MApp (MApp (MRaw (Const x,
+         MFun (M1, S Minus, MFun (M2, S Minus, body_M))), m1), m2),
+     accum ||> add_mtypes_equal M1 M2)
+  end
+
+(* mdata -> sign -> term -> accumulator -> mterm * accumulator *)
+fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
+  let
+    (* typ -> mtyp *)
+    val mtype_for = fresh_mtype_for_type mdata
+    (* term -> accumulator -> mterm * accumulator *)
+    val do_term = consider_term mdata
+    (* sign -> term -> accumulator -> mterm * accumulator *)
+    fun do_formula _ t (_, UnsolvableCSet) =
+        (MRaw (t, dummy_M), unsolvable_accum)
+      | do_formula sn t accum =
         let
-          (* term -> accumulator -> accumulator *)
-          val do_co_formula = do_formula sn
-          val do_contra_formula = do_formula (negate sn)
-          (* string -> typ -> term -> accumulator *)
-          fun do_quantifier quant_s abs_T body_t =
+          (* styp -> string -> typ -> term -> mterm * accumulator *)
+          fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
             let
-              val abs_C = ctype_for abs_T
+              val abs_M = mtype_for abs_T
               val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
-              val cset = cset |> side_cond ? add_ctype_is_right_total abs_C
+              val (body_m, accum) =
+                accum ||> side_cond ? add_mtype_is_right_total abs_M
+                      |>> push_bound abs_M |> do_formula sn body_t
+              val body_M = mtype_of_mterm body_m
             in
-              (gamma |> push_bound abs_C, cset) |> do_co_formula body_t
-                                                |>> pop_bound
+              (MApp (MRaw (Const quant_x, MFun (abs_M, S Minus, body_M)),
+                     MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
+               accum |>> pop_bound)
             end
-          (* typ -> term -> accumulator *)
-          fun do_bounded_quantifier abs_T body_t =
-            accum |>> push_bound (ctype_for abs_T) |> do_co_formula body_t
-                  |>> pop_bound
-          (* term -> term -> accumulator *)
-          fun do_equals t1 t2 =
+          (* styp -> term -> term -> mterm * accumulator *)
+          fun do_equals x t1 t2 =
             case sn of
-              Plus => do_term t accum |> snd
-            | Minus => let
-                         val (C1, accum) = do_term t1 accum
-                         val (C2, accum) = do_term t2 accum
-                       in accum ||> add_ctypes_equal C1 C2 end
+              Plus => do_term t accum
+            | Minus => consider_general_equals mdata x t1 t2 accum
         in
           case t of
-            Const (s0 as @{const_name all}, _) $ Abs (_, T1, t1) =>
-            do_quantifier s0 T1 t1
-          | Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2
-          | @{const "==>"} $ t1 $ t2 =>
-            accum |> do_contra_formula t1 |> do_co_formula t2
-          | @{const Trueprop} $ t1 => do_co_formula t1 accum
-          | @{const Not} $ t1 => do_contra_formula t1 accum
-          | Const (@{const_name All}, _)
-            $ Abs (_, T1, t1 as @{const "op -->"} $ (_ $ Bound 0) $ _) =>
-            do_bounded_quantifier T1 t1
-          | Const (s0 as @{const_name All}, _) $ Abs (_, T1, t1) =>
-            do_quantifier s0 T1 t1
-          | Const (@{const_name Ex}, _)
-            $ Abs (_, T1, t1 as @{const "op &"} $ (_ $ Bound 0) $ _) =>
-            do_bounded_quantifier T1 t1
-          | Const (s0 as @{const_name Ex}, _) $ Abs (_, T1, t1) =>
-            do_quantifier s0 T1 t1
-          | Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2
-          | @{const "op &"} $ t1 $ t2 =>
-            accum |> do_co_formula t1 |> do_co_formula t2
-          | @{const "op |"} $ t1 $ t2 =>
-            accum |> do_co_formula t1 |> do_co_formula t2
-          | @{const "op -->"} $ t1 $ t2 =>
-            accum |> do_contra_formula t1 |> do_co_formula t2
-          | Const (@{const_name If}, _) $ t1 $ t2 $ t3 =>
-            accum |> do_term t1 |> snd |> fold do_co_formula [t2, t3]
-          | Const (@{const_name Let}, _) $ t1 $ t2 =>
-            do_co_formula (betapply (t2, t1)) accum
-          | _ => do_term t accum |> snd
+            Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+            do_quantifier x s1 T1 t1
+          | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
+          | @{const Trueprop} $ t1 =>
+            let val (m1, accum) = do_formula sn t1 accum in
+              (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
+                     m1), accum)
+            end
+          | @{const Not} $ t1 =>
+            let val (m1, accum) = do_formula (negate sn) t1 accum in
+              (MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1),
+               accum)
+            end
+          | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
+            do_quantifier x s1 T1 t1
+          | Const (x0 as (s0 as @{const_name Ex}, T0))
+            $ (t1 as Abs (s1, T1, t1')) =>
+            (case sn of
+               Plus => do_quantifier x0 s1 T1 t1'
+             | Minus =>
+               (* ### do elsewhere *)
+               do_term (@{const Not}
+                        $ (HOLogic.eq_const (domain_type T0) $ t1
+                           $ Abs (Name.uu, T1, @{const False}))) accum)
+          | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+            do_equals x t1 t2
+          | (t0 as Const (s0, _)) $ t1 $ t2 =>
+            if s0 = @{const_name "==>"} orelse s0 = @{const_name "op &"} orelse
+               s0 = @{const_name "op |"} orelse s0 = @{const_name "op -->"} then
+              let
+                val impl = (s0 = @{const_name "==>"} orelse
+                           s0 = @{const_name "op -->"})
+                val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
+                val (m2, accum) = do_formula sn t2 accum
+              in
+                (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
+                 accum)
+              end 
+            else
+              do_term t accum
+          | _ => do_term t accum
         end
-        |> tap (fn _ => print_g ("\<Gamma> \<turnstile> " ^
-                                 Syntax.string_of_term ctxt t ^
-                                 " : o\<^sup>" ^ string_for_sign sn))
+        |> tap (fn (m, _) =>
+                   print_g ("\<Gamma> \<turnstile> " ^
+                            string_for_mterm ctxt m ^ " : o\<^sup>" ^
+                            string_for_sign sn))
   in do_formula end
 
 (* The harmless axiom optimization below is somewhat too aggressive in the face
@@ -877,79 +988,110 @@
   |> (forall (member (op =) harmless_consts o original_name o fst)
       orf exists (member (op =) bounteous_consts o fst))
 
-(* cdata -> sign -> term -> accumulator -> accumulator *)
-fun consider_nondefinitional_axiom (cdata as {hol_ctxt, ...}) sn t =
-  not (is_harmless_axiom hol_ctxt t) ? consider_general_formula cdata sn t
+(* mdata -> term -> accumulator -> mterm * accumulator *)
+fun consider_nondefinitional_axiom (mdata as {hol_ctxt, ...}) t =
+  if is_harmless_axiom hol_ctxt t then pair (MRaw (t, dummy_M))
+  else consider_general_formula mdata Plus t
 
-(* cdata -> term -> accumulator -> accumulator *)
-fun consider_definitional_axiom (cdata as {hol_ctxt as {thy, ...}, ...}) t =
+(* mdata -> term -> accumulator -> mterm * accumulator *)
+fun consider_definitional_axiom (mdata as {hol_ctxt as {thy, ...}, ...}) t =
   if not (is_constr_pattern_formula thy t) then
-    consider_nondefinitional_axiom cdata Plus t
+    consider_nondefinitional_axiom mdata t
   else if is_harmless_axiom hol_ctxt t then
-    I
+    pair (MRaw (t, dummy_M))
   else
     let
-      (* term -> accumulator -> ctype * accumulator *)
-      val do_term = consider_term cdata
-      (* typ -> term -> accumulator -> accumulator *)
-      fun do_all abs_T body_t accum =
-        let val abs_C = fresh_ctype_for_type cdata abs_T in
-          accum |>> push_bound abs_C |> do_formula body_t |>> pop_bound
+      (* typ -> mtyp *)
+      val mtype_for = fresh_mtype_for_type mdata
+      (* term -> accumulator -> mterm * accumulator *)
+      val do_term = consider_term mdata
+      (* term -> string -> typ -> term -> accumulator -> mterm * accumulator *)
+      fun do_all quant_t abs_s abs_T body_t accum =
+        let
+          val abs_M = mtype_for abs_T
+          val (body_m, accum) =
+            accum |>> push_bound abs_M |> do_formula body_t
+          val body_M = mtype_of_mterm body_m
+        in
+          (MApp (MRaw (quant_t, MFun (abs_M, S Minus, body_M)),
+                 MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
+           accum |>> pop_bound)
         end
-      (* term -> term -> accumulator -> accumulator *)
-      and do_implies t1 t2 = do_term t1 #> snd #> do_formula t2
-      and do_equals t1 t2 accum =
+      (* term -> term -> term -> accumulator -> mterm * accumulator *)
+      and do_conjunction t0 t1 t2 accum =
         let
-          val (C1, accum) = do_term t1 accum
-          val (C2, accum) = do_term t2 accum
-        in accum ||> add_ctypes_equal C1 C2 end
+          val (m1, accum) = do_formula t1 accum
+          val (m2, accum) = do_formula t2 accum
+        in
+          (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
+        end
+      and do_implies t0 t1 t2 accum =
+        let
+          val (m1, accum) = do_term t1 accum
+          val (m2, accum) = do_formula t2 accum
+        in
+          (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
+        end
       (* term -> accumulator -> accumulator *)
-      and do_formula _ (_, UnsolvableCSet) = unsolvable_accum
+      and do_formula t (_, UnsolvableCSet) =
+          (MRaw (t, dummy_M), unsolvable_accum)
         | do_formula t accum =
           case t of
-            Const (@{const_name all}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum
+            (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+            do_all t0 s1 T1 t1 accum
           | @{const Trueprop} $ t1 => do_formula t1 accum
-          | Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2 accum
-          | @{const "==>"} $ t1 $ t2 => do_implies t1 t2 accum
-          | @{const Pure.conjunction} $ t1 $ t2 =>
-            accum |> do_formula t1 |> do_formula t2
-          | Const (@{const_name All}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum
-          | Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 accum
-          | @{const "op &"} $ t1 $ t2 => accum |> do_formula t1 |> do_formula t2
-          | @{const "op -->"} $ t1 $ t2 => do_implies t1 t2 accum
+          | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
+            consider_general_equals mdata x t1 t2 accum
+          | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+          | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
+            do_conjunction t0 t1 t2 accum
+          | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
+            do_all t0 s0 T1 t1 accum
+          | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+            consider_general_equals mdata x t1 t2 accum
+          | (t0 as @{const "op &"}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
+          | (t0 as @{const "op -->"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
           | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
                              \do_formula", [t])
     in do_formula t end
 
-(* Proof.context -> literal list -> term -> ctype -> string *)
-fun string_for_ctype_of_term ctxt lits t C =
+(* Proof.context -> literal list -> term -> mtyp -> string *)
+fun string_for_mtype_of_term ctxt lits t M =
   Syntax.string_of_term ctxt t ^ " : " ^
-  string_for_ctype (instantiate_ctype lits C)
+  string_for_mtype (instantiate_mtype lits M)
 
-(* theory -> literal list -> ctype_context -> unit *)
-fun print_ctype_context ctxt lits ({frees, consts, ...} : ctype_context) =
-  map (fn (x, C) => string_for_ctype_of_term ctxt lits (Free x) C) frees @
-  map (fn (x, C) => string_for_ctype_of_term ctxt lits (Const x) C) consts
+(* theory -> literal list -> mtype_context -> unit *)
+fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
+  map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
+  map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
   |> cat_lines |> print_g
 
-(* hol_context -> bool -> typ -> sign -> term list -> term list -> term
-   -> bool *)
-fun formulas_monotonic (hol_ctxt as {ctxt, ...}) binarize alpha_T sn def_ts
-                       nondef_ts core_t =
+(* ('a -> 'b -> 'c * 'd) -> 'a -> 'c list * 'b -> 'c list * 'd *)
+fun gather f t (ms, accum) =
+  let val (m, accum) = f t accum in (m :: ms, accum) end
+
+(* hol_context -> bool -> typ -> term list * term list -> bool *)
+fun formulas_monotonic (hol_ctxt as {ctxt, ...}) binarize alpha_T
+                       (nondef_ts, def_ts) =
   let
-    val _ = print_g ("****** " ^ string_for_ctype CAlpha ^ " is " ^
+    val _ = print_g ("****** Monotonicity analysis: " ^
+                     string_for_mtype MAlpha ^ " is " ^
                      Syntax.string_of_typ ctxt alpha_T)
-    val cdata as {max_fresh, ...} = initial_cdata hol_ctxt binarize alpha_T
-    val (gamma, cset) =
-      (initial_gamma, slack)
-      |> fold (consider_definitional_axiom cdata) def_ts
-      |> fold (consider_nondefinitional_axiom cdata Plus) nondef_ts
-      |> consider_general_formula cdata sn core_t
+    val mdata as {max_fresh, constr_cache, ...} =
+      initial_mdata hol_ctxt binarize alpha_T
+
+    val accum = (initial_gamma, slack)
+    val (nondef_ms, accum) =
+      ([], accum) |> gather (consider_general_formula mdata Plus) (hd nondef_ts)
+                  |> fold (gather (consider_nondefinitional_axiom mdata))
+                          (tl nondef_ts)
+    val (def_ms, (gamma, cset)) =
+      ([], accum) |> fold (gather (consider_definitional_axiom mdata)) def_ts
   in
     case solve (!max_fresh) cset of
-      SOME lits => (print_ctype_context ctxt lits gamma; true)
+      SOME lits => (print_mtype_context ctxt lits gamma; true)
     | _ => false
   end
-  handle CTYPE (loc, Cs) => raise BAD (loc, commas (map string_for_ctype Cs))
+  handle MTYPE (loc, Ms) => raise BAD (loc, commas (map string_for_mtype Ms))
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -400,13 +400,13 @@
   | name_ord (_, BoundName _) = GREATER
   | name_ord (FreeName (s1, T1, _), FreeName (s2, T2, _)) =
     (case fast_string_ord (s1, s2) of
-       EQUAL => TermOrd.typ_ord (T1, T2)
+       EQUAL => Term_Ord.typ_ord (T1, T2)
      | ord => ord)
   | name_ord (FreeName _, _) = LESS
   | name_ord (_, FreeName _) = GREATER
   | name_ord (ConstName (s1, T1, _), ConstName (s2, T2, _)) =
     (case fast_string_ord (s1, s2) of
-       EQUAL => TermOrd.typ_ord (T1, T2)
+       EQUAL => Term_Ord.typ_ord (T1, T2)
      | ord => ord)
   | name_ord (u1, u2) = raise NUT ("Nitpick_Nut.name_ord", [u1, u2])
 
@@ -1031,7 +1031,7 @@
             if is_opt_rep R then
               triad_fn (fn polar => s_op2 Subset T (Formula polar) u1 u2)
             else if opt andalso polar = Pos andalso
-                    not (is_concrete_type datatypes (type_of u1)) then
+                    not (is_concrete_type datatypes true (type_of u1)) then
               Cst (False, T, Formula Pos)
             else
               s_op2 Subset T R u1 u2
@@ -1057,7 +1057,7 @@
               else opt_opt_case ()
           in
             if unsound orelse polar = Neg orelse
-               is_concrete_type datatypes (type_of u1) then
+               is_concrete_type datatypes true (type_of u1) then
               case (is_opt_rep (rep_of u1'), is_opt_rep (rep_of u2')) of
                 (true, true) => opt_opt_case ()
               | (true, false) => hybrid_case u1'
@@ -1159,7 +1159,7 @@
                 let val quant_u = s_op2 oper T (Formula polar) u1' u2' in
                   if def orelse
                      (unsound andalso (polar = Pos) = (oper = All)) orelse
-                     is_complete_type datatypes (type_of u1) then
+                     is_complete_type datatypes true (type_of u1) then
                     quant_u
                   else
                     let
@@ -1202,7 +1202,7 @@
                       else unopt_R |> opt ? opt_rep ofs T
               val u = Op2 (oper, T, R, u1', sub u2)
             in
-              if is_complete_type datatypes T orelse not opt1 then
+              if is_complete_type datatypes true T orelse not opt1 then
                 u
               else
                 let
--- a/src/HOL/Tools/Nitpick/nitpick_peephole.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_peephole.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -126,7 +126,7 @@
 val norm_frac_rel = (4, 0)
 
 (* int -> bool -> rel_expr *)
-fun atom_for_bool j0 = Atom o Integer.add j0 o int_for_bool
+fun atom_for_bool j0 = Atom o Integer.add j0 o int_from_bool
 (* bool -> formula *)
 fun formula_for_bool b = if b then True else False
 
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -9,8 +9,7 @@
 sig
   type hol_context = Nitpick_HOL.hol_context
   val preprocess_term :
-    hol_context -> term
-    -> ((term list * term list) * (bool * bool)) * term * bool
+    hol_context -> term -> term list * term list * bool * bool * bool
 end
 
 structure Nitpick_Preproc : NITPICK_PREPROC =
@@ -136,6 +135,59 @@
                                      (index_seq 0 (length arg_Ts)) arg_Ts)
          end
 
+(* (term -> term) -> int -> term -> term *)
+fun coerce_bound_no f j t =
+  case t of
+    t1 $ t2 => coerce_bound_no f j t1 $ coerce_bound_no f j t2
+  | Abs (s, T, t') => Abs (s, T, coerce_bound_no f (j + 1) t')
+  | Bound j' => if j' = j then f t else t
+  | _ => t
+(* hol_context -> typ -> typ -> term -> term *)
+fun coerce_bound_0_in_term hol_ctxt new_T old_T =
+  old_T <> new_T ? coerce_bound_no (coerce_term hol_ctxt [new_T] old_T new_T) 0
+(* hol_context -> typ list -> typ -> typ -> term -> term *)
+and coerce_term (hol_ctxt as {thy, stds, fast_descrs, ...}) Ts new_T old_T t =
+  if old_T = new_T then
+    t
+  else
+    case (new_T, old_T) of
+      (Type (new_s, new_Ts as [new_T1, new_T2]),
+       Type ("fun", [old_T1, old_T2])) =>
+      (case eta_expand Ts t 1 of
+         Abs (s, _, t') =>
+         Abs (s, new_T1,
+              t' |> coerce_bound_0_in_term hol_ctxt new_T1 old_T1
+                 |> coerce_term hol_ctxt (new_T1 :: Ts) new_T2 old_T2)
+         |> Envir.eta_contract
+         |> new_s <> "fun"
+            ? construct_value thy stds
+                  (@{const_name FunBox}, Type ("fun", new_Ts) --> new_T)
+                  o single
+       | t' => raise TERM ("Nitpick_Preproc.coerce_term", [t']))
+    | (Type (new_s, new_Ts as [new_T1, new_T2]),
+       Type (old_s, old_Ts as [old_T1, old_T2])) =>
+      if old_s = @{type_name fun_box} orelse
+         old_s = @{type_name pair_box} orelse old_s = "*" then
+        case constr_expand hol_ctxt old_T t of
+          Const (old_s, _) $ t1 =>
+          if new_s = "fun" then
+            coerce_term hol_ctxt Ts new_T (Type ("fun", old_Ts)) t1
+          else
+            construct_value thy stds
+                (old_s, Type ("fun", new_Ts) --> new_T)
+                [coerce_term hol_ctxt Ts (Type ("fun", new_Ts))
+                             (Type ("fun", old_Ts)) t1]
+        | Const _ $ t1 $ t2 =>
+          construct_value thy stds
+              (if new_s = "*" then @{const_name Pair}
+               else @{const_name PairBox}, new_Ts ---> new_T)
+              [coerce_term hol_ctxt Ts new_T1 old_T1 t1,
+               coerce_term hol_ctxt Ts new_T2 old_T2 t2]
+        | t' => raise TERM ("Nitpick_Preproc.coerce_term", [t'])
+      else
+        raise TYPE ("Nitpick_Preproc.coerce_term", [new_T, old_T], [t])
+    | _ => raise TYPE ("Nitpick_Preproc.coerce_term", [new_T, old_T], [t])
+
 (* hol_context -> bool -> term -> term *)
 fun box_fun_and_pair_in_term (hol_ctxt as {thy, stds, fast_descrs, ...}) def
                              orig_t =
@@ -146,60 +198,6 @@
       | box_relational_operator_type (Type ("*", Ts)) =
         Type ("*", map (box_type hol_ctxt InPair) Ts)
       | box_relational_operator_type T = T
-    (* (term -> term) -> int -> term -> term *)
-    fun coerce_bound_no f j t =
-      case t of
-        t1 $ t2 => coerce_bound_no f j t1 $ coerce_bound_no f j t2
-      | Abs (s, T, t') => Abs (s, T, coerce_bound_no f (j + 1) t')
-      | Bound j' => if j' = j then f t else t
-      | _ => t
-    (* typ -> typ -> term -> term *)
-    fun coerce_bound_0_in_term new_T old_T =
-      old_T <> new_T ? coerce_bound_no (coerce_term [new_T] old_T new_T) 0
-    (* typ list -> typ -> term -> term *)
-    and coerce_term Ts new_T old_T t =
-      if old_T = new_T then
-        t
-      else
-        case (new_T, old_T) of
-          (Type (new_s, new_Ts as [new_T1, new_T2]),
-           Type ("fun", [old_T1, old_T2])) =>
-          (case eta_expand Ts t 1 of
-             Abs (s, _, t') =>
-             Abs (s, new_T1,
-                  t' |> coerce_bound_0_in_term new_T1 old_T1
-                     |> coerce_term (new_T1 :: Ts) new_T2 old_T2)
-             |> Envir.eta_contract
-             |> new_s <> "fun"
-                ? construct_value thy stds
-                      (@{const_name FunBox}, Type ("fun", new_Ts) --> new_T)
-                      o single
-           | t' => raise TERM ("Nitpick_Preproc.box_fun_and_pair_in_term.\
-                               \coerce_term", [t']))
-        | (Type (new_s, new_Ts as [new_T1, new_T2]),
-           Type (old_s, old_Ts as [old_T1, old_T2])) =>
-          if old_s = @{type_name fun_box} orelse
-             old_s = @{type_name pair_box} orelse old_s = "*" then
-            case constr_expand hol_ctxt old_T t of
-              Const (@{const_name FunBox}, _) $ t1 =>
-              if new_s = "fun" then
-                coerce_term Ts new_T (Type ("fun", old_Ts)) t1
-              else
-                construct_value thy stds
-                    (@{const_name FunBox}, Type ("fun", new_Ts) --> new_T)
-                    [coerce_term Ts (Type ("fun", new_Ts))
-                                 (Type ("fun", old_Ts)) t1]
-            | Const _ $ t1 $ t2 =>
-              construct_value thy stds
-                  (if new_s = "*" then @{const_name Pair}
-                   else @{const_name PairBox}, new_Ts ---> new_T)
-                  [coerce_term Ts new_T1 old_T1 t1,
-                   coerce_term Ts new_T2 old_T2 t2]
-            | t' => raise TERM ("Nitpick_Preproc.box_fun_and_pair_in_term.\
-                                \coerce_term", [t'])
-          else
-            raise TYPE ("coerce_term", [new_T, old_T], [t])
-        | _ => raise TYPE ("coerce_term", [new_T, old_T], [t])
     (* indexname * typ -> typ * term -> typ option list -> typ option list *)
     fun add_boxed_types_for_var (z as (_, T)) (T', t') =
       case t' of
@@ -249,10 +247,10 @@
       let
         val (t1, t2) = pairself (do_term new_Ts old_Ts Neut) (t1, t2)
         val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2)
-        val T = [T1, T2] |> sort TermOrd.typ_ord |> List.last
+        val T = [T1, T2] |> sort Term_Ord.typ_ord |> List.last
       in
         list_comb (Const (s0, T --> T --> body_type T0),
-                   map2 (coerce_term new_Ts T) [T1, T2] [t1, t2])
+                   map2 (coerce_term hol_ctxt new_Ts T) [T1, T2] [t1, t2])
       end
     (* string -> typ -> term *)
     and do_description_operator s T =
@@ -320,7 +318,7 @@
           val T' = hd (snd (dest_Type (hd Ts1)))
           val t2 = Abs (s, T', do_term (T' :: new_Ts) (T :: old_Ts) Neut t2')
           val T2 = fastype_of1 (new_Ts, t2)
-          val t2 = coerce_term new_Ts (hd Ts1) T2 t2
+          val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2
         in
           betapply (if s1 = "fun" then
                       t1
@@ -336,7 +334,7 @@
           val (s1, Ts1) = dest_Type T1
           val t2 = do_term new_Ts old_Ts Neut t2
           val T2 = fastype_of1 (new_Ts, t2)
-          val t2 = coerce_term new_Ts (hd Ts1) T2 t2
+          val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2
         in
           betapply (if s1 = "fun" then
                       t1
@@ -474,6 +472,19 @@
           (list_comb (t, args), seen)
   in aux [] 0 t [] [] |> fst end
 
+val let_var_prefix = nitpick_prefix ^ "l"
+val let_inline_threshold = 32
+
+(* int -> typ -> term -> (term -> term) -> term *)
+fun hol_let n abs_T body_T f t =
+  if n * size_of_term t <= let_inline_threshold then
+    f t
+  else
+    let val z = ((let_var_prefix, 0), abs_T) in
+      Const (@{const_name Let}, abs_T --> (abs_T --> body_T) --> body_T)
+      $ t $ abs_var z (incr_boundvars 1 (f (Var z)))
+    end
+
 (* hol_context -> bool -> term -> term *)
 fun destroy_pulled_out_constrs (hol_ctxt as {thy, stds, ...}) axiom t =
   let
@@ -508,14 +519,19 @@
           if z1 = z2 andalso num_occs_of_var z1 = 2 then @{const True}
           else raise SAME ()
         | (Const (x as (s, T)), args) =>
-          let val arg_Ts = binder_types T in
-            if length arg_Ts = length args andalso
-               (is_constr thy stds x orelse s = @{const_name Pair}) andalso
+          let
+            val (arg_Ts, dataT) = strip_type T
+            val n = length arg_Ts
+          in
+            if length args = n andalso
+               (is_constr thy stds x orelse s = @{const_name Pair} orelse
+                x = (@{const_name Suc}, nat_T --> nat_T)) andalso
                (not careful orelse not (is_Var t1) orelse
                 String.isPrefix val_var_prefix (fst (fst (dest_Var t1)))) then
-              discriminate_value hol_ctxt x t1 ::
-              map3 (sel_eq x t1) (index_seq 0 (length args)) arg_Ts args
-              |> foldr1 s_conj
+                hol_let (n + 1) dataT bool_T
+                    (fn t1 => discriminate_value hol_ctxt x t1 ::
+                              map3 (sel_eq x t1) (index_seq 0 n) arg_Ts args
+                              |> foldr1 s_conj) t1
             else
               raise SAME ()
           end
@@ -799,7 +815,7 @@
       | call_sets [] uss vs = vs :: call_sets uss [] []
       | call_sets ([] :: _) _ _ = []
       | call_sets ((t :: ts) :: tss) uss vs =
-        OrdList.insert TermOrd.term_ord t vs |> call_sets tss (ts :: uss)
+        OrdList.insert Term_Ord.term_ord t vs |> call_sets tss (ts :: uss)
     val sets = call_sets (fun_calls t [] []) [] []
     val indexed_sets = sets ~~ (index_seq 0 (length sets))
   in
@@ -814,7 +830,7 @@
 (* hol_context -> styp -> term list -> (int * term option) list *)
 fun static_args_in_terms hol_ctxt x =
   map (static_args_in_term hol_ctxt x)
-  #> fold1 (OrdList.inter (prod_ord int_ord (option_ord TermOrd.term_ord)))
+  #> fold1 (OrdList.inter (prod_ord int_ord (option_ord Term_Ord.term_ord)))
 
 (* (int * term option) list -> (int * term) list -> int list *)
 fun overlapping_indices [] _ = []
@@ -866,7 +882,7 @@
                 val _ = not (null fixed_js) orelse raise SAME ()
                 val fixed_args = filter_indices fixed_js args
                 val vars = fold Term.add_vars fixed_args []
-                           |> sort (TermOrd.fast_indexname_ord o pairself fst)
+                           |> sort (Term_Ord.fast_indexname_ord o pairself fst)
                 val bound_js = fold (fn t => fn js => add_loose_bnos (t, 0, js))
                                     fixed_args []
                                |> sort int_ord
@@ -956,7 +972,7 @@
     fun generality (js, _, _) = ~(length js)
     (* special_triple -> special_triple -> bool *)
     fun is_more_specific (j1, t1, x1) (j2, t2, x2) =
-      x1 <> x2 andalso OrdList.subset (prod_ord int_ord TermOrd.term_ord)
+      x1 <> x2 andalso OrdList.subset (prod_ord int_ord Term_Ord.term_ord)
                                       (j2 ~~ t2, j1 ~~ t1)
     (* typ -> special_triple list -> special_triple list -> special_triple list
        -> term list -> term list *)
@@ -1020,7 +1036,7 @@
 (* Prevents divergence in case of cyclic or infinite axiom dependencies. *)
 val axioms_max_depth = 255
 
-(* hol_context -> term -> (term list * term list) * (bool * bool) *)
+(* hol_context -> term -> term list * term list * bool * bool *)
 fun axioms_for_term
         (hol_ctxt as {thy, ctxt, max_bisim_depth, stds, user_axioms,
                       fast_descrs, evals, def_table, nondef_table, user_nondefs,
@@ -1170,10 +1186,9 @@
                      |> user_axioms = SOME true
                         ? fold (add_nondef_axiom 1) mono_user_nondefs
     val defs = defs @ special_congruence_axioms hol_ctxt xs
-  in
-    ((defs, nondefs), (user_axioms = SOME true orelse null mono_user_nondefs,
-                       null poly_user_nondefs))
-  end
+    val got_all_mono_user_axioms =
+      (user_axioms = SOME true orelse null mono_user_nondefs)
+  in (t :: nondefs, defs, got_all_mono_user_axioms, null poly_user_nondefs) end
 
 (** Simplification of constructor/selector terms **)
 
@@ -1275,7 +1290,7 @@
 
 (* Maximum number of quantifiers in a cluster for which the exponential
    algorithm is used. Larger clusters use a heuristic inspired by Claessen &
-   Sörensson's polynomial binary splitting procedure (p. 5 of their MODEL 2003
+   Soerensson's polynomial binary splitting procedure (p. 5 of their MODEL 2003
    paper). *)
 val quantifier_cluster_threshold = 7
 
@@ -1386,29 +1401,29 @@
 
 (** Preprocessor entry point **)
 
-(* hol_context -> term
-   -> ((term list * term list) * (bool * bool)) * term * bool *)
+(* hol_context -> term -> term list * term list * bool * bool * bool *)
 fun preprocess_term (hol_ctxt as {thy, stds, binary_ints, destroy_constrs,
                                   boxes, skolemize, uncurry, ...}) t =
   let
     val skolem_depth = if skolemize then 4 else ~1
-    val (((def_ts, nondef_ts), (got_all_mono_user_axioms, no_poly_user_axioms)),
-         core_t) = t |> unfold_defs_in_term hol_ctxt
-                     |> close_form
-                     |> skolemize_term_and_more hol_ctxt skolem_depth
-                     |> specialize_consts_in_term hol_ctxt 0
-                     |> `(axioms_for_term hol_ctxt)
+    val (nondef_ts, def_ts, got_all_mono_user_axioms, no_poly_user_axioms) =
+      t |> unfold_defs_in_term hol_ctxt
+        |> close_form
+        |> skolemize_term_and_more hol_ctxt skolem_depth
+        |> specialize_consts_in_term hol_ctxt 0
+        |> axioms_for_term hol_ctxt
     val binarize =
       is_standard_datatype thy stds nat_T andalso
       case binary_ints of
         SOME false => false
-      | _ => forall may_use_binary_ints (core_t :: def_ts @ nondef_ts) andalso
+      | _ => forall may_use_binary_ints (nondef_ts @ def_ts) andalso
              (binary_ints = SOME true orelse
-              exists should_use_binary_ints (core_t :: def_ts @ nondef_ts))
+              exists should_use_binary_ints (nondef_ts @ def_ts))
     val box = exists (not_equal (SOME false) o snd) boxes
+    val uncurry = uncurry andalso box
     val table =
-      Termtab.empty |> uncurry
-        ? fold (add_to_uncurry_table thy) (core_t :: def_ts @ nondef_ts)
+      Termtab.empty
+      |> uncurry ? fold (add_to_uncurry_table thy) (nondef_ts @ def_ts)
     (* bool -> term -> term *)
     fun do_rest def =
       binarize ? binarize_nat_and_int_in_term
@@ -1425,10 +1440,10 @@
       #> push_quantifiers_inward
       #> close_form
       #> Term.map_abs_vars shortest_name
+    val nondef_ts = map (do_rest false) nondef_ts
+    val def_ts = map (do_rest true) def_ts
   in
-    (((map (do_rest true) def_ts, map (do_rest false) nondef_ts),
-      (got_all_mono_user_axioms, no_poly_user_axioms)),
-     do_rest false core_t, binarize)
+    (nondef_ts, def_ts, got_all_mono_user_axioms, no_poly_user_axioms, binarize)
   end
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_rep.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -299,7 +299,7 @@
      | z => Func z)
   | best_non_opt_set_rep_for_type scope T = best_one_rep_for_type scope T
 fun best_set_rep_for_type (scope as {datatypes, ...}) T =
-  (if is_exact_type datatypes T then best_non_opt_set_rep_for_type
+  (if is_exact_type datatypes true T then best_non_opt_set_rep_for_type
    else best_opt_set_rep_for_type) scope T
 fun best_non_opt_symmetric_reps_for_fun_type (scope as {ofs, ...})
                                              (Type ("fun", [T1, T2])) =
--- a/src/HOL/Tools/Nitpick/nitpick_scope.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -23,8 +23,8 @@
     card: int,
     co: bool,
     standard: bool,
-    complete: bool,
-    concrete: bool,
+    complete: bool * bool,
+    concrete: bool * bool,
     deep: bool,
     constrs: constr_spec list}
 
@@ -39,9 +39,9 @@
 
   val datatype_spec : dtype_spec list -> typ -> dtype_spec option
   val constr_spec : dtype_spec list -> styp -> constr_spec
-  val is_complete_type : dtype_spec list -> typ -> bool
-  val is_concrete_type : dtype_spec list -> typ -> bool
-  val is_exact_type : dtype_spec list -> typ -> bool
+  val is_complete_type : dtype_spec list -> bool -> typ -> bool
+  val is_concrete_type : dtype_spec list -> bool -> typ -> bool
+  val is_exact_type : dtype_spec list -> bool -> typ -> bool
   val offset_of_type : int Typtab.table -> typ -> int
   val spec_of_type : scope -> typ -> int * int
   val pretties_for_scope : scope -> bool -> Pretty.T list
@@ -51,7 +51,7 @@
   val all_scopes :
     hol_context -> bool -> int -> (typ option * int list) list
     -> (styp option * int list) list -> (styp option * int list) list
-    -> int list -> int list -> typ list -> typ list -> typ list
+    -> int list -> int list -> typ list -> typ list -> typ list -> typ list
     -> int * scope list
 end;
 
@@ -74,8 +74,8 @@
   card: int,
   co: bool,
   standard: bool,
-  complete: bool,
-  concrete: bool,
+  complete: bool * bool,
+  concrete: bool * bool,
   deep: bool,
   constrs: constr_spec list}
 
@@ -105,22 +105,24 @@
       SOME c => c
     | NONE => constr_spec dtypes x
 
-(* dtype_spec list -> typ -> bool *)
-fun is_complete_type dtypes (Type ("fun", [T1, T2])) =
-    is_concrete_type dtypes T1 andalso is_complete_type dtypes T2
-  | is_complete_type dtypes (Type ("*", Ts)) =
-    forall (is_complete_type dtypes) Ts
-  | is_complete_type dtypes T =
+(* dtype_spec list -> bool -> typ -> bool *)
+fun is_complete_type dtypes facto (Type ("fun", [T1, T2])) =
+    is_concrete_type dtypes facto T1 andalso is_complete_type dtypes facto T2
+  | is_complete_type dtypes facto (Type ("*", Ts)) =
+    forall (is_complete_type dtypes facto) Ts
+  | is_complete_type dtypes facto T =
     not (is_integer_like_type T) andalso not (is_bit_type T) andalso
-    #complete (the (datatype_spec dtypes T))
+    fun_from_pair (#complete (the (datatype_spec dtypes T))) facto
     handle Option.Option => true
-and is_concrete_type dtypes (Type ("fun", [T1, T2])) =
-    is_complete_type dtypes T1 andalso is_concrete_type dtypes T2
-  | is_concrete_type dtypes (Type ("*", Ts)) =
-    forall (is_concrete_type dtypes) Ts
-  | is_concrete_type dtypes T =
-    #concrete (the (datatype_spec dtypes T)) handle Option.Option => true
-fun is_exact_type dtypes = is_complete_type dtypes andf is_concrete_type dtypes
+and is_concrete_type dtypes facto (Type ("fun", [T1, T2])) =
+    is_complete_type dtypes facto T1 andalso is_concrete_type dtypes facto T2
+  | is_concrete_type dtypes facto (Type ("*", Ts)) =
+    forall (is_concrete_type dtypes facto) Ts
+  | is_concrete_type dtypes facto T =
+    fun_from_pair (#concrete (the (datatype_spec dtypes T))) facto
+    handle Option.Option => true
+fun is_exact_type dtypes facto =
+  is_complete_type dtypes facto andf is_concrete_type dtypes facto
 
 (* int Typtab.table -> typ -> int *)
 fun offset_of_type ofs T =
@@ -461,15 +463,18 @@
      explicit_max = max, total = total} :: constrs
   end
 
-(* hol_context -> (typ * int) list -> typ -> bool *)
-fun has_exact_card hol_ctxt card_assigns T =
+(* hol_context -> bool -> typ list -> (typ * int) list -> typ -> bool *)
+fun has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T =
   let val card = card_of_type card_assigns T in
-    card = bounded_exact_card_of_type hol_ctxt (card + 1) 0 card_assigns T
+    card = bounded_exact_card_of_type hol_ctxt
+               (if facto then finitizable_dataTs else []) (card + 1) 0
+               card_assigns T
   end
 
-(* hol_context -> bool -> typ list -> scope_desc -> typ * int -> dtype_spec *)
+(* hol_context -> bool -> typ list -> typ list -> scope_desc -> typ * int
+   -> dtype_spec *)
 fun datatype_spec_from_scope_descriptor (hol_ctxt as {thy, stds, ...}) binarize
-        deep_dataTs (desc as (card_assigns, _)) (T, card) =
+        deep_dataTs finitizable_dataTs (desc as (card_assigns, _)) (T, card) =
   let
     val deep = member (op =) deep_dataTs T
     val co = is_codatatype thy T
@@ -478,10 +483,16 @@
     val self_recs = map (is_self_recursive_constr_type o snd) xs
     val (num_self_recs, num_non_self_recs) =
       List.partition I self_recs |> pairself length
-    val complete = has_exact_card hol_ctxt card_assigns T
-    val concrete = is_word_type T orelse
-                   (xs |> maps (binder_types o snd) |> maps binder_types
-                       |> forall (has_exact_card hol_ctxt card_assigns))
+    (* bool -> bool *)
+    fun is_complete facto =
+      has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T
+    fun is_concrete facto =
+      is_word_type T orelse
+      xs |> maps (binder_types o snd) |> maps binder_types
+         |> forall (has_exact_card hol_ctxt facto finitizable_dataTs
+                                   card_assigns)
+    val complete = pair_from_fun is_complete
+    val concrete = pair_from_fun is_concrete
     (* int -> int *)
     fun sum_dom_cards max =
       map (domain_card max card_assigns o snd) xs |> Integer.sum
@@ -494,13 +505,15 @@
      concrete = concrete, deep = deep, constrs = constrs}
   end
 
-(* hol_context -> bool -> int -> typ list -> scope_desc -> scope *)
+(* hol_context -> bool -> int -> typ list -> typ list -> scope_desc -> scope *)
 fun scope_from_descriptor (hol_ctxt as {thy, stds, ...}) binarize sym_break
-                          deep_dataTs (desc as (card_assigns, _)) =
+                          deep_dataTs finitizable_dataTs
+                          (desc as (card_assigns, _)) =
   let
     val datatypes =
       map (datatype_spec_from_scope_descriptor hol_ctxt binarize deep_dataTs
-               desc) (filter (is_datatype thy stds o fst) card_assigns)
+                                               finitizable_dataTs desc)
+          (filter (is_datatype thy stds o fst) card_assigns)
     val bits = card_of_type card_assigns @{typ signed_bit} - 1
                handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
                       card_of_type card_assigns @{typ unsigned_bit}
@@ -530,10 +543,10 @@
 
 (* hol_context -> bool -> int -> (typ option * int list) list
    -> (styp option * int list) list -> (styp option * int list) list -> int list
-   -> typ list -> typ list -> typ list -> int * scope list *)
+   -> typ list -> typ list -> typ list ->typ list -> int * scope list *)
 fun all_scopes (hol_ctxt as {thy, ...}) binarize sym_break cards_assigns
                maxes_assigns iters_assigns bitss bisim_depths mono_Ts nonmono_Ts
-               deep_dataTs =
+               deep_dataTs finitizable_dataTs =
   let
     val cards_assigns = repair_cards_assigns_wrt_boxing thy mono_Ts
                                                         cards_assigns
@@ -550,7 +563,7 @@
     (length all - length head,
      descs |> length descs <= distinct_threshold ? distinct (op =)
            |> map (scope_from_descriptor hol_ctxt binarize sym_break
-                                         deep_dataTs))
+                                         deep_dataTs finitizable_dataTs))
   end
 
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_tests.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -19,9 +19,7 @@
 open Nitpick_Nut
 open Nitpick_Kodkod
 
-val settings =
-  [("solver", "\"zChaff\""),
-   ("skolem_depth", "-1")]
+val settings = [("solver", "\"DefaultSAT4J\"")]
 
 fun cast_to_rep R u = Op1 (Cast, type_of u, R, u)
 
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -20,7 +20,9 @@
   val nitpick_prefix : string
   val curry3 : ('a * 'b * 'c -> 'd) -> 'a -> 'b -> 'c -> 'd
   val pairf : ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c
-  val int_for_bool : bool -> int
+  val pair_from_fun : (bool -> 'a) -> 'a * 'a
+  val fun_from_pair : 'a * 'a -> bool -> 'a
+  val int_from_bool : bool -> int
   val nat_minus : int -> int -> int
   val reasonable_power : int -> int -> int
   val exact_log : int -> int -> int
@@ -84,8 +86,13 @@
 (* ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c *)
 fun pairf f g x = (f x, g x)
 
+(* (bool -> 'a) -> 'a * 'a *)
+fun pair_from_fun f = (f false, f true)
+(* 'a * 'a -> bool -> 'a *)
+fun fun_from_pair (f, t) b = if b then t else f
+
 (* bool -> int *)
-fun int_for_bool b = if b then 1 else 0
+fun int_from_bool b = if b then 1 else 0
 (* int -> int -> int *)
 fun nat_minus i j = if i > j then i - j else 0
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -83,7 +83,7 @@
   else
     let
       fun get_specs ts = map_filter (fn t =>
-        TermGraph.get_node gr t |>
+        Term_Graph.get_node gr t |>
         (fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))
         ts
       val _ = print_step options ("Preprocessing scc of " ^
@@ -123,12 +123,12 @@
     val _ = print_step options "Fetching definitions from theory..."
     val gr = Output.cond_timeit (!Quickcheck.timing) "preprocess-obtain graph"
           (fn () => Predicate_Compile_Data.obtain_specification_graph options thy t
-          |> (fn gr => TermGraph.subgraph (member (op =) (TermGraph.all_succs gr [t])) gr))
+          |> (fn gr => Term_Graph.subgraph (member (op =) (Term_Graph.all_succs gr [t])) gr))
     val _ = if !present_graph then Predicate_Compile_Data.present_graph gr else ()
   in
     Output.cond_timeit (!Quickcheck.timing) "preprocess-process"
       (fn () => (fold_rev (preprocess_strong_conn_constnames options gr)
-        (TermGraph.strong_conn gr) thy))
+        (Term_Graph.strong_conn gr) thy))
   end
 
 fun extract_options (((expected_modes, proposed_modes), (compilation, raw_options)), const) =
@@ -151,7 +151,7 @@
       skip_proof = chk "skip_proof",
       function_flattening = not (chk "no_function_flattening"),
       fail_safe_function_flattening = false,
-      no_topmost_reordering = false,
+      no_topmost_reordering = (chk "no_topmost_reordering"),
       no_higher_order_predicate = [],
       inductify = chk "inductify",
       compilation = compilation
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -4,9 +4,7 @@
 Auxilary functions for predicate compiler.
 *)
 
-(* FIXME proper signature *)
-
-structure TermGraph = Graph(type key = term val ord = TermOrd.fast_term_ord);
+(* FIXME proper signature! *)
 
 structure Predicate_Compile_Aux =
 struct
@@ -556,7 +554,8 @@
 }
 
 val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes",
-  "show_mode_inference", "show_compilation", "skip_proof", "inductify", "no_function_flattening"]
+  "show_mode_inference", "show_compilation", "skip_proof", "inductify", "no_function_flattening",
+  "no_topmost_reordering"]
 
 val compilation_names = [("pred", Pred),
   (*("random", Random), ("depth_limited", Depth_Limited), ("annotated", Annotated),*)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -897,7 +897,7 @@
 
 (** mode analysis **)
 
-(* options for mode analysis  are: #use_random, #reorder_premises *)
+type mode_analysis_options = {use_random : bool, reorder_premises : bool, infer_pos_and_neg_modes : bool}
 
 fun is_constrt thy =
   let
@@ -1178,7 +1178,7 @@
   tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
     commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
 
-fun select_mode_prem mode_analysis_options thy pol (modes, (pos_modes, neg_modes)) vs ps =
+fun select_mode_prem (mode_analysis_options : mode_analysis_options) thy pol (modes, (pos_modes, neg_modes)) vs ps =
   let
     fun choose_mode_of_prem (Prem t) = partial_hd
         (sort (deriv_ord2 thy modes t) (all_derivations_of thy pos_modes vs t))
@@ -1194,7 +1194,7 @@
       SOME (hd ps, choose_mode_of_prem (hd ps))
   end
 
-fun check_mode_clause' mode_analysis_options thy param_vs (modes :
+fun check_mode_clause' (mode_analysis_options : mode_analysis_options) thy param_vs (modes :
   (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
   let
     val vTs = distinct (op =) (fold Term.add_frees (map term_of_prem ps) (fold Term.add_frees ts []))
@@ -2933,7 +2933,7 @@
             "computed values: " ^ Syntax.string_of_term ctxt elems ^ "\n")
   in
     if k = ~1 orelse length ts < k then elems
-      else Const (@{const_name Set.union}, setT --> setT --> setT) $ elems $ cont
+      else Const (@{const_abbrev Set.union}, setT --> setT --> setT) $ elems $ cont
   end;
 
 fun values_cmd print_modes param_user_modes options k raw_t state =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -14,9 +14,9 @@
   
   val get_specification : Predicate_Compile_Aux.options -> theory -> term -> thm list
   val obtain_specification_graph :
-    Predicate_Compile_Aux.options -> theory -> term -> thm list TermGraph.T
+    Predicate_Compile_Aux.options -> theory -> term -> thm list Term_Graph.T
     
-  val present_graph : thm list TermGraph.T -> unit
+  val present_graph : thm list Term_Graph.T -> unit
   val normalize_equation : theory -> thm -> thm
 end;
 
@@ -257,7 +257,7 @@
 fun obtain_specification_graph options thy t =
   let
     fun is_nondefining_const (c, T) = member (op =) logic_operator_names c
-    fun has_code_pred_intros (c, T) = is_some (try (Predicate_Compile_Core.intros_of thy) c)
+    fun has_code_pred_intros (c, T) = can (Predicate_Compile_Core.intros_of thy) c
     fun case_consts (c, T) = is_some (Datatype.info_of_case thy c)
     fun is_datatype_constructor (c, T) = is_some (Datatype.info_of_constr thy (c, T))
     fun defiants_of specs =
@@ -272,14 +272,22 @@
       |> map Const
       (*
       |> filter is_defining_constname*)
-    fun extend t =
-      let
-        val specs = get_specification options thy t
-          (*|> Predicate_Compile_Set.unfold_set_notation*)
-        (*val _ = print_specification options thy constname specs*)
-      in (specs, defiants_of specs) end;
+    fun extend t gr =
+      if can (Term_Graph.get_node gr) t then gr
+      else
+        let
+          val specs = get_specification options thy t
+            (*|> Predicate_Compile_Set.unfold_set_notation*)
+          (*val _ = print_specification options thy constname specs*)
+          val us = defiants_of specs
+        in
+          gr
+          |> Term_Graph.new_node (t, specs)
+          |> fold extend us
+          |> fold (fn u => Term_Graph.add_edge (t, u)) us
+        end
   in
-    TermGraph.extend extend t TermGraph.empty
+    extend t Term_Graph.empty
   end;
 
 
@@ -288,11 +296,11 @@
     fun eq_cname (Const (c1, _), Const (c2, _)) = (c1 = c2)
     fun string_of_const (Const (c, _)) = c
       | string_of_const _ = error "string_of_const: unexpected term"
-    val constss = TermGraph.strong_conn gr;
+    val constss = Term_Graph.strong_conn gr;
     val mapping = Termtab.empty |> fold (fn consts => fold (fn const =>
       Termtab.update (const, consts)) consts) constss;
     fun succs consts = consts
-      |> maps (TermGraph.imm_succs gr)
+      |> maps (Term_Graph.imm_succs gr)
       |> subtract eq_cname consts
       |> map (the o Termtab.lookup mapping)
       |> distinct (eq_list eq_cname);
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -22,7 +22,8 @@
 structure Fun_Pred = Theory_Data
 (
   type T = (term * term) Item_Net.T;
-  val empty = Item_Net.init (op aconv o pairself fst) (single o fst);
+  val empty = Item_Net.init ((op aconv o pairself fst) : (term * term) * (term * term) -> bool)
+    (single o fst);
   val extend = I;
   val merge = Item_Net.merge;
 )
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -23,7 +23,7 @@
 fun datatype_names_of_case_name thy case_name =
   map (#1 o #2) (#descr (the (Datatype_Data.info_of_case thy case_name)))
 
-fun make_case_rewrites new_type_names descr sorts thy =
+fun make_case_distribs new_type_names descr sorts thy =
   let
     val case_combs = Datatype_Prop.make_case_combs new_type_names descr sorts thy "f";
     fun make comb =
@@ -58,7 +58,7 @@
     val typ_names = the_default [Tcon] (#alt_names info)
   in
     map (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop)
-      (make_case_rewrites typ_names [descr] sorts thy)
+      (make_case_distribs typ_names [descr] sorts thy)
   end
 
 fun instantiated_case_rewrites thy Tcon =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -10,7 +10,8 @@
   val test_ref :
     ((unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option) Unsynchronized.ref
   val tracing : bool Unsynchronized.ref;
-  val quickcheck_compile_term : bool -> bool -> Proof.context -> term -> int -> term list option
+  val quickcheck_compile_term : bool -> bool -> 
+    Proof.context -> bool -> term -> int -> term list option * (bool list * bool);
 (*  val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *)
   val quiet : bool Unsynchronized.ref;
   val nrandom : int Unsynchronized.ref;
@@ -216,7 +217,7 @@
     fun try' j =
       if j <= i then
         let
-          val _ = priority ("Executing depth " ^ string_of_int j)
+          val _ = if quiet then () else priority ("Executing depth " ^ string_of_int j)
         in
           case f j handle Match => (if quiet then ()
              else warning "Exception Match raised during quickcheck"; NONE)
@@ -230,11 +231,12 @@
 
 (* quickcheck interface functions *)
 
-fun compile_term' options ctxt t =
+fun compile_term' options ctxt report t =
   let
     val c = compile_term options ctxt t
+    val dummy_report = ([], false)
   in
-    (fn size => try_upto (!quiet) (c size (!nrandom)) (!depth))
+    fn size => (try_upto (!quiet) (c size (!nrandom)) (!depth), dummy_report)
   end
 
 fun quickcheck_compile_term function_flattening fail_safe_function_flattening ctxt t =
--- a/src/HOL/Tools/Qelim/cooper.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Qelim/cooper.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -16,7 +16,7 @@
 
 exception COOPER of string * exn;
 fun simp_thms_conv ctxt =
-  Simplifier.rewrite (Simplifier.context ctxt HOL_basic_ss addsimps simp_thms);
+  Simplifier.rewrite (Simplifier.context ctxt HOL_basic_ss addsimps @{thms simp_thms});
 val FWD = Drule.implies_elim_list;
 
 val true_tm = @{cterm "True"};
@@ -514,8 +514,8 @@
 
 local
  val pcv = Simplifier.rewrite
-     (HOL_basic_ss addsimps (simp_thms @ (List.take(ex_simps,4))
-                      @ [not_all,all_not_ex, ex_disj_distrib]))
+     (HOL_basic_ss addsimps (@{thms simp_thms} @ List.take(@{thms ex_simps}, 4)
+                      @ [not_all, all_not_ex, @{thm ex_disj_distrib}]))
  val postcv = Simplifier.rewrite presburger_ss
  fun conv ctxt p =
   let val _ = ()
--- a/src/HOL/Tools/Qelim/langford.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Qelim/langford.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -210,7 +210,7 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
--- a/src/HOL/Tools/Qelim/presburger.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Qelim/presburger.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -103,13 +103,13 @@
  let 
    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
    fun gen x t = Thm.capply (all (ctyp_of_term x)) (Thm.cabs x t)
-   val ts = sort (fn (a,b) => TermOrd.fast_term_ord (term_of a, term_of b)) (f p)
+   val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
    val p' = fold_rev gen ts p
  in implies_intr p' (implies_elim st (fold forall_elim ts (assume p'))) end));
 
 local
 val ss1 = comp_ss
-  addsimps simp_thms @ [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] 
+  addsimps @{thms simp_thms} @ [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] 
       @ map (fn r => r RS sym) 
         [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, 
          @{thm "zmult_int"}]
@@ -120,7 +120,7 @@
             @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, 
             @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}]
   addcongs [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
-val div_mod_ss = HOL_basic_ss addsimps simp_thms 
+val div_mod_ss = HOL_basic_ss addsimps @{thms simp_thms}
   @ map (symmetric o mk_meta_eq) 
     [@{thm "dvd_eq_mod_eq_0"},
      @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"}, 
--- a/src/HOL/Tools/Qelim/qelim.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Qelim/qelim.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -55,9 +55,10 @@
 (* Instantiation of some parameter for most common cases *)
 
 local
-val pcv = Simplifier.rewrite
-                 (HOL_basic_ss addsimps (simp_thms @ ex_simps @ all_simps)
-                     @ [@{thm "all_not_ex"}, not_all, ex_disj_distrib]);
+val pcv =
+  Simplifier.rewrite
+    (HOL_basic_ss addsimps (@{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps} @
+        [@{thm all_not_ex}, not_all, @{thm ex_disj_distrib}]));
 
 in fun standard_qelim_conv atcv ncv qcv p =
       gen_qelim_conv pcv pcv pcv cons (Thm.add_cterm_frees p []) atcv ncv qcv p
--- a/src/HOL/Tools/Quotient/quotient_term.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Quotient/quotient_term.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -265,7 +265,7 @@
   | is_eq _ = false
 
 fun mk_rel_compose (trm1, trm2) =
-  Const (@{const_name "rel_conj"}, dummyT) $ trm1 $ trm2
+  Const (@{const_abbrev "rel_conj"}, dummyT) $ trm1 $ trm2
 
 fun get_relmap ctxt s =
 let
--- a/src/HOL/Tools/Quotient/quotient_typ.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -7,6 +7,9 @@
 
 signature QUOTIENT_TYPE =
 sig
+  val add_quotient_type: ((string list * binding * mixfix) * (typ * term)) * thm
+    -> Proof.context -> (thm * thm) * local_theory
+
   val quotient_type: ((string list * binding * mixfix) * (typ * term)) list
     -> Proof.context -> Proof.state
 
@@ -128,7 +131,7 @@
 
 
 (* main function for constructing a quotient type *)
-fun mk_quotient_type (((vs, qty_name, mx), (rty, rel)), equiv_thm) lthy =
+fun add_quotient_type (((vs, qty_name, mx), (rty, rel)), equiv_thm) lthy =
 let
   (* generates the typedef *)
   val ((qty_full_name, typedef_info), lthy1) = typedef_make (vs, qty_name, mx, rel, rty) lthy
@@ -264,7 +267,7 @@
   val goals = map (mk_goal o snd) quot_list
 
   fun after_qed thms lthy =
-    fold_map mk_quotient_type (quot_list ~~ thms) lthy |> snd
+    fold_map add_quotient_type (quot_list ~~ thms) lthy |> snd
 in
   theorem after_qed goals lthy
 end
@@ -296,7 +299,7 @@
 val quotspec_parser =
     OuterParse.and_list1
      ((OuterParse.type_args -- OuterParse.binding) --
-        OuterParse.opt_infix -- (OuterParse.$$$ "=" |-- OuterParse.typ) --
+        OuterParse.opt_mixfix -- (OuterParse.$$$ "=" |-- OuterParse.typ) --
          (OuterParse.$$$ "/" |-- OuterParse.term))
 
 val _ = OuterKeyword.keyword "/"
--- a/src/HOL/Tools/TFL/post.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/TFL/post.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -125,7 +125,7 @@
 
 (*lcp: curry the predicate of the induction rule*)
 fun curry_rule rl =
-  SplitRule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl;
+  Split_Rule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl;
 
 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
 val meta_outer =
--- a/src/HOL/Tools/arith_data.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/arith_data.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -120,7 +120,7 @@
   in fn ss => ALLGOALS (simp_tac (Simplifier.inherit_context ss ss0)) end;
 
 fun simplify_meta_eq rules =
-  let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules
+  let val ss0 = HOL_basic_ss addeqcongs [@{thm eq_cong2}] addsimps rules
   in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end
 
 fun trans_tac NONE  = all_tac
--- a/src/HOL/Tools/inductive.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/inductive.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -183,7 +183,7 @@
   in
     case concl_of thm of
       Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
-    | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
+    | _ $ (Const (@{const_name "op ="}, _) $ _ $ _) => eq2mono thm
     | _ $ (Const (@{const_name Orderings.less_eq}, _) $ _ $ _) =>
       dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
         (resolve_tac [@{thm le_funI}, @{thm le_boolI'}])) thm))
@@ -300,7 +300,7 @@
       else err_in_prem ctxt err_name rule prem "Non-atomic premise";
   in
     (case concl of
-       Const ("Trueprop", _) $ t =>
+       Const (@{const_name Trueprop}, _) $ t =>
          if head_of t mem cs then
            (check_ind (err_in_rule ctxt err_name rule') t;
             List.app check_prem (prems ~~ aprems))
--- a/src/HOL/Tools/inductive_codegen.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/inductive_codegen.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -8,7 +8,8 @@
 sig
   val add : string option -> int option -> attribute
   val test_fn : (int * int * int -> term list option) Unsynchronized.ref
-  val test_term: Proof.context -> term -> int -> term list option
+  val test_term:
+    Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
   val setup : theory -> theory
   val quickcheck_setup : theory -> theory
 end;
@@ -51,12 +52,13 @@
     fun thyname_of s = (case optmod of
       NONE => Codegen.thyname_of_const thy s | SOME s => s);
   in (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
-      SOME (Const ("op =", _), [t, _]) => (case head_of t of
-        Const (s, _) =>
-          CodegenData.put {intros = intros, graph = graph,
-             eqns = eqns |> Symtab.map_default (s, [])
-               (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
-      | _ => (warn thm; thy))
+      SOME (Const (@{const_name "op ="}, _), [t, _]) =>
+        (case head_of t of
+          Const (s, _) =>
+            CodegenData.put {intros = intros, graph = graph,
+               eqns = eqns |> Symtab.map_default (s, [])
+                 (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
+        | _ => (warn thm; thy))
     | SOME (Const (s, _), _) =>
         let
           val cs = fold Term.add_const_names (Thm.prems_of thm) [];
@@ -186,7 +188,7 @@
         end)) (AList.lookup op = modes name)
 
   in (case strip_comb t of
-      (Const ("op =", Type (_, [T, _])), _) =>
+      (Const (@{const_name "op ="}, Type (_, [T, _])), _) =>
         [Mode ((([], [1]), false), [1], []), Mode ((([], [2]), false), [2], [])] @
         (if is_eqT T then [Mode ((([], [1, 2]), false), [1, 2], [])] else [])
     | (Const (name, _), args) => the_default default (mk_modes name args)
@@ -577,17 +579,20 @@
       fun get_nparms s = if s mem names then SOME nparms else
         Option.map #3 (get_clauses thy s);
 
-      fun dest_prem (_ $ (Const ("op :", _) $ t $ u)) = Prem ([t], Envir.beta_eta_contract u, true)
-        | dest_prem (_ $ ((eq as Const ("op =", _)) $ t $ u)) = Prem ([t, u], eq, false)
+      fun dest_prem (_ $ (Const (@{const_name "op :"}, _) $ t $ u)) =
+            Prem ([t], Envir.beta_eta_contract u, true)
+        | dest_prem (_ $ ((eq as Const (@{const_name "op ="}, _)) $ t $ u)) =
+            Prem ([t, u], eq, false)
         | dest_prem (_ $ t) =
             (case strip_comb t of
-               (v as Var _, ts) => if v mem args then Prem (ts, v, false) else Sidecond t
-             | (c as Const (s, _), ts) => (case get_nparms s of
-                 NONE => Sidecond t
-               | SOME k =>
-                   let val (ts1, ts2) = chop k ts
-                   in Prem (ts2, list_comb (c, ts1), false) end)
-             | _ => Sidecond t);
+              (v as Var _, ts) => if v mem args then Prem (ts, v, false) else Sidecond t
+            | (c as Const (s, _), ts) =>
+                (case get_nparms s of
+                  NONE => Sidecond t
+                | SOME k =>
+                    let val (ts1, ts2) = chop k ts
+                    in Prem (ts2, list_comb (c, ts1), false) end)
+            | _ => Sidecond t);
 
       fun add_clause intr (clauses, arities) =
         let
@@ -600,7 +605,7 @@
           (AList.update op = (name, these (AList.lookup op = clauses name) @
              [(ts2, prems)]) clauses,
            AList.update op = (name, (map (fn U => (case strip_type U of
-                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
+                 (Rs as _ :: _, @{typ bool}) => SOME (length Rs)
                | _ => NONE)) Ts,
              length Us)) arities)
         end;
@@ -632,7 +637,7 @@
     val (ts1, ts2) = chop k ts;
     val u = list_comb (Const (s, T), ts1);
 
-    fun mk_mode (Const ("dummy_pattern", _)) ((ts, mode), i) = ((ts, mode), i + 1)
+    fun mk_mode (Const (@{const_name dummy_pattern}, _)) ((ts, mode), i) = ((ts, mode), i + 1)
       | mk_mode t ((ts, mode), i) = ((ts @ [t], mode @ [i]), i + 1);
 
     val module' = if_library thyname module;
@@ -715,7 +720,7 @@
   end;
 
 fun inductive_codegen thy defs dep module brack t gr  = (case strip_comb t of
-    (Const ("Collect", _), [u]) =>
+    (Const (@{const_name Collect}, _), [u]) =>
       let val (r, Ts, fs) = HOLogic.strip_psplits u
       in case strip_comb r of
           (Const (s, T), ts) =>
@@ -805,7 +810,7 @@
   Config.declare false "ind_quickcheck_size_offset" (Config.Int 0);
 val size_offset = Config.int size_offset_value;
 
-fun test_term ctxt t =
+fun test_term ctxt report t =
   let
     val thy = ProofContext.theory_of ctxt;
     val (xs, p) = strip_abs t;
@@ -848,9 +853,10 @@
     val start = Config.get ctxt depth_start;
     val offset = Config.get ctxt size_offset;
     val test_fn' = !test_fn;
-    fun test k = deepen bound (fn i =>
+    val dummy_report = ([], false)
+    fun test k = (deepen bound (fn i =>
       (priority ("Search depth: " ^ string_of_int i);
-       test_fn' (i, values, k+offset))) start;
+       test_fn' (i, values, k+offset))) start, dummy_report);
   in test end;
 
 val quickcheck_setup =
--- a/src/HOL/Tools/inductive_realizer.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/inductive_realizer.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -21,7 +21,7 @@
     [name] => name
   | _ => error ("name_of_thm: bad proof of theorem\n" ^ Display.string_of_thm_without_context thm));
 
-val all_simps = map (symmetric o mk_meta_eq) (thms "HOL.all_simps");
+val all_simps = map (symmetric o mk_meta_eq) @{thms HOL.all_simps};
 
 fun prf_of thm =
   let
@@ -57,7 +57,7 @@
 
 fun relevant_vars prop = List.foldr (fn
       (Var ((a, i), T), vs) => (case strip_type T of
-        (_, Type (s, _)) => if s mem ["bool"] then (a, T) :: vs else vs
+        (_, Type (s, _)) => if s mem [@{type_name bool}] then (a, T) :: vs else vs
       | _ => vs)
     | (_, vs) => vs) [] (OldTerm.term_vars prop);
 
@@ -150,9 +150,10 @@
 
     fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
       | is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
-      | is_meta (Const ("Trueprop", _) $ t) = (case head_of t of
-          Const (s, _) => can (Inductive.the_inductive ctxt) s
-        | _ => true)
+      | is_meta (Const (@{const_name Trueprop}, _) $ t) =
+          (case head_of t of
+            Const (s, _) => can (Inductive.the_inductive ctxt) s
+          | _ => true)
       | is_meta _ = false;
 
     fun fun_of ts rts args used (prem :: prems) =
@@ -174,7 +175,7 @@
                     (Free (r, U) :: Free (x, T) :: args) (x :: r :: used) prems'
                 end
               else (case strip_type T of
-                  (Ts, Type ("*", [T1, T2])) =>
+                  (Ts, Type (@{type_name "*"}, [T1, T2])) =>
                     let
                       val fx = Free (x, Ts ---> T1);
                       val fr = Free (r, Ts ---> T2);
@@ -211,8 +212,9 @@
       let
         val fs = map (fn (rule, (ivs, intr)) =>
           fun_of_prem thy rsets vs params rule ivs intr) (prems ~~ intrs)
-      in if dummy then Const ("HOL.default_class.default",
-          HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs
+      in
+        if dummy then Const (@{const_name default},
+            HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs
         else fs
       end) (premss ~~ dummies);
     val frees = fold Term.add_frees fs [];
@@ -439,7 +441,7 @@
         val r = if null Ps then Extraction.nullt
           else list_abs (map (pair "x") Ts, list_comb (Const (case_name, T),
             (if dummy then
-               [Abs ("x", HOLogic.unitT, Const ("HOL.default_class.default", body_type T))]
+               [Abs ("x", HOLogic.unitT, Const (@{const_name default}, body_type T))]
              else []) @
             map reorder2 (intrs ~~ (length prems - 1 downto 0)) @
             [Bound (length prems)]));
@@ -508,7 +510,7 @@
 val setup =
   Attrib.setup @{binding ind_realizer}
     ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |--
-      Scan.option (Scan.lift (Args.colon) |-- Scan.repeat1 Args.const))) >> rlz_attrib)
+      Scan.option (Scan.lift (Args.colon) |-- Scan.repeat1 (Args.const true)))) >> rlz_attrib)
     "add realizers for inductive set";
 
 end;
--- a/src/HOL/Tools/inductive_set.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/inductive_set.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -33,10 +33,10 @@
 
 val collect_mem_simproc =
   Simplifier.simproc (theory "Set") "Collect_mem" ["Collect t"] (fn thy => fn ss =>
-    fn S as Const ("Collect", Type ("fun", [_, T])) $ t =>
+    fn S as Const (@{const_name Collect}, Type ("fun", [_, T])) $ t =>
          let val (u, _, ps) = HOLogic.strip_psplits t
          in case u of
-           (c as Const ("op :", _)) $ q $ S' =>
+           (c as Const (@{const_name "op :"}, _)) $ q $ S' =>
              (case try (HOLogic.strip_ptuple ps) q of
                 NONE => NONE
               | SOME ts =>
@@ -74,18 +74,20 @@
         in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
           (p (fold (Logic.all o Var) vs t) f)
         end;
-      fun mkop "op &" T x = SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x)
-        | mkop "op |" T x = SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x)
+      fun mkop @{const_name "op &"} T x =
+            SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x)
+        | mkop @{const_name "op |"} T x =
+            SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x)
         | mkop _ _ _ = NONE;
       fun mk_collect p T t =
         let val U = HOLogic.dest_setT T
         in HOLogic.Collect_const U $
           HOLogic.mk_psplits (HOLogic.flat_tuple_paths p) U HOLogic.boolT t
         end;
-      fun decomp (Const (s, _) $ ((m as Const ("op :",
+      fun decomp (Const (s, _) $ ((m as Const (@{const_name "op :"},
             Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) =
               mkop s T (m, p, S, mk_collect p T (head_of u))
-        | decomp (Const (s, _) $ u $ ((m as Const ("op :",
+        | decomp (Const (s, _) $ u $ ((m as Const (@{const_name "op :"},
             Type (_, [_, Type (_, [T, _])]))) $ p $ S)) =
               mkop s T (m, p, mk_collect p T (head_of u), S)
         | decomp _ = NONE;
@@ -263,13 +265,13 @@
 
 fun add ctxt thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
   case prop_of thm of
-    Const ("Trueprop", _) $ (Const ("op =", Type (_, [T, _])) $ lhs $ rhs) =>
+    Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, Type (_, [T, _])) $ lhs $ rhs) =>
       (case body_type T of
-         Type ("bool", []) =>
+         @{typ bool} =>
            let
              val thy = Context.theory_of ctxt;
              fun factors_of t fs = case strip_abs_body t of
-                 Const ("op :", _) $ u $ S =>
+                 Const (@{const_name "op :"}, _) $ u $ S =>
                    if is_Free S orelse is_Var S then
                      let val ps = HOLogic.flat_tuple_paths u
                      in (SOME ps, (S, ps) :: fs) end
@@ -279,7 +281,7 @@
              val (pfs, fs) = fold_map factors_of ts [];
              val ((h', ts'), fs') = (case rhs of
                  Abs _ => (case strip_abs_body rhs of
-                     Const ("op :", _) $ u $ S =>
+                     Const (@{const_name "op :"}, _) $ u $ S =>
                        (strip_comb S, SOME (HOLogic.flat_tuple_paths u))
                    | _ => error "member symbol on right-hand side expected")
                | _ => (strip_comb rhs, NONE))
@@ -412,7 +414,7 @@
     val {set_arities, pred_arities, to_pred_simps, ...} =
       PredSetConvData.get (Context.Proof lthy);
     fun infer (Abs (_, _, t)) = infer t
-      | infer (Const ("op :", _) $ t $ u) =
+      | infer (Const (@{const_name "op :"}, _) $ t $ u) =
           infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u)
       | infer (t $ u) = infer t #> infer u
       | infer _ = I;
--- a/src/HOL/Tools/lin_arith.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/lin_arith.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -42,7 +42,7 @@
 val not_lessD = @{thm linorder_not_less} RS iffD1;
 val not_leD = @{thm linorder_not_le} RS iffD1;
 
-fun mk_Eq thm = thm RS Eq_FalseI handle THM _ => thm RS Eq_TrueI;
+fun mk_Eq thm = thm RS @{thm Eq_FalseI} handle THM _ => thm RS @{thm Eq_TrueI};
 
 val mk_Trueprop = HOLogic.mk_Trueprop;
 
@@ -703,8 +703,8 @@
   val nnf_simpset =
     empty_ss setmkeqTrue mk_eq_True
     setmksimps (mksimps mksimps_pairs)
-    addsimps [imp_conv_disj, iff_conv_conj_imp, de_Morgan_disj, de_Morgan_conj,
-      not_all, not_ex, not_not]
+    addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj},
+      @{thm de_Morgan_conj}, not_all, not_ex, not_not]
   fun prem_nnf_tac ss = full_simp_tac (Simplifier.inherit_context ss nnf_simpset)
 in
 
@@ -823,7 +823,7 @@
       addsimprocs [ab_group_add_cancel.sum_conv, ab_group_add_cancel.rel_conv]
        (*abel_cancel helps it work in abstract algebraic domains*)
       addsimprocs Nat_Arith.nat_cancel_sums_add
-      addcongs [if_weak_cong],
+      addcongs [@{thm if_weak_cong}],
     number_of = number_of}) #>
   add_discrete_type @{type_name nat};
 
--- a/src/HOL/Tools/meson.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/meson.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -517,10 +517,10 @@
   nnf_ss also includes the one-point simprocs,
   which are needed to avoid the various one-point theorems from generating junk clauses.*)
 val nnf_simps =
-     [simp_implies_def, Ex1_def, Ball_def, Bex_def, if_True,
-      if_False, if_cancel, if_eq_cancel, cases_simp];
+  [@{thm simp_implies_def}, @{thm Ex1_def}, @{thm Ball_def},@{thm  Bex_def}, @{thm if_True},
+    @{thm if_False}, @{thm if_cancel}, @{thm if_eq_cancel}, @{thm cases_simp}];
 val nnf_extra_simps =
-      thms"split_ifs" @ ex_simps @ all_simps @ simp_thms;
+  @{thms split_ifs} @ @{thms ex_simps} @ @{thms all_simps} @ @{thms simp_thms};
 
 val nnf_ss =
   HOL_basic_ss addsimps nnf_extra_simps
@@ -685,7 +685,7 @@
 
 (*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing
   double-negations.*)
-val negate_head = rewrite_rule [atomize_not, not_not RS eq_reflection];
+val negate_head = rewrite_rule [@{thm atomize_not}, not_not RS eq_reflection];
 
 (*Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*)
 fun select_literal i cl = negate_head (make_last i cl);
--- a/src/HOL/Tools/nat_numeral_simprocs.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -80,7 +80,7 @@
 
 (*Express t as a product of (possibly) a numeral with other factors, sorted*)
 fun dest_coeff t =
-    let val ts = sort TermOrd.term_ord (dest_prod t)
+    let val ts = sort Term_Ord.term_ord (dest_prod t)
         val (n, _, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, one, ts)
     in (n, mk_prod ts') end;
--- a/src/HOL/Tools/numeral_simprocs.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/numeral_simprocs.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -74,7 +74,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort TermOrd.term_ord (dest_prod t)
+    let val ts = sort Term_Ord.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod (Term.fastype_of t) ts') end;
@@ -124,12 +124,12 @@
     EQUAL => int_ord (Int.sign i, Int.sign j) 
   | ord => ord);
 
-(*This resembles TermOrd.term_ord, but it puts binary numerals before other
+(*This resembles Term_Ord.term_ord, but it puts binary numerals before other
   non-atomic terms.*)
 local open Term 
 in 
 fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) =
-      (case numterm_ord (t, u) of EQUAL => TermOrd.typ_ord (T, U) | ord => ord)
+      (case numterm_ord (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
   | numterm_ord
      (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
      num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
@@ -139,7 +139,7 @@
       (case int_ord (size_of_term t, size_of_term u) of
         EQUAL =>
           let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
-            (case TermOrd.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
+            (case Term_Ord.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
           end
       | ord => ord)
 and numterms_ord (ts, us) = list_ord numterm_ord (ts, us)
--- a/src/HOL/Tools/numeral_syntax.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/numeral_syntax.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -69,7 +69,7 @@
 
 in
 
-fun numeral_tr' show_sorts (*"number_of"*) (Type ("fun", [_, T])) (t :: ts) =  (* FIXME @{type_syntax} *)
+fun numeral_tr' show_sorts (*"number_of"*) (Type (@{type_syntax fun}, [_, T])) (t :: ts) =
       let val t' =
         if not (! show_types) andalso can Term.dest_Type T then syntax_numeral t
         else Syntax.const Syntax.constrainC $ syntax_numeral t $ Syntax.term_of_typ show_sorts T
--- a/src/HOL/Tools/quickcheck_generators.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/quickcheck_generators.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -15,8 +15,11 @@
     -> string list -> string list * string list -> typ list * typ list
     -> term list * (term * term) list
   val ensure_random_datatype: Datatype.config -> string list -> theory -> theory
-  val compile_generator_expr: theory -> term -> int -> term list option
+  val compile_generator_expr:
+    theory -> bool -> term -> int -> term list option * (bool list * bool)
   val eval_ref: (unit -> int -> seed -> term list option * seed) option Unsynchronized.ref
+  val eval_report_ref:
+    (unit -> int -> seed -> (term list option * (bool list * bool)) * seed) option Unsynchronized.ref
   val setup: theory -> theory
 end;
 
@@ -350,6 +353,10 @@
     (unit -> int -> int * int -> term list option * (int * int)) option Unsynchronized.ref =
   Unsynchronized.ref NONE;
 
+val eval_report_ref :
+    (unit -> int -> seed -> (term list option * (bool list * bool)) * seed) option Unsynchronized.ref =
+  Unsynchronized.ref NONE;
+
 val target = "Quickcheck";
 
 fun mk_generator_expr thy prop Ts =
@@ -360,7 +367,7 @@
     val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
     val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
     val check = @{term "If :: bool => term list option => term list option => term list option"}
-      $ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option "} $ terms);
+      $ result $ @{term "None :: term list option"} $ (@{term "Some :: term list => term list option"} $ terms);
     val return = @{term "Pair :: term list option => Random.seed => term list option * Random.seed"};
     fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
     fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"});
@@ -375,13 +382,69 @@
       (Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i);
   in Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check)) end;
 
-fun compile_generator_expr thy t =
+fun mk_reporting_generator_expr thy prop Ts =
+  let
+    val bound_max = length Ts - 1;
+    val bounds = map_index (fn (i, ty) =>
+      (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
+    fun strip_imp (Const("op -->",_) $ A $ B) = apfst (cons A) (strip_imp B)
+      | strip_imp A = ([], A)
+    val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds);
+    val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds)
+    val (assms, concl) = strip_imp prop'
+    val return =
+      @{term "Pair :: term list option * (bool list * bool) => Random.seed => (term list option * (bool list * bool)) * Random.seed"};
+    fun mk_assms_report i =
+      HOLogic.mk_prod (@{term "None :: term list option"},
+        HOLogic.mk_prod (HOLogic.mk_list @{typ "bool"}
+          (replicate i @{term "True"} @ replicate (length assms - i) @{term "False"}),
+        @{term "False"}))
+    fun mk_concl_report b =
+      HOLogic.mk_prod (HOLogic.mk_list @{typ "bool"} (replicate (length assms) @{term "True"}),
+        if b then @{term True} else @{term False})
+    val If =
+      @{term "If :: bool => term list option * (bool list * bool) => term list option * (bool list * bool) => term list option * (bool list * bool)"}
+    val concl_check = If $ concl $
+      HOLogic.mk_prod (@{term "None :: term list option"}, mk_concl_report true) $
+      HOLogic.mk_prod (@{term "Some :: term list  => term list option"} $ terms, mk_concl_report false)
+    val check = fold_rev (fn (i, assm) => fn t => If $ assm $ t $ mk_assms_report i)
+      (map_index I assms) concl_check
+    fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
+    fun mk_termtyp T = HOLogic.mk_prodT (T, @{typ "unit => term"});
+    fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
+      liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
+    fun mk_split T = Sign.mk_const thy
+      (@{const_name split}, [T, @{typ "unit => term"},
+        liftT @{typ "term list option * (bool list * bool)"} @{typ Random.seed}]);
+    fun mk_scomp_split T t t' =
+      mk_scomp (mk_termtyp T) @{typ "term list option * (bool list * bool)"} @{typ Random.seed} t
+        (mk_split T $ Abs ("", T, Abs ("", @{typ "unit => term"}, t')));
+    fun mk_bindclause (_, _, i, T) = mk_scomp_split T
+      (Sign.mk_const thy (@{const_name Quickcheck.random}, [T]) $ Bound i);
+  in
+    Abs ("n", @{typ code_numeral}, fold_rev mk_bindclause bounds (return $ check))
+  end
+
+fun compile_generator_expr thy report t =
   let
     val Ts = (map snd o fst o strip_abs) t;
-    val t' = mk_generator_expr thy t Ts;
-    val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
-      (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
-  in compile #> Random_Engine.run end;
+  in
+    if report then
+      let
+        val t' = mk_reporting_generator_expr thy t Ts;
+        val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_report_ref", eval_report_ref)
+          (fn proc => fn g => fn s => g s #>> ((apfst o Option.map o map) proc)) thy t' [];
+      in
+        compile #> Random_Engine.run
+      end
+    else
+      let
+        val t' = mk_generator_expr thy t Ts;
+        val compile = Code_Eval.eval (SOME target) ("Quickcheck_Generators.eval_ref", eval_ref)
+          (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+        val dummy_report = ([], false)
+      in fn s => ((compile #> Random_Engine.run) s, dummy_report) end
+  end;
 
 
 (** setup **)
--- a/src/HOL/Tools/record.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/record.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -762,8 +762,7 @@
     mk_ext (field_types_tr t)
   end;
 
-(* FIXME @{type_syntax} *)
-fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const "Product_Type.unit") ctxt t
+fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const @{type_syntax unit}) ctxt t
   | record_type_tr _ ts = raise TERM ("record_type_tr", ts);
 
 fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t
@@ -1101,7 +1100,7 @@
 fun get_accupd_simps thy term defset =
   let
     val (acc, [body]) = strip_comb term;
-    val upd_funs = sort_distinct TermOrd.fast_term_ord (get_upd_funs body);
+    val upd_funs = sort_distinct Term_Ord.fast_term_ord (get_upd_funs body);
     fun get_simp upd =
       let
         (* FIXME fresh "f" (!?) *)
@@ -1417,7 +1416,7 @@
         | name =>
             (case get_equalities thy name of
               NONE => NONE
-            | SOME thm => SOME (thm RS Eq_TrueI)))
+            | SOME thm => SOME (thm RS @{thm Eq_TrueI})))
       | _ => NONE));
 
 
@@ -1463,7 +1462,7 @@
         fun prove prop =
           prove_global true thy [] prop
             (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy)
-                addsimps simp_thms addsimprocs [record_split_simproc (K ~1)]) 1);
+                addsimps @{thms simp_thms} addsimprocs [record_split_simproc (K ~1)]) 1);
 
         fun mkeq (lr, Teq, (sel, Tsel), x) i =
           if is_selector thy sel then
--- a/src/HOL/Tools/sat_funcs.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/sat_funcs.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -320,7 +320,7 @@
 	(* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *)
 	(* terms sorted in descending order, while only linear for terms      *)
 	(* sorted in ascending order                                          *)
-	val sorted_clauses = sort (TermOrd.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
+	val sorted_clauses = sort (Term_Ord.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
 	val _ = if !trace_sat then
 			tracing ("Sorted non-trivial clauses:\n" ^
 			  cat_lines (map (Syntax.string_of_term ctxt o Thm.term_of) sorted_clauses))
--- a/src/HOL/Tools/simpdata.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/simpdata.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -10,11 +10,11 @@
 structure Quantifier1 = Quantifier1Fun
 (struct
   (*abstract syntax*)
-  fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t)
+  fun dest_eq ((c as Const(@{const_name "op ="},_)) $ s $ t) = SOME (c, s, t)
     | dest_eq _ = NONE;
-  fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t)
+  fun dest_conj ((c as Const(@{const_name "op &"},_)) $ s $ t) = SOME (c, s, t)
     | dest_conj _ = NONE;
-  fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t)
+  fun dest_imp ((c as Const(@{const_name "op -->"},_)) $ s $ t) = SOME (c, s, t)
     | dest_imp _ = NONE;
   val conj = HOLogic.conj
   val imp  = HOLogic.imp
@@ -43,9 +43,9 @@
 
 fun mk_eq th = case concl_of th
   (*expects Trueprop if not == *)
-  of Const ("==",_) $ _ $ _ => th
-   | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq th
-   | _ $ (Const ("Not", _) $ _) => th RS @{thm Eq_FalseI}
+  of Const (@{const_name "=="},_) $ _ $ _ => th
+   | _ $ (Const (@{const_name "op ="}, _) $ _ $ _) => mk_meta_eq th
+   | _ $ (Const (@{const_name "Not"}, _) $ _) => th RS @{thm Eq_FalseI}
    | _ => th RS @{thm Eq_TrueI}
 
 fun mk_eq_True r =
@@ -57,7 +57,7 @@
 
 fun lift_meta_eq_to_obj_eq i st =
   let
-    fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
+    fun count_imp (Const (@{const_name HOL.simp_implies}, _) $ P $ Q) = 1 + count_imp Q
       | count_imp _ = 0;
     val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1)))
   in if j = 0 then @{thm meta_eq_to_obj_eq}
@@ -65,7 +65,7 @@
       let
         val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
         fun mk_simp_implies Q = fold_rev (fn R => fn S =>
-          Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Ps Q
+          Const (@{const_name HOL.simp_implies}, propT --> propT --> propT) $ R $ S) Ps Q
         val aT = TFree ("'a", HOLogic.typeS);
         val x = Free ("x", aT);
         val y = Free ("y", aT)
@@ -98,7 +98,7 @@
           else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm];
       in
         case concl_of thm
-          of Const ("Trueprop", _) $ p => (case head_of p
+          of Const (@{const_name Trueprop}, _) $ p => (case head_of p
             of Const (a, _) => (case AList.lookup (op =) pairs a
               of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm])
               | NONE => [thm])
@@ -159,9 +159,12 @@
   (Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())");
 
 val mksimps_pairs =
-  [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
-   ("All", [@{thm spec}]), ("True", []), ("False", []),
-   ("HOL.If", [@{thm if_bool_eq_conj} RS @{thm iffD1}])];
+ [(@{const_name "op -->"}, [@{thm mp}]),
+  (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}]),
+  (@{const_name All}, [@{thm spec}]),
+  (@{const_name True}, []),
+  (@{const_name False}, []),
+  (@{const_name If}, [@{thm if_bool_eq_conj} RS @{thm iffD1}])];
 
 val HOL_basic_ss =
   Simplifier.global_context @{theory} empty_ss
--- a/src/HOL/Tools/split_rule.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/split_rule.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -4,49 +4,34 @@
 Some tools for managing tupled arguments and abstractions in rules.
 *)
 
-signature BASIC_SPLIT_RULE =
+signature SPLIT_RULE =
 sig
+  val split_rule_var: term -> thm -> thm
+  val split_rule_goal: Proof.context -> string list list -> thm -> thm
   val split_rule: thm -> thm
   val complete_split_rule: thm -> thm
-end;
-
-signature SPLIT_RULE =
-sig
-  include BASIC_SPLIT_RULE
-  val split_rule_var: term -> thm -> thm
-  val split_rule_goal: Proof.context -> string list list -> thm -> thm
   val setup: theory -> theory
 end;
 
-structure SplitRule: SPLIT_RULE =
+structure Split_Rule: SPLIT_RULE =
 struct
 
-
-
-(** theory context references **)
-
-val split_conv = thm "split_conv";
-val fst_conv = thm "fst_conv";
-val snd_conv = thm "snd_conv";
-
-fun internal_split_const (Ta, Tb, Tc) =
-  Const ("Product_Type.internal_split", [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc);
-
-val internal_split_def = thm "internal_split_def";
-val internal_split_conv = thm "internal_split_conv";
-
-
-
 (** split rules **)
 
-val eval_internal_split = hol_simplify [internal_split_def] o hol_simplify [internal_split_conv];
+fun internal_split_const (Ta, Tb, Tc) =
+  Const (@{const_name Product_Type.internal_split},
+    [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc);
+
+val eval_internal_split =
+  hol_simplify [@{thm internal_split_def}] o hol_simplify [@{thm internal_split_conv}];
+
 val remove_internal_split = eval_internal_split o split_all;
 
 
 (*In ap_split S T u, term u expects separate arguments for the factors of S,
   with result type T.  The call creates a new term expecting one argument
   of type S.*)
-fun ap_split (Type ("*", [T1, T2])) T3 u =
+fun ap_split (Type (@{type_name "*"}, [T1, T2])) T3 u =
       internal_split_const (T1, T2, T3) $
       Abs ("v", T1,
           ap_split T2 T3
@@ -127,7 +112,7 @@
   THEN' hyp_subst_tac
   THEN' K prune_params_tac;
 
-val split_rule_ss = HOL_basic_ss addsimps [split_conv, fst_conv, snd_conv];
+val split_rule_ss = HOL_basic_ss addsimps [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}];
 
 fun split_rule_goal ctxt xss rl =
   let
@@ -159,5 +144,3 @@
 
 end;
 
-structure BasicSplitRule: BASIC_SPLIT_RULE = SplitRule;
-open BasicSplitRule;
--- a/src/HOL/Tools/string_syntax.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/string_syntax.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -71,7 +71,7 @@
       [] =>
         Syntax.Appl
           [Syntax.Constant Syntax.constrainC,
-            Syntax.Constant @{const_syntax Nil}, Syntax.Constant "string"]  (* FIXME @{type_syntax} *)
+            Syntax.Constant @{const_syntax Nil}, Syntax.Constant @{type_syntax string}]
     | cs => mk_string cs)
   | string_ast_tr asts = raise AST ("string_tr", asts);
 
--- a/src/HOL/Tools/typedef.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Tools/typedef.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -255,7 +255,7 @@
     (Scan.optional (P.$$$ "(" |--
         ((P.$$$ "open" >> K false) -- Scan.option P.binding ||
           P.binding >> (fn s => (true, SOME s))) --| P.$$$ ")") (true, NONE) --
-      (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
+      (P.type_args -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
       Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding))
     >> (fn ((((((def, opt_name), (vs, t)), mx), A), morphs)) =>
         Toplevel.print o Toplevel.theory_to_proof
--- a/src/HOL/Typerep.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Typerep.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -25,15 +25,16 @@
   fun typerep_tr (*"_TYPEREP"*) [ty] =
         Syntax.const @{const_syntax typerep} $
           (Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax "TYPE"} $
-            (Syntax.const "itself" $ ty))  (* FIXME @{type_syntax} *)
+            (Syntax.const @{type_syntax itself} $ ty))
     | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
 in [(@{syntax_const "_TYPEREP"}, typerep_tr)] end
 *}
 
 typed_print_translation {*
 let
-  fun typerep_tr' show_sorts (*"typerep"*)  (* FIXME @{type_syntax} *)
-          (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) =
+  fun typerep_tr' show_sorts (*"typerep"*)
+          (Type (@{type_syntax fun}, [Type (@{type_syntax itself}, [T]), _]))
+          (Const (@{const_syntax TYPE}, _) :: ts) =
         Term.list_comb
           (Syntax.const @{syntax_const "_TYPEREP"} $ Syntax.term_of_typ show_sorts T, ts)
     | typerep_tr' _ T ts = raise Match;
--- a/src/HOL/UNITY/Comp.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -27,23 +27,20 @@
 
 end
 
-constdefs
-  component_of :: "'a program =>'a program=> bool"
-                                    (infixl "component'_of" 50)
+definition component_of :: "'a program =>'a program=> bool" (infixl "component'_of" 50) where
   "F component_of H == \<exists>G. F ok G & F\<squnion>G = H"
 
-  strict_component_of :: "'a program\<Rightarrow>'a program=> bool"
-                                    (infixl "strict'_component'_of" 50)
+definition strict_component_of :: "'a program\<Rightarrow>'a program=> bool" (infixl "strict'_component'_of" 50)  where
   "F strict_component_of H == F component_of H & F\<noteq>H"
 
-  preserves :: "('a=>'b) => 'a program set"
+definition preserves :: "('a=>'b) => 'a program set" where
     "preserves v == \<Inter>z. stable {s. v s = z}"
 
-  localize  :: "('a=>'b) => 'a program => 'a program"
+definition localize :: "('a=>'b) => 'a program => 'a program" where
   "localize v F == mk_program(Init F, Acts F,
                               AllowedActs F \<inter> (\<Union>G \<in> preserves v. Acts G))"
 
-  funPair      :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c"
+definition funPair :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c" where
   "funPair f g == %x. (f x, g x)"
 
 
--- a/src/HOL/UNITY/Comp/AllocImpl.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/AllocImpl.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -20,8 +20,7 @@
   "'b merge" +
   dummy :: 'a       (*dummy field for new variables*)
 
-constdefs
-  non_dummy :: "('a,'b) merge_d => 'b merge"
+definition non_dummy :: "('a,'b) merge_d => 'b merge" where
     "non_dummy s == (|In = In s, Out = Out s, iOut = iOut s|)"
 
 record 'b distr =
--- a/src/HOL/UNITY/Comp/Counter.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/Counter.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -16,25 +16,21 @@
 datatype name = C | c nat
 types state = "name=>int"
 
-consts  
-  sum  :: "[nat,state]=>int"
-  sumj :: "[nat, nat, state]=>int"
-
-primrec (* sum I s = sigma_{i<I}. s (c i) *)
+primrec sum  :: "[nat,state]=>int" where
+  (* sum I s = sigma_{i<I}. s (c i) *)
   "sum 0 s = 0"
-  "sum (Suc i) s = s (c i) + sum i s"
+| "sum (Suc i) s = s (c i) + sum i s"
 
-primrec
+primrec sumj :: "[nat, nat, state]=>int" where
   "sumj 0 i s = 0"
-  "sumj (Suc n) i s = (if n=i then sum n s else s (c n) + sumj n i s)"
+| "sumj (Suc n) i s = (if n=i then sum n s else s (c n) + sumj n i s)"
   
 types command = "(state*state)set"
 
-constdefs
-  a :: "nat=>command"
+definition a :: "nat=>command" where
  "a i == {(s, s'). s'=s(c i:= s (c i) + 1, C:= s C + 1)}"
 
-  Component :: "nat => state program"
+definition Component :: "nat => state program" where
   "Component i ==
     mk_total_program({s. s C = 0 & s (c i) = 0}, {a i},
                      \<Union>G \<in> preserves (%s. s (c i)). Acts G)"
--- a/src/HOL/UNITY/Comp/Counterc.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/Counterc.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -21,25 +21,21 @@
   C :: "state=>int"
   c :: "state=>nat=>int"
 
-consts  
-  sum  :: "[nat,state]=>int"
-  sumj :: "[nat, nat, state]=>int"
-
-primrec (* sum I s = sigma_{i<I}. c s i *)
+primrec sum  :: "[nat,state]=>int" where
+  (* sum I s = sigma_{i<I}. c s i *)
   "sum 0 s = 0"
-  "sum (Suc i) s = (c s) i + sum i s"
+| "sum (Suc i) s = (c s) i + sum i s"
 
-primrec
+primrec sumj :: "[nat, nat, state]=>int" where
   "sumj 0 i s = 0"
-  "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)"
+| "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)"
   
 types command = "(state*state)set"
 
-constdefs
-  a :: "nat=>command"
+definition a :: "nat=>command" where
  "a i == {(s, s'). (c s') i = (c s) i + 1 & (C s') = (C s) + 1}"
  
-  Component :: "nat => state program"
+definition Component :: "nat => state program" where
   "Component i == mk_total_program({s. C s = 0 & (c s) i = 0},
                                    {a i},
                                    \<Union>G \<in> preserves (%s. (c s) i). Acts G)"
@@ -127,4 +123,4 @@
 apply (auto intro!: sum0 p2_p3)
 done
 
-end  
+end
--- a/src/HOL/UNITY/Comp/Priority.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/Priority.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Priority
-    ID:         $Id$
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 *)
@@ -22,52 +21,50 @@
 
 text{*Following the definitions given in section 4.4 *}
 
-constdefs
-  highest :: "[vertex, (vertex*vertex)set]=>bool"
+definition highest :: "[vertex, (vertex*vertex)set]=>bool" where
   "highest i r == A i r = {}"
     --{* i has highest priority in r *}
   
-  lowest :: "[vertex, (vertex*vertex)set]=>bool"
+definition lowest :: "[vertex, (vertex*vertex)set]=>bool" where
   "lowest i r == R i r = {}"
     --{* i has lowest priority in r *}
 
-  act :: command
+definition act :: command where
   "act i == {(s, s'). s'=reverse i s & highest i s}"
 
-  Component :: "vertex=>state program"
+definition Component :: "vertex=>state program" where
   "Component i == mk_total_program({init}, {act i}, UNIV)"
     --{* All components start with the same initial state *}
 
 
 text{*Some Abbreviations *}
-constdefs
-  Highest :: "vertex=>state set"
+definition Highest :: "vertex=>state set" where
   "Highest i == {s. highest i s}"
 
-  Lowest :: "vertex=>state set"
+definition Lowest :: "vertex=>state set" where
   "Lowest i == {s. lowest i s}"
 
-  Acyclic :: "state set"
+definition Acyclic :: "state set" where
   "Acyclic == {s. acyclic s}"
 
 
-  Maximal :: "state set"
+definition Maximal :: "state set" where
     --{* Every ``above'' set has a maximal vertex*}
   "Maximal == \<Inter>i. {s. ~highest i s-->(\<exists>j \<in> above i  s. highest j s)}"
 
-  Maximal' :: "state set"
+definition Maximal' :: "state set" where
     --{* Maximal vertex: equivalent definition*}
   "Maximal' == \<Inter>i. Highest i Un (\<Union>j. {s. j \<in> above i s} Int Highest j)"
 
   
-  Safety :: "state set"
+definition Safety :: "state set" where
   "Safety == \<Inter>i. {s. highest i s --> (\<forall>j \<in> neighbors i s. ~highest j s)}"
 
 
   (* Composition of a finite set of component;
      the vertex 'UNIV' is finite by assumption *)
   
-  system :: "state program"
+definition system :: "state program" where
   "system == JN i. Component i"
 
 
--- a/src/HOL/UNITY/Comp/PriorityAux.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/PriorityAux.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -12,38 +12,37 @@
 
 typedecl vertex
   
-constdefs
-  symcl :: "(vertex*vertex)set=>(vertex*vertex)set"
+definition symcl :: "(vertex*vertex)set=>(vertex*vertex)set" where
   "symcl r == r \<union> (r^-1)"
     --{* symmetric closure: removes the orientation of a relation*}
 
-  neighbors :: "[vertex, (vertex*vertex)set]=>vertex set"
+definition neighbors :: "[vertex, (vertex*vertex)set]=>vertex set" where
   "neighbors i r == ((r \<union> r^-1)``{i}) - {i}"
     --{* Neighbors of a vertex i *}
 
-  R :: "[vertex, (vertex*vertex)set]=>vertex set"
+definition R :: "[vertex, (vertex*vertex)set]=>vertex set" where
   "R i r == r``{i}"
 
-  A :: "[vertex, (vertex*vertex)set]=>vertex set"
+definition A :: "[vertex, (vertex*vertex)set]=>vertex set" where
   "A i r == (r^-1)``{i}"
 
-  reach :: "[vertex, (vertex*vertex)set]=> vertex set"
+definition reach :: "[vertex, (vertex*vertex)set]=> vertex set" where
   "reach i r == (r^+)``{i}"
     --{* reachable and above vertices: the original notation was R* and A* *}
 
-  above :: "[vertex, (vertex*vertex)set]=> vertex set"
+definition above :: "[vertex, (vertex*vertex)set]=> vertex set" where
   "above i r == ((r^-1)^+)``{i}"  
 
-  reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set"
+definition reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set" where
   "reverse i r == (r - {(x,y). x=i | y=i} \<inter> r) \<union> ({(x,y). x=i|y=i} \<inter> r)^-1"
 
-  derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool"
+definition derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where
     --{* The original definition *}
   "derive1 i r q == symcl r = symcl q &
                     (\<forall>k k'. k\<noteq>i & k'\<noteq>i -->((k,k'):r) = ((k,k'):q)) &
                     A i r = {} & R i q = {}"
 
-  derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool"
+definition derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where
     --{* Our alternative definition *}
   "derive i r q == A i r = {} & (q = reverse i r)"
 
--- a/src/HOL/UNITY/Comp/Progress.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/Progress.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Progress
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   2003  University of Cambridge
 
@@ -13,11 +12,10 @@
 subsection {*The Composition of Two Single-Assignment Programs*}
 text{*Thesis Section 4.4.2*}
 
-constdefs
-  FF :: "int program"
+definition FF :: "int program" where
     "FF == mk_total_program (UNIV, {range (\<lambda>x. (x, x+1))}, UNIV)"
 
-  GG :: "int program"
+definition GG :: "int program" where
     "GG == mk_total_program (UNIV, {range (\<lambda>x. (x, 2*x))}, UNIV)"
 
 subsubsection {*Calculating @{term "wens_set FF (atLeast k)"}*}
--- a/src/HOL/UNITY/Comp/TimerArray.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Comp/TimerArray.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/TimerArray.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -10,14 +9,13 @@
 
 types 'a state = "nat * 'a"   (*second component allows new variables*)
 
-constdefs
-  count  :: "'a state => nat"
+definition count :: "'a state => nat" where
     "count s == fst s"
   
-  decr  :: "('a state * 'a state) set"
+definition decr  :: "('a state * 'a state) set" where
     "decr == UN n uu. {((Suc n, uu), (n,uu))}"
   
-  Timer :: "'a state program"
+definition Timer :: "'a state program" where
     "Timer == mk_total_program (UNIV, {decr}, UNIV)"
 
 
--- a/src/HOL/UNITY/Constrains.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Constrains.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -32,22 +32,21 @@
   | Acts:  "[| act: Acts F;  s \<in> reachable F;  (s,s'): act |]
             ==> s' \<in> reachable F"
 
-constdefs
-  Constrains :: "['a set, 'a set] => 'a program set"  (infixl "Co" 60)
+definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where
     "A Co B == {F. F \<in> (reachable F \<inter> A)  co  B}"
 
-  Unless  :: "['a set, 'a set] => 'a program set"     (infixl "Unless" 60)
+definition Unless  :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) where
     "A Unless B == (A-B) Co (A \<union> B)"
 
-  Stable     :: "'a set => 'a program set"
+definition Stable     :: "'a set => 'a program set" where
     "Stable A == A Co A"
 
   (*Always is the weak form of "invariant"*)
-  Always :: "'a set => 'a program set"
+definition Always :: "'a set => 'a program set" where
     "Always A == {F. Init F \<subseteq> A} \<inter> Stable A"
 
   (*Polymorphic in both states and the meaning of \<le> *)
-  Increasing :: "['a => 'b::{order}] => 'a program set"
+definition Increasing :: "['a => 'b::{order}] => 'a program set" where
     "Increasing f == \<Inter>z. Stable {s. z \<le> f s}"
 
 
--- a/src/HOL/UNITY/FP.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/FP.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/FP
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -10,12 +9,10 @@
 
 theory FP imports UNITY begin
 
-constdefs
-
-  FP_Orig :: "'a program => 'a set"
+definition FP_Orig :: "'a program => 'a set" where
     "FP_Orig F == Union{A. ALL B. F : stable (A Int B)}"
 
-  FP :: "'a program => 'a set"
+definition FP :: "'a program => 'a set" where
     "FP F == {s. F : stable {s}}"
 
 lemma stable_FP_Orig_Int: "F : stable (FP_Orig F Int B)"
--- a/src/HOL/UNITY/Follows.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Follows.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -7,10 +7,7 @@
 
 theory Follows imports SubstAx ListOrder Multiset begin
 
-constdefs
-
-  Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set"
-                 (infixl "Fols" 65)
+definition Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set" (infixl "Fols" 65) where
    "f Fols g == Increasing g \<inter> Increasing f Int
                 Always {s. f s \<le> g s} Int
                 (\<Inter>k. {s. k \<le> g s} LeadsTo {s. k \<le> f s})"
--- a/src/HOL/UNITY/Guar.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Guar.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -22,51 +22,47 @@
 text{*Existential and Universal properties.  I formalize the two-program
       case, proving equivalence with Chandy and Sanders's n-ary definitions*}
 
-constdefs
-
-  ex_prop  :: "'a program set => bool"
+definition ex_prop :: "'a program set => bool" where
    "ex_prop X == \<forall>F G. F ok G -->F \<in> X | G \<in> X --> (F\<squnion>G) \<in> X"
 
-  strict_ex_prop  :: "'a program set => bool"
+definition strict_ex_prop  :: "'a program set => bool" where
    "strict_ex_prop X == \<forall>F G.  F ok G --> (F \<in> X | G \<in> X) = (F\<squnion>G \<in> X)"
 
-  uv_prop  :: "'a program set => bool"
+definition uv_prop  :: "'a program set => bool" where
    "uv_prop X == SKIP \<in> X & (\<forall>F G. F ok G --> F \<in> X & G \<in> X --> (F\<squnion>G) \<in> X)"
 
-  strict_uv_prop  :: "'a program set => bool"
+definition strict_uv_prop  :: "'a program set => bool" where
    "strict_uv_prop X == 
       SKIP \<in> X & (\<forall>F G. F ok G --> (F \<in> X & G \<in> X) = (F\<squnion>G \<in> X))"
 
 
 text{*Guarantees properties*}
 
-constdefs
-
-  guar :: "['a program set, 'a program set] => 'a program set"
-          (infixl "guarantees" 55)  (*higher than membership, lower than Co*)
+definition guar :: "['a program set, 'a program set] => 'a program set" (infixl "guarantees" 55) where
+          (*higher than membership, lower than Co*)
    "X guarantees Y == {F. \<forall>G. F ok G --> F\<squnion>G \<in> X --> F\<squnion>G \<in> Y}"
   
 
   (* Weakest guarantees *)
-   wg :: "['a program, 'a program set] =>  'a program set"
+definition wg :: "['a program, 'a program set] => 'a program set" where
   "wg F Y == Union({X. F \<in> X guarantees Y})"
 
    (* Weakest existential property stronger than X *)
-   wx :: "('a program) set => ('a program)set"
+definition wx :: "('a program) set => ('a program)set" where
    "wx X == Union({Y. Y \<subseteq> X & ex_prop Y})"
   
   (*Ill-defined programs can arise through "Join"*)
-  welldef :: "'a program set"
+definition welldef :: "'a program set" where
   "welldef == {F. Init F \<noteq> {}}"
   
-  refines :: "['a program, 'a program, 'a program set] => bool"
-                        ("(3_ refines _ wrt _)" [10,10,10] 10)
+definition refines :: "['a program, 'a program, 'a program set] => bool"
+                        ("(3_ refines _ wrt _)" [10,10,10] 10) where
   "G refines F wrt X ==
      \<forall>H. (F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X) --> 
          (G\<squnion>H \<in> welldef \<inter> X)"
 
-  iso_refines :: "['a program, 'a program, 'a program set] => bool"
-                              ("(3_ iso'_refines _ wrt _)" [10,10,10] 10)
+definition iso_refines :: "['a program, 'a program, 'a program set] => bool"
+                              ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where
   "G iso_refines F wrt X ==
    F \<in> welldef \<inter> X --> G \<in> welldef \<inter> X"
 
--- a/src/HOL/UNITY/Lift_prog.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Lift_prog.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Lift_prog.thy
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1999  University of Cambridge
 
@@ -10,30 +9,28 @@
 
 theory Lift_prog imports Rename begin
 
-constdefs
-
-  insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)"
+definition insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)" where
     "insert_map i z f k == if k<i then f k
                            else if k=i then z
                            else f(k - 1)"
 
-  delete_map :: "[nat, nat=>'b] => (nat=>'b)"
+definition delete_map :: "[nat, nat=>'b] => (nat=>'b)" where
     "delete_map i g k == if k<i then g k else g (Suc k)"
 
-  lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c"
+definition lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c" where
     "lift_map i == %(s,(f,uu)). (insert_map i s f, uu)"
 
-  drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)"
+definition drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)" where
     "drop_map i == %(g, uu). (g i, (delete_map i g, uu))"
 
-  lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set"
+definition lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set" where
     "lift_set i A == lift_map i ` A"
 
-  lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program"
+definition lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program" where
     "lift i == rename (lift_map i)"
 
   (*simplifies the expression of specifications*)
-  sub :: "['a, 'a=>'b] => 'b"
+definition sub :: "['a, 'a=>'b] => 'b" where
     "sub == %i f. f i"
 
 
--- a/src/HOL/UNITY/ListOrder.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/ListOrder.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -42,11 +42,10 @@
 
 end
 
-constdefs
-  Le :: "(nat*nat) set"
+definition Le :: "(nat*nat) set" where
     "Le == {(x,y). x <= y}"
 
-  Ge :: "(nat*nat) set"
+definition  Ge :: "(nat*nat) set" where
     "Ge == {(x,y). y <= x}"
 
 abbreviation
--- a/src/HOL/UNITY/PPROD.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/PPROD.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -10,9 +10,8 @@
 
 theory PPROD imports Lift_prog begin
 
-constdefs
-  PLam  :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program]
-            => ((nat=>'b) * 'c) program"
+definition PLam :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program]
+            => ((nat=>'b) * 'c) program" where
     "PLam I F == \<Squnion>i \<in> I. lift i (F i)"
 
 syntax
--- a/src/HOL/UNITY/ProgressSets.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/ProgressSets.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -19,13 +19,12 @@
 
 subsection {*Complete Lattices and the Operator @{term cl}*}
 
-constdefs
-  lattice :: "'a set set => bool"
+definition lattice :: "'a set set => bool" where
    --{*Meier calls them closure sets, but they are just complete lattices*}
    "lattice L ==
          (\<forall>M. M \<subseteq> L --> \<Inter>M \<in> L) & (\<forall>M. M \<subseteq> L --> \<Union>M \<in> L)"
 
-  cl :: "['a set set, 'a set] => 'a set"
+definition cl :: "['a set set, 'a set] => 'a set" where
    --{*short for ``closure''*}
    "cl L r == \<Inter>{x. x\<in>L & r \<subseteq> x}"
 
@@ -143,12 +142,11 @@
 text{*A progress set satisfies certain closure conditions and is a 
 simple way of including the set @{term "wens_set F B"}.*}
 
-constdefs 
-  closed :: "['a program, 'a set, 'a set,  'a set set] => bool"
+definition closed :: "['a program, 'a set, 'a set,  'a set set] => bool" where
    "closed F T B L == \<forall>M. \<forall>act \<in> Acts F. B\<subseteq>M & T\<inter>M \<in> L -->
                               T \<inter> (B \<union> wp act M) \<in> L"
 
-  progress_set :: "['a program, 'a set, 'a set] => 'a set set set"
+definition progress_set :: "['a program, 'a set, 'a set] => 'a set set set" where
    "progress_set F T B ==
       {L. lattice L & B \<in> L & T \<in> L & closed F T B L}"
 
@@ -324,12 +322,11 @@
 subsubsection {*Lattices and Relations*}
 text{*From Meier's thesis, section 4.5.3*}
 
-constdefs
-  relcl :: "'a set set => ('a * 'a) set"
+definition relcl :: "'a set set => ('a * 'a) set" where
     -- {*Derived relation from a lattice*}
     "relcl L == {(x,y). y \<in> cl L {x}}"
   
-  latticeof :: "('a * 'a) set => 'a set set"
+definition latticeof :: "('a * 'a) set => 'a set set" where
     -- {*Derived lattice from a relation: the set of upwards-closed sets*}
     "latticeof r == {X. \<forall>s t. s \<in> X & (s,t) \<in> r --> t \<in> X}"
 
@@ -405,8 +402,7 @@
 
 subsubsection {*Decoupling Theorems*}
 
-constdefs
-  decoupled :: "['a program, 'a program] => bool"
+definition decoupled :: "['a program, 'a program] => bool" where
    "decoupled F G ==
         \<forall>act \<in> Acts F. \<forall>B. G \<in> stable B --> G \<in> stable (wp act B)"
 
@@ -469,8 +465,7 @@
 subsection{*Composition Theorems Based on Monotonicity and Commutativity*}
 
 subsubsection{*Commutativity of @{term "cl L"} and assignment.*}
-constdefs 
-  commutes :: "['a program, 'a set, 'a set,  'a set set] => bool"
+definition commutes :: "['a program, 'a set, 'a set,  'a set set] => bool" where
    "commutes F T B L ==
        \<forall>M. \<forall>act \<in> Acts F. B \<subseteq> M --> 
            cl L (T \<inter> wp act M) \<subseteq> T \<inter> (B \<union> wp act (cl L (T\<inter>M)))"
@@ -511,8 +506,7 @@
 
 
 text{*Possibly move to Relation.thy, after @{term single_valued}*}
-constdefs
-  funof :: "[('a*'b)set, 'a] => 'b"
+definition funof :: "[('a*'b)set, 'a] => 'b" where
    "funof r == (\<lambda>x. THE y. (x,y) \<in> r)"
 
 lemma funof_eq: "[|single_valued r; (x,y) \<in> r|] ==> funof r x = y"
--- a/src/HOL/UNITY/Project.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Project.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -11,19 +11,18 @@
 
 theory Project imports Extend begin
 
-constdefs
-  projecting :: "['c program => 'c set, 'a*'b => 'c, 
-                  'a program, 'c program set, 'a program set] => bool"
+definition projecting :: "['c program => 'c set, 'a*'b => 'c, 
+                  'a program, 'c program set, 'a program set] => bool" where
     "projecting C h F X' X ==
        \<forall>G. extend h F\<squnion>G \<in> X' --> F\<squnion>project h (C G) G \<in> X"
 
-  extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, 
-                 'c program set, 'a program set] => bool"
+definition extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, 
+                 'c program set, 'a program set] => bool" where
     "extending C h F Y' Y ==
        \<forall>G. extend h F  ok G --> F\<squnion>project h (C G) G \<in> Y
               --> extend h F\<squnion>G \<in> Y'"
 
-  subset_closed :: "'a set set => bool"
+definition subset_closed :: "'a set set => bool" where
     "subset_closed U == \<forall>A \<in> U. Pow A \<subseteq> U"  
 
 
--- a/src/HOL/UNITY/Rename.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Rename.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -8,9 +8,7 @@
 
 theory Rename imports Extend begin
 
-constdefs
-  
-  rename :: "['a => 'b, 'a program] => 'b program"
+definition rename :: "['a => 'b, 'a program] => 'b program" where
     "rename h == extend (%(x,u::unit). h x)"
 
 declare image_inv_f_f [simp] image_surj_f_inv_f [simp]
--- a/src/HOL/UNITY/Simple/Channel.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/Channel.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -15,8 +15,7 @@
 consts
   F :: "state program"
 
-constdefs
-  minSet :: "nat set => nat option"
+definition minSet :: "nat set => nat option" where
     "minSet A == if A={} then None else Some (LEAST x. x \<in> A)"
 
 axioms
--- a/src/HOL/UNITY/Simple/Common.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/Common.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -24,11 +24,10 @@
   fasc:  "m \<le> ftime n"
   gasc:  "m \<le> gtime n"
 
-constdefs
-  common :: "nat set"
+definition common :: "nat set" where
     "common == {n. ftime n = n & gtime n = n}"
 
-  maxfg :: "nat => nat set"
+definition maxfg :: "nat => nat set" where
     "maxfg m == {t. t \<le> max (ftime m) (gtime m)}"
 
 
--- a/src/HOL/UNITY/Simple/NSP_Bad.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/NSP_Bad.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -53,8 +53,7 @@
                & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \<in> set s3}"
 
 
-constdefs
-  Nprg :: "state program"
+definition Nprg :: "state program" where
     (*Initial trace is empty*)
     "Nprg == mk_total_program({[]}, {Fake, NS1, NS2, NS3}, UNIV)"
 
--- a/src/HOL/UNITY/Simple/Reach.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/Reach.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -17,21 +17,19 @@
 
   edges :: "(vertex*vertex) set"
 
-constdefs
-
-  asgt  :: "[vertex,vertex] => (state*state) set"
+definition asgt :: "[vertex,vertex] => (state*state) set" where
     "asgt u v == {(s,s'). s' = s(v:= s u | s v)}"
 
-  Rprg :: "state program"
+definition Rprg :: "state program" where
     "Rprg == mk_total_program ({%v. v=init}, \<Union>(u,v)\<in>edges. {asgt u v}, UNIV)"
 
-  reach_invariant :: "state set"
+definition reach_invariant :: "state set" where
     "reach_invariant == {s. (\<forall>v. s v --> (init, v) \<in> edges^*) & s init}"
 
-  fixedpoint :: "state set"
+definition fixedpoint :: "state set" where
     "fixedpoint == {s. \<forall>(u,v)\<in>edges. s u --> s v}"
 
-  metric :: "state => nat"
+definition metric :: "state => nat" where
     "metric s == card {v. ~ s v}"
 
 
--- a/src/HOL/UNITY/Simple/Reachability.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/Reachability.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -25,23 +25,22 @@
     base: "v \<in> V ==> ((v,v) \<in> REACHABLE)"
   | step: "((u,v) \<in> REACHABLE) & (v,w) \<in> E ==> ((u,w) \<in> REACHABLE)"
 
-constdefs
-  reachable :: "vertex => state set"
+definition reachable :: "vertex => state set" where
   "reachable p == {s. reach s p}"
 
-  nmsg_eq :: "nat => edge  => state set"
+definition nmsg_eq :: "nat => edge  => state set" where
   "nmsg_eq k == %e. {s. nmsg s e = k}"
 
-  nmsg_gt :: "nat => edge  => state set"
+definition nmsg_gt :: "nat => edge  => state set" where
   "nmsg_gt k  == %e. {s. k < nmsg s e}"
 
-  nmsg_gte :: "nat => edge => state set"
+definition nmsg_gte :: "nat => edge => state set" where
   "nmsg_gte k == %e. {s. k \<le> nmsg s e}"
 
-  nmsg_lte  :: "nat => edge => state set"
+definition nmsg_lte  :: "nat => edge => state set" where
   "nmsg_lte k  == %e. {s. nmsg s e \<le> k}"
 
-  final :: "state set"
+definition final :: "state set" where
   "final == (\<Inter>v\<in>V. reachable v <==> {s. (root, v) \<in> REACHABLE}) \<inter> 
             (INTER E (nmsg_eq 0))"
 
--- a/src/HOL/UNITY/Simple/Token.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Simple/Token.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Token
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 *)
@@ -24,17 +23,16 @@
   proc  :: "nat => pstate"
 
 
-constdefs
-  HasTok :: "nat => state set"
+definition HasTok :: "nat => state set" where
     "HasTok i == {s. token s = i}"
 
-  H :: "nat => state set"
+definition H :: "nat => state set" where
     "H i == {s. proc s i = Hungry}"
 
-  E :: "nat => state set"
+definition E :: "nat => state set" where
     "E i == {s. proc s i = Eating}"
 
-  T :: "nat => state set"
+definition T :: "nat => state set" where
     "T i == {s. proc s i = Thinking}"
 
 
--- a/src/HOL/UNITY/SubstAx.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/SubstAx.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/SubstAx
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -10,15 +9,14 @@
 
 theory SubstAx imports WFair Constrains begin
 
-constdefs
-   Ensures :: "['a set, 'a set] => 'a program set"    (infixl "Ensures" 60)
+definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where
     "A Ensures B == {F. F \<in> (reachable F \<inter> A) ensures B}"
 
-   LeadsTo :: "['a set, 'a set] => 'a program set"    (infixl "LeadsTo" 60)
+definition LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) where
     "A LeadsTo B == {F. F \<in> (reachable F \<inter> A) leadsTo B}"
 
-syntax (xsymbols)
-  "op LeadsTo" :: "['a set, 'a set] => 'a program set" (infixl " \<longmapsto>w " 60)
+notation (xsymbols)
+  LeadsTo  (infixl " \<longmapsto>w " 60)
 
 
 text{*Resembles the previous definition of LeadsTo*}
--- a/src/HOL/UNITY/Transformers.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/Transformers.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -20,16 +20,15 @@
 subsection{*Defining the Predicate Transformers @{term wp},
    @{term awp} and  @{term wens}*}
 
-constdefs
-  wp :: "[('a*'a) set, 'a set] => 'a set"  
+definition wp :: "[('a*'a) set, 'a set] => 'a set" where  
     --{*Dijkstra's weakest-precondition operator (for an individual command)*}
     "wp act B == - (act^-1 `` (-B))"
 
-  awp :: "['a program, 'a set] => 'a set"  
+definition awp :: "['a program, 'a set] => 'a set" where
     --{*Dijkstra's weakest-precondition operator (for a program)*}
     "awp F B == (\<Inter>act \<in> Acts F. wp act B)"
 
-  wens :: "['a program, ('a*'a) set, 'a set] => 'a set"  
+definition wens :: "['a program, ('a*'a) set, 'a set] => 'a set" where
     --{*The weakest-ensures transformer*}
     "wens F act B == gfp(\<lambda>X. (wp act B \<inter> awp F (B \<union> X)) \<union> B)"
 
@@ -335,11 +334,10 @@
 
 text{*Next, we express the @{term "wens_set"} for single-assignment programs*}
 
-constdefs
-  wens_single_finite :: "[('a*'a) set, 'a set, nat] => 'a set"  
+definition wens_single_finite :: "[('a*'a) set, 'a set, nat] => 'a set" where  
     "wens_single_finite act B k == \<Union>i \<in> atMost k. (wp act ^^ i) B"
 
-  wens_single :: "[('a*'a) set, 'a set] => 'a set"  
+definition wens_single :: "[('a*'a) set, 'a set] => 'a set" where
     "wens_single act B == \<Union>i. (wp act ^^ i) B"
 
 lemma wens_single_Un_eq:
--- a/src/HOL/UNITY/UNITY.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/UNITY.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -17,40 +17,39 @@
                    allowed :: ('a * 'a)set set). Id \<in> acts & Id: allowed}" 
   by blast
 
-constdefs
-  Acts :: "'a program => ('a * 'a)set set"
+definition Acts :: "'a program => ('a * 'a)set set" where
     "Acts F == (%(init, acts, allowed). acts) (Rep_Program F)"
 
-  "constrains" :: "['a set, 'a set] => 'a program set"  (infixl "co"     60)
+definition "constrains" :: "['a set, 'a set] => 'a program set"  (infixl "co"     60) where
     "A co B == {F. \<forall>act \<in> Acts F. act``A \<subseteq> B}"
 
-  unless  :: "['a set, 'a set] => 'a program set"  (infixl "unless" 60)
+definition unless  :: "['a set, 'a set] => 'a program set"  (infixl "unless" 60)  where
     "A unless B == (A-B) co (A \<union> B)"
 
-  mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
-                   => 'a program"
+definition mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
+                   => 'a program" where
     "mk_program == %(init, acts, allowed).
                       Abs_Program (init, insert Id acts, insert Id allowed)"
 
-  Init :: "'a program => 'a set"
+definition Init :: "'a program => 'a set" where
     "Init F == (%(init, acts, allowed). init) (Rep_Program F)"
 
-  AllowedActs :: "'a program => ('a * 'a)set set"
+definition AllowedActs :: "'a program => ('a * 'a)set set" where
     "AllowedActs F == (%(init, acts, allowed). allowed) (Rep_Program F)"
 
-  Allowed :: "'a program => 'a program set"
+definition Allowed :: "'a program => 'a program set" where
     "Allowed F == {G. Acts G \<subseteq> AllowedActs F}"
 
-  stable     :: "'a set => 'a program set"
+definition stable     :: "'a set => 'a program set" where
     "stable A == A co A"
 
-  strongest_rhs :: "['a program, 'a set] => 'a set"
+definition strongest_rhs :: "['a program, 'a set] => 'a set" where
     "strongest_rhs F A == Inter {B. F \<in> A co B}"
 
-  invariant :: "'a set => 'a program set"
+definition invariant :: "'a set => 'a program set" where
     "invariant A == {F. Init F \<subseteq> A} \<inter> stable A"
 
-  increasing :: "['a => 'b::{order}] => 'a program set"
+definition increasing :: "['a => 'b::{order}] => 'a program set" where
     --{*Polymorphic in both states and the meaning of @{text "\<le>"}*}
     "increasing f == \<Inter>z. stable {s. z \<le> f s}"
 
@@ -357,20 +356,19 @@
 
 subsection{*Partial versus Total Transitions*}
 
-constdefs
-  totalize_act :: "('a * 'a)set => ('a * 'a)set"
+definition totalize_act :: "('a * 'a)set => ('a * 'a)set" where
     "totalize_act act == act \<union> Id_on (-(Domain act))"
 
-  totalize :: "'a program => 'a program"
+definition totalize :: "'a program => 'a program" where
     "totalize F == mk_program (Init F,
                                totalize_act ` Acts F,
                                AllowedActs F)"
 
-  mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
-                   => 'a program"
+definition mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
+                   => 'a program" where
     "mk_total_program args == totalize (mk_program args)"
 
-  all_total :: "'a program => bool"
+definition all_total :: "'a program => bool" where
     "all_total F == \<forall>act \<in> Acts F. Domain act = UNIV"
   
 lemma insert_Id_image_Acts: "f Id = Id ==> insert Id (f`Acts F) = f ` Acts F"
--- a/src/HOL/UNITY/WFair.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/UNITY/WFair.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/WFair
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -33,16 +32,17 @@
 is the impossibility law for leads-to.
 *}
 
-constdefs
+definition
 
   --{*This definition specifies conditional fairness.  The rest of the theory
       is generic to all forms of fairness.  To get weak fairness, conjoin
       the inclusion below with @{term "A \<subseteq> Domain act"}, which specifies 
       that the action is enabled over all of @{term A}.*}
-  transient :: "'a set => 'a program set"
+  transient :: "'a set => 'a program set" where
     "transient A == {F. \<exists>act\<in>Acts F. act``A \<subseteq> -A}"
 
-  ensures :: "['a set, 'a set] => 'a program set"       (infixl "ensures" 60)
+definition
+  ensures :: "['a set, 'a set] => 'a program set"       (infixl "ensures" 60) where
     "A ensures B == (A-B co A \<union> B) \<inter> transient (A-B)"
 
 
@@ -59,18 +59,16 @@
   | Union:  "\<forall>A \<in> S. (A,B) \<in> leads F ==> (Union S, B) \<in> leads F"
 
 
-constdefs
-
-  leadsTo :: "['a set, 'a set] => 'a program set"    (infixl "leadsTo" 60)
+definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where
      --{*visible version of the LEADS-TO relation*}
     "A leadsTo B == {F. (A,B) \<in> leads F}"
   
-  wlt :: "['a program, 'a set] => 'a set"
+definition wlt :: "['a program, 'a set] => 'a set" where
      --{*predicate transformer: the largest set that leads to @{term B}*}
     "wlt F B == Union {A. F \<in> A leadsTo B}"
 
-syntax (xsymbols)
-  "op leadsTo" :: "['a set, 'a set] => 'a program set" (infixl "\<longmapsto>" 60)
+notation (xsymbols)
+  leadsTo  (infixl "\<longmapsto>" 60)
 
 
 subsection{*transient*}
--- a/src/HOL/Word/WordDefinition.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Word/WordDefinition.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -165,14 +165,13 @@
 where
   "word_pred a = word_of_int (Int.pred (uint a))"
 
-constdefs
-  udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50)
+definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where
   "a udvd b == EX n>=0. uint b = n * uint a"
 
-  word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50)
+definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) where
   "a <=s b == sint a <= sint b"
 
-  word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50)
+definition word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50) where
   "(x <s y) == (x <=s y & x ~= y)"
 
 
@@ -223,89 +222,81 @@
 
 end
 
-constdefs
-  setBit :: "'a :: len0 word => nat => 'a word" 
+definition setBit :: "'a :: len0 word => nat => 'a word" where 
   "setBit w n == set_bit w n True"
 
-  clearBit :: "'a :: len0 word => nat => 'a word" 
+definition clearBit :: "'a :: len0 word => nat => 'a word" where
   "clearBit w n == set_bit w n False"
 
 
 subsection "Shift operations"
 
-constdefs
-  sshiftr1 :: "'a :: len word => 'a word" 
+definition sshiftr1 :: "'a :: len word => 'a word" where 
   "sshiftr1 w == word_of_int (bin_rest (sint w))"
 
-  bshiftr1 :: "bool => 'a :: len word => 'a word"
+definition bshiftr1 :: "bool => 'a :: len word => 'a word" where
   "bshiftr1 b w == of_bl (b # butlast (to_bl w))"
 
-  sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55)
+definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) where
   "w >>> n == (sshiftr1 ^^ n) w"
 
-  mask :: "nat => 'a::len word"
+definition mask :: "nat => 'a::len word" where
   "mask n == (1 << n) - 1"
 
-  revcast :: "'a :: len0 word => 'b :: len0 word"
+definition revcast :: "'a :: len0 word => 'b :: len0 word" where
   "revcast w ==  of_bl (takefill False (len_of TYPE('b)) (to_bl w))"
 
-  slice1 :: "nat => 'a :: len0 word => 'b :: len0 word"
+definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" where
   "slice1 n w == of_bl (takefill False n (to_bl w))"
 
-  slice :: "nat => 'a :: len0 word => 'b :: len0 word"
+definition slice :: "nat => 'a :: len0 word => 'b :: len0 word" where
   "slice n w == slice1 (size w - n) w"
 
 
 subsection "Rotation"
 
-constdefs
-  rotater1 :: "'a list => 'a list"
+definition rotater1 :: "'a list => 'a list" where
   "rotater1 ys == 
     case ys of [] => [] | x # xs => last ys # butlast ys"
 
-  rotater :: "nat => 'a list => 'a list"
+definition rotater :: "nat => 'a list => 'a list" where
   "rotater n == rotater1 ^^ n"
 
-  word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word"
+definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" where
   "word_rotr n w == of_bl (rotater n (to_bl w))"
 
-  word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word"
+definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" where
   "word_rotl n w == of_bl (rotate n (to_bl w))"
 
-  word_roti :: "int => 'a :: len0 word => 'a :: len0 word"
+definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word" where
   "word_roti i w == if i >= 0 then word_rotr (nat i) w
                     else word_rotl (nat (- i)) w"
 
 
 subsection "Split and cat operations"
 
-constdefs
-  word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word"
+definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" where
   "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))"
 
-  word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)"
+definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" where
   "word_split a == 
    case bin_split (len_of TYPE ('c)) (uint a) of 
      (u, v) => (word_of_int u, word_of_int v)"
 
-  word_rcat :: "'a :: len0 word list => 'b :: len0 word"
+definition word_rcat :: "'a :: len0 word list => 'b :: len0 word" where
   "word_rcat ws == 
   word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))"
 
-  word_rsplit :: "'a :: len0 word => 'b :: len word list"
+definition word_rsplit :: "'a :: len0 word => 'b :: len word list" where
   "word_rsplit w == 
   map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))"
 
-constdefs
-  -- "Largest representable machine integer."
-  max_word :: "'a::len word"
+definition max_word :: "'a::len word" -- "Largest representable machine integer." where
   "max_word \<equiv> word_of_int (2 ^ len_of TYPE('a) - 1)"
 
-consts 
-  of_bool :: "bool \<Rightarrow> 'a::len word"
-primrec
+primrec of_bool :: "bool \<Rightarrow> 'a::len word" where
   "of_bool False = 0"
-  "of_bool True = 1"
+| "of_bool True = 1"
 
 
 lemmas of_nth_def = word_set_bits_def
--- a/src/HOL/Word/WordGenLib.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/Word/WordGenLib.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -344,8 +344,7 @@
   apply (case_tac "1+n = 0", auto)
   done
 
-constdefs
-  word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a"
+definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a" where
   "word_rec forZero forSuc n \<equiv> nat_rec forZero (forSuc \<circ> of_nat) (unat n)"
 
 lemma word_rec_0: "word_rec z s 0 = z"
--- a/src/HOL/ZF/Games.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ZF/Games.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,4 +1,4 @@
-(*  Title:      HOL/ZF/Games.thy
+(*  Title:      HOL/ZF/MainZF.thy/Games.thy
     Author:     Steven Obua
 
 An application of HOLZF: Partizan Games.  See "Partizan Games in
@@ -9,12 +9,13 @@
 imports MainZF
 begin
 
-constdefs
-  fixgames :: "ZF set \<Rightarrow> ZF set"
+definition fixgames :: "ZF set \<Rightarrow> ZF set" where
   "fixgames A \<equiv> { Opair l r | l r. explode l \<subseteq> A & explode r \<subseteq> A}"
-  games_lfp :: "ZF set"
+
+definition games_lfp :: "ZF set" where
   "games_lfp \<equiv> lfp fixgames"
-  games_gfp :: "ZF set"
+
+definition games_gfp :: "ZF set" where
   "games_gfp \<equiv> gfp fixgames"
 
 lemma mono_fixgames: "mono (fixgames)"
@@ -42,12 +43,13 @@
     by auto
 qed
 
-constdefs
-  left_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool"
+definition left_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" where
   "left_option g opt \<equiv> (Elem opt (Fst g))"
-  right_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool"
+
+definition right_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" where
   "right_option g opt \<equiv> (Elem opt (Snd g))"
-  is_option_of :: "(ZF * ZF) set"
+
+definition is_option_of :: "(ZF * ZF) set" where
   "is_option_of \<equiv> { (opt, g) | opt g. g \<in> games_gfp \<and> (left_option g opt \<or> right_option g opt) }"
 
 lemma games_lfp_subset_gfp: "games_lfp \<subseteq> games_gfp"
@@ -190,14 +192,16 @@
 typedef game = games_lfp
   by (blast intro: games_lfp_nonempty)
 
-constdefs
-  left_options :: "game \<Rightarrow> game zet"
+definition left_options :: "game \<Rightarrow> game zet" where
   "left_options g \<equiv> zimage Abs_game (zexplode (Fst (Rep_game g)))"
-  right_options :: "game \<Rightarrow> game zet"
+
+definition right_options :: "game \<Rightarrow> game zet" where
   "right_options g \<equiv> zimage Abs_game (zexplode (Snd (Rep_game g)))"
-  options :: "game \<Rightarrow> game zet"
+
+definition options :: "game \<Rightarrow> game zet" where
   "options g \<equiv> zunion (left_options g) (right_options g)"
-  Game :: "game zet \<Rightarrow> game zet \<Rightarrow> game"
+
+definition Game :: "game zet \<Rightarrow> game zet \<Rightarrow> game" where
   "Game L R \<equiv> Abs_game (Opair (zimplode (zimage Rep_game L)) (zimplode (zimage Rep_game R)))"
   
 lemma Repl_Rep_game_Abs_game: "\<forall> e. Elem e z \<longrightarrow> e \<in> games_lfp \<Longrightarrow> Repl z (Rep_game o Abs_game) = z"
@@ -295,8 +299,7 @@
   apply simp
   done
 
-constdefs
-  option_of :: "(game * game) set"
+definition option_of :: "(game * game) set" where
   "option_of \<equiv> image (\<lambda> (option, g). (Abs_game option, Abs_game g)) is_option_of"
 
 lemma option_to_is_option_of: "((option, g) \<in> option_of) = ((Rep_game option, Rep_game g) \<in> is_option_of)"
@@ -344,13 +347,12 @@
     right_option_def[symmetric] left_option_def[symmetric])
   done
 
-consts
+function
   neg_game :: "game \<Rightarrow> game"
-
-recdef neg_game "option_of"
-  "neg_game g = Game (zimage neg_game (right_options g)) (zimage neg_game (left_options g))"
-
-declare neg_game.simps[simp del]
+where
+  [simp del]: "neg_game g = Game (zimage neg_game (right_options g)) (zimage neg_game (left_options g))"
+by auto
+termination by (relation "option_of") auto
 
 lemma "neg_game (neg_game g) = g"
   apply (induct g rule: neg_game.induct)
@@ -362,17 +364,16 @@
   apply (auto simp add: zet_ext_eq zimage_iff)
   done
 
-consts
+function
   ge_game :: "(game * game) \<Rightarrow> bool" 
-
-recdef ge_game "(gprod_2_1 option_of)"
-  "ge_game (G, H) = (\<forall> x. if zin x (right_options G) then (
+where
+  [simp del]: "ge_game (G, H) = (\<forall> x. if zin x (right_options G) then (
                           if zin x (left_options H) then \<not> (ge_game (H, x) \<or> (ge_game (x, G))) 
                                                     else \<not> (ge_game (H, x)))
                           else (if zin x (left_options H) then \<not> (ge_game (x, G)) else True))"
-(hints simp: gprod_2_1_def)
-
-declare ge_game.simps [simp del]
+by auto
+termination by (relation "(gprod_2_1 option_of)") 
+ (simp, auto simp: gprod_2_1_def)
 
 lemma ge_game_eq: "ge_game (G, H) = (\<forall> x. (zin x (right_options G) \<longrightarrow> \<not> ge_game (H, x)) \<and> (zin x (left_options H) \<longrightarrow> \<not> ge_game (x, G)))"
   apply (subst ge_game.simps[where G=G and H=H])
@@ -437,8 +438,7 @@
   qed
 qed
         
-constdefs
-  eq_game :: "game \<Rightarrow> game \<Rightarrow> bool"
+definition eq_game :: "game \<Rightarrow> game \<Rightarrow> bool" where
   "eq_game G H \<equiv> ge_game (G, H) \<and> ge_game (H, G)" 
 
 lemma eq_game_sym: "(eq_game G H) = (eq_game H G)"
@@ -501,23 +501,21 @@
 lemma eq_game_trans: "eq_game a b \<Longrightarrow> eq_game b c \<Longrightarrow> eq_game a c"
   by (auto simp add: eq_game_def intro: ge_game_trans)
 
-constdefs
-  zero_game :: game
-  "zero_game \<equiv> Game zempty zempty"
-
-consts 
-  plus_game :: "game * game \<Rightarrow> game"
+definition zero_game :: game
+ where  "zero_game \<equiv> Game zempty zempty"
 
-recdef plus_game "gprod_2_2 option_of"
-  "plus_game (G, H) = Game (zunion (zimage (\<lambda> g. plus_game (g, H)) (left_options G))
-                                   (zimage (\<lambda> h. plus_game (G, h)) (left_options H)))
-                           (zunion (zimage (\<lambda> g. plus_game (g, H)) (right_options G))
-                                   (zimage (\<lambda> h. plus_game (G, h)) (right_options H)))"
-(hints simp add: gprod_2_2_def)
+function 
+  plus_game :: "game \<Rightarrow> game \<Rightarrow> game"
+where
+  [simp del]: "plus_game G H = Game (zunion (zimage (\<lambda> g. plus_game g H) (left_options G))
+                                   (zimage (\<lambda> h. plus_game G h) (left_options H)))
+                           (zunion (zimage (\<lambda> g. plus_game g H) (right_options G))
+                                   (zimage (\<lambda> h. plus_game G h) (right_options H)))"
+by auto
+termination by (relation "gprod_2_2 option_of")
+  (simp, auto simp: gprod_2_2_def)
 
-declare plus_game.simps[simp del]
-
-lemma plus_game_comm: "plus_game (G, H) = plus_game (H, G)"
+lemma plus_game_comm: "plus_game G H = plus_game H G"
 proof (induct G H rule: plus_game.induct)
   case (1 G H)
   show ?case
@@ -540,11 +538,11 @@
 lemma right_zero_game[simp]: "right_options (zero_game) = zempty"
   by (simp add: zero_game_def)
 
-lemma plus_game_zero_right[simp]: "plus_game (G, zero_game) = G"
+lemma plus_game_zero_right[simp]: "plus_game G zero_game = G"
 proof -
   { 
     fix G H
-    have "H = zero_game \<longrightarrow> plus_game (G, H) = G "
+    have "H = zero_game \<longrightarrow> plus_game G H = G "
     proof (induct G H rule: plus_game.induct, rule impI)
       case (goal1 G H)
       note induct_hyp = prems[simplified goal1, simplified] and prems
@@ -552,7 +550,7 @@
         apply (simp only: plus_game.simps[where G=G and H=H])
         apply (simp add: game_ext_eq prems)
         apply (auto simp add: 
-          zimage_cong[where f = "\<lambda> g. plus_game (g, zero_game)" and g = "id"] 
+          zimage_cong[where f = "\<lambda> g. plus_game g zero_game" and g = "id"] 
           induct_hyp)
         done
     qed
@@ -560,7 +558,7 @@
   then show ?thesis by auto
 qed
 
-lemma plus_game_zero_left: "plus_game (zero_game, G) = G"
+lemma plus_game_zero_left: "plus_game zero_game G = G"
   by (simp add: plus_game_comm)
 
 lemma left_imp_options[simp]: "zin opt (left_options g) \<Longrightarrow> zin opt (options g)"
@@ -570,11 +568,11 @@
   by (simp add: options_def zunion)
 
 lemma left_options_plus: 
-  "left_options (plus_game (u, v)) =  zunion (zimage (\<lambda>g. plus_game (g, v)) (left_options u)) (zimage (\<lambda>h. plus_game (u, h)) (left_options v))" 
+  "left_options (plus_game u v) =  zunion (zimage (\<lambda>g. plus_game g v) (left_options u)) (zimage (\<lambda>h. plus_game u h) (left_options v))" 
   by (subst plus_game.simps, simp)
 
 lemma right_options_plus:
-  "right_options (plus_game (u, v)) =  zunion (zimage (\<lambda>g. plus_game (g, v)) (right_options u)) (zimage (\<lambda>h. plus_game (u, h)) (right_options v))"
+  "right_options (plus_game u v) =  zunion (zimage (\<lambda>g. plus_game g v) (right_options u)) (zimage (\<lambda>h. plus_game u h) (right_options v))"
   by (subst plus_game.simps, simp)
 
 lemma left_options_neg: "left_options (neg_game u) = zimage neg_game (right_options u)"  
@@ -583,32 +581,32 @@
 lemma right_options_neg: "right_options (neg_game u) = zimage neg_game (left_options u)"
   by (subst neg_game.simps, simp)
   
-lemma plus_game_assoc: "plus_game (plus_game (F, G), H) = plus_game (F, plus_game (G, H))"
+lemma plus_game_assoc: "plus_game (plus_game F G) H = plus_game F (plus_game G H)"
 proof -
   { 
     fix a
-    have "\<forall> F G H. a = [F, G, H] \<longrightarrow> plus_game (plus_game (F, G), H) = plus_game (F, plus_game (G, H))"
+    have "\<forall> F G H. a = [F, G, H] \<longrightarrow> plus_game (plus_game F G) H = plus_game F (plus_game G H)"
     proof (induct a rule: induct_game, (rule impI | rule allI)+)
       case (goal1 x F G H)
-      let ?L = "plus_game (plus_game (F, G), H)"
-      let ?R = "plus_game (F, plus_game (G, H))"
+      let ?L = "plus_game (plus_game F G) H"
+      let ?R = "plus_game F (plus_game G H)"
       note options_plus = left_options_plus right_options_plus
       {
         fix opt
         note hyp = goal1(1)[simplified goal1(2), rule_format] 
-        have F: "zin opt (options F)  \<Longrightarrow> plus_game (plus_game (opt, G), H) = plus_game (opt, plus_game (G, H))"
+        have F: "zin opt (options F)  \<Longrightarrow> plus_game (plus_game opt G) H = plus_game opt (plus_game G H)"
           by (blast intro: hyp lprod_3_3)
-        have G: "zin opt (options G) \<Longrightarrow> plus_game (plus_game (F, opt), H) = plus_game (F, plus_game (opt, H))"
+        have G: "zin opt (options G) \<Longrightarrow> plus_game (plus_game F opt) H = plus_game F (plus_game opt H)"
           by (blast intro: hyp lprod_3_4)
-        have H: "zin opt (options H) \<Longrightarrow> plus_game (plus_game (F, G), opt) = plus_game (F, plus_game (G, opt))" 
+        have H: "zin opt (options H) \<Longrightarrow> plus_game (plus_game F G) opt = plus_game F (plus_game G opt)" 
           by (blast intro: hyp lprod_3_5)
         note F and G and H
       }
       note induct_hyp = this
       have "left_options ?L = left_options ?R \<and> right_options ?L = right_options ?R"
         by (auto simp add: 
-          plus_game.simps[where G="plus_game (F,G)" and H=H]
-          plus_game.simps[where G="F" and H="plus_game (G,H)"] 
+          plus_game.simps[where G="plus_game F G" and H=H]
+          plus_game.simps[where G="F" and H="plus_game G H"] 
           zet_ext_eq zunion zimage_iff options_plus
           induct_hyp left_imp_options right_imp_options)
       then show ?case
@@ -618,7 +616,7 @@
   then show ?thesis by auto
 qed
 
-lemma neg_plus_game: "neg_game (plus_game (G, H)) = plus_game(neg_game G, neg_game H)"
+lemma neg_plus_game: "neg_game (plus_game G H) = plus_game (neg_game G) (neg_game H)"
 proof (induct G H rule: plus_game.induct)
   case (1 G H)
   note opt_ops = 
@@ -626,26 +624,26 @@
     left_options_neg right_options_neg  
   show ?case
     by (auto simp add: opt_ops
-      neg_game.simps[of "plus_game (G,H)"]
+      neg_game.simps[of "plus_game G H"]
       plus_game.simps[of "neg_game G" "neg_game H"]
       Game_ext zet_ext_eq zunion zimage_iff prems)
 qed
 
-lemma eq_game_plus_inverse: "eq_game (plus_game (x, neg_game x)) zero_game"
+lemma eq_game_plus_inverse: "eq_game (plus_game x (neg_game x)) zero_game"
 proof (induct x rule: wf_induct[OF wf_option_of])
   case (goal1 x)
   { fix y
     assume "zin y (options x)"
-    then have "eq_game (plus_game (y, neg_game y)) zero_game"
+    then have "eq_game (plus_game y (neg_game y)) zero_game"
       by (auto simp add: prems)
   }
   note ihyp = this
   {
     fix y
     assume y: "zin y (right_options x)"
-    have "\<not> (ge_game (zero_game, plus_game (y, neg_game x)))"
+    have "\<not> (ge_game (zero_game, plus_game y (neg_game x)))"
       apply (subst ge_game.simps, simp)
-      apply (rule exI[where x="plus_game (y, neg_game y)"])
+      apply (rule exI[where x="plus_game y (neg_game y)"])
       apply (auto simp add: ihyp[of y, simplified y right_imp_options eq_game_def])
       apply (auto simp add: left_options_plus left_options_neg zunion zimage_iff intro: prems)
       done
@@ -654,9 +652,9 @@
   {
     fix y
     assume y: "zin y (left_options x)"
-    have "\<not> (ge_game (zero_game, plus_game (x, neg_game y)))"
+    have "\<not> (ge_game (zero_game, plus_game x (neg_game y)))"
       apply (subst ge_game.simps, simp)
-      apply (rule exI[where x="plus_game (y, neg_game y)"])
+      apply (rule exI[where x="plus_game y (neg_game y)"])
       apply (auto simp add: ihyp[of y, simplified y left_imp_options eq_game_def])
       apply (auto simp add: left_options_plus zunion zimage_iff intro: prems)
       done
@@ -665,9 +663,9 @@
   {
     fix y
     assume y: "zin y (left_options x)"
-    have "\<not> (ge_game (plus_game (y, neg_game x), zero_game))"
+    have "\<not> (ge_game (plus_game y (neg_game x), zero_game))"
       apply (subst ge_game.simps, simp)
-      apply (rule exI[where x="plus_game (y, neg_game y)"])
+      apply (rule exI[where x="plus_game y (neg_game y)"])
       apply (auto simp add: ihyp[of y, simplified y left_imp_options eq_game_def])
       apply (auto simp add: right_options_plus right_options_neg zunion zimage_iff intro: prems)
       done
@@ -676,9 +674,9 @@
   {
     fix y
     assume y: "zin y (right_options x)"
-    have "\<not> (ge_game (plus_game (x, neg_game y), zero_game))"
+    have "\<not> (ge_game (plus_game x (neg_game y), zero_game))"
       apply (subst ge_game.simps, simp)
-      apply (rule exI[where x="plus_game (y, neg_game y)"])
+      apply (rule exI[where x="plus_game y (neg_game y)"])
       apply (auto simp add: ihyp[of y, simplified y right_imp_options eq_game_def])
       apply (auto simp add: right_options_plus zunion zimage_iff intro: prems)
       done
@@ -686,28 +684,28 @@
   note case4 = this
   show ?case
     apply (simp add: eq_game_def)
-    apply (simp add: ge_game.simps[of "plus_game (x, neg_game x)" "zero_game"])
-    apply (simp add: ge_game.simps[of "zero_game" "plus_game (x, neg_game x)"])
+    apply (simp add: ge_game.simps[of "plus_game x (neg_game x)" "zero_game"])
+    apply (simp add: ge_game.simps[of "zero_game" "plus_game x (neg_game x)"])
     apply (simp add: right_options_plus left_options_plus right_options_neg left_options_neg zunion zimage_iff)
     apply (auto simp add: case1 case2 case3 case4)
     done
 qed
 
-lemma ge_plus_game_left: "ge_game (y,z) = ge_game(plus_game (x, y), plus_game (x, z))"
+lemma ge_plus_game_left: "ge_game (y,z) = ge_game (plus_game x y, plus_game x z)"
 proof -
   { fix a
-    have "\<forall> x y z. a = [x,y,z] \<longrightarrow> ge_game (y,z) = ge_game(plus_game (x, y), plus_game (x, z))"
+    have "\<forall> x y z. a = [x,y,z] \<longrightarrow> ge_game (y,z) = ge_game (plus_game x y, plus_game x z)"
     proof (induct a rule: induct_game, (rule impI | rule allI)+)
       case (goal1 a x y z)
       note induct_hyp = goal1(1)[rule_format, simplified goal1(2)]
       { 
-        assume hyp: "ge_game(plus_game (x, y), plus_game (x, z))"
+        assume hyp: "ge_game(plus_game x y, plus_game x z)"
         have "ge_game (y, z)"
         proof -
           { fix yr
             assume yr: "zin yr (right_options y)"
-            from hyp have "\<not> (ge_game (plus_game (x, z), plus_game (x, yr)))"
-              by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
+            from hyp have "\<not> (ge_game (plus_game x z, plus_game x yr))"
+              by (auto simp add: ge_game_eq[of "plus_game x y" "plus_game x z"]
                 right_options_plus zunion zimage_iff intro: yr)
             then have "\<not> (ge_game (z, yr))"
               apply (subst induct_hyp[where y="[x, z, yr]", of "x" "z" "yr"])
@@ -717,8 +715,8 @@
           note yr = this
           { fix zl
             assume zl: "zin zl (left_options z)"
-            from hyp have "\<not> (ge_game (plus_game (x, zl), plus_game (x, y)))"
-              by (auto simp add: ge_game_eq[of "plus_game (x,y)" "plus_game(x,z)"]
+            from hyp have "\<not> (ge_game (plus_game x zl, plus_game x y))"
+              by (auto simp add: ge_game_eq[of "plus_game x y" "plus_game x z"]
                 left_options_plus zunion zimage_iff intro: zl)
             then have "\<not> (ge_game (zl, y))"
               apply (subst goal1(1)[rule_format, where y="[x, zl, y]", of "x" "zl" "y"])
@@ -738,11 +736,11 @@
         {
           fix x'
           assume x': "zin x' (right_options x)"
-          assume hyp: "ge_game (plus_game (x, z), plus_game (x', y))"
-          then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
-            by (auto simp add: ge_game_eq[of "plus_game (x,z)" "plus_game (x', y)"] 
+          assume hyp: "ge_game (plus_game x z, plus_game x' y)"
+          then have n: "\<not> (ge_game (plus_game x' y, plus_game x' z))"
+            by (auto simp add: ge_game_eq[of "plus_game x z" "plus_game x' y"] 
               right_options_plus zunion zimage_iff intro: x')
-          have t: "ge_game (plus_game (x', y), plus_game (x', z))"
+          have t: "ge_game (plus_game x' y, plus_game x' z)"
             apply (subst induct_hyp[symmetric])
             apply (auto intro: lprod_3_3 x' yz)
             done
@@ -752,11 +750,11 @@
         {
           fix x'
           assume x': "zin x' (left_options x)"
-          assume hyp: "ge_game (plus_game (x', z), plus_game (x, y))"
-          then have n: "\<not> (ge_game (plus_game (x', y), plus_game (x', z)))"
-            by (auto simp add: ge_game_eq[of "plus_game (x',z)" "plus_game (x, y)"] 
+          assume hyp: "ge_game (plus_game x' z, plus_game x y)"
+          then have n: "\<not> (ge_game (plus_game x' y, plus_game x' z))"
+            by (auto simp add: ge_game_eq[of "plus_game x' z" "plus_game x y"] 
               left_options_plus zunion zimage_iff intro: x')
-          have t: "ge_game (plus_game (x', y), plus_game (x', z))"
+          have t: "ge_game (plus_game x' y, plus_game x' z)"
             apply (subst induct_hyp[symmetric])
             apply (auto intro: lprod_3_3 x' yz)
             done
@@ -766,7 +764,7 @@
         {
           fix y'
           assume y': "zin y' (right_options y)"
-          assume hyp: "ge_game (plus_game(x, z), plus_game (x, y'))"
+          assume hyp: "ge_game (plus_game x z, plus_game x y')"
           then have "ge_game(z, y')"
             apply (subst induct_hyp[of "[x, z, y']" "x" "z" "y'"])
             apply (auto simp add: hyp lprod_3_6 y')
@@ -779,7 +777,7 @@
         {
           fix z'
           assume z': "zin z' (left_options z)"
-          assume hyp: "ge_game (plus_game(x, z'), plus_game (x, y))"
+          assume hyp: "ge_game (plus_game x z', plus_game x y)"
           then have "ge_game(z', y)"
             apply (subst induct_hyp[of "[x, z', y]" "x" "z'" "y"])
             apply (auto simp add: hyp lprod_3_7 z')
@@ -789,7 +787,7 @@
           with z' have "False" by (auto simp add: ge_game_leftright_refl)
         }
         note case4 = this   
-        have "ge_game(plus_game (x, y), plus_game (x, z))"
+        have "ge_game(plus_game x y, plus_game x z)"
           apply (subst ge_game_eq)
           apply (auto simp add: right_options_plus left_options_plus zunion zimage_iff)
           apply (auto intro: case1 case2 case3 case4)
@@ -803,7 +801,7 @@
   then show ?thesis by blast
 qed
 
-lemma ge_plus_game_right: "ge_game (y,z) = ge_game(plus_game (y, x), plus_game (z, x))"
+lemma ge_plus_game_right: "ge_game (y,z) = ge_game(plus_game y x, plus_game z x)"
   by (simp add: ge_plus_game_left plus_game_comm)
 
 lemma ge_neg_game: "ge_game (neg_game x, neg_game y) = ge_game (y, x)"
@@ -838,8 +836,7 @@
   then show ?thesis by blast
 qed
 
-constdefs 
-  eq_game_rel :: "(game * game) set"
+definition eq_game_rel :: "(game * game) set" where
   "eq_game_rel \<equiv> { (p, q) . eq_game p q }"
 
 typedef Pg = "UNIV//eq_game_rel"
@@ -865,7 +862,7 @@
   Pg_minus_def: "- G = contents (\<Union> g \<in> Rep_Pg G. {Abs_Pg (eq_game_rel `` {neg_game g})})"
 
 definition
-  Pg_plus_def: "G + H = contents (\<Union> g \<in> Rep_Pg G. \<Union> h \<in> Rep_Pg H. {Abs_Pg (eq_game_rel `` {plus_game (g,h)})})"
+  Pg_plus_def: "G + H = contents (\<Union> g \<in> Rep_Pg G. \<Union> h \<in> Rep_Pg H. {Abs_Pg (eq_game_rel `` {plus_game g h})})"
 
 definition
   Pg_diff_def: "G - H = G + (- (H::Pg))"
@@ -891,14 +888,14 @@
   apply (simp add: eq_game_rel_def)
   done
 
-lemma char_Pg_plus[simp]: "Abs_Pg (eq_game_rel `` {g}) + Abs_Pg (eq_game_rel `` {h}) = Abs_Pg (eq_game_rel `` {plus_game (g, h)})"
+lemma char_Pg_plus[simp]: "Abs_Pg (eq_game_rel `` {g}) + Abs_Pg (eq_game_rel `` {h}) = Abs_Pg (eq_game_rel `` {plus_game g h})"
 proof -
-  have "(\<lambda> g h. {Abs_Pg (eq_game_rel `` {plus_game (g, h)})}) respects2 eq_game_rel" 
+  have "(\<lambda> g h. {Abs_Pg (eq_game_rel `` {plus_game g h})}) respects2 eq_game_rel" 
     apply (simp add: congruent2_def)
     apply (auto simp add: eq_game_rel_def eq_game_def)
-    apply (rule_tac y="plus_game (y1, z2)" in ge_game_trans)
+    apply (rule_tac y="plus_game y1 z2" in ge_game_trans)
     apply (simp add: ge_plus_game_left[symmetric] ge_plus_game_right[symmetric])+
-    apply (rule_tac y="plus_game (z1, y2)" in ge_game_trans)
+    apply (rule_tac y="plus_game z1 y2" in ge_game_trans)
     apply (simp add: ge_plus_game_left[symmetric] ge_plus_game_right[symmetric])+
     done
   then show ?thesis
--- a/src/HOL/ZF/HOLZF.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ZF/HOLZF.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -19,16 +19,19 @@
   Repl :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF) \<Rightarrow> ZF" and
   Inf :: ZF
 
-constdefs
-  Upair:: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+definition Upair :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "Upair a b == Repl (Power (Power Empty)) (% x. if x = Empty then a else b)"
-  Singleton:: "ZF \<Rightarrow> ZF"
+
+definition Singleton:: "ZF \<Rightarrow> ZF" where
   "Singleton x == Upair x x"
-  union :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+
+definition union :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "union A B == Sum (Upair A B)"
-  SucNat:: "ZF \<Rightarrow> ZF"
+
+definition SucNat:: "ZF \<Rightarrow> ZF" where
   "SucNat x == union x (Singleton x)"
-  subset :: "ZF \<Rightarrow> ZF \<Rightarrow> bool"
+
+definition subset :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" where
   "subset A B == ! x. Elem x A \<longrightarrow> Elem x B"
 
 axioms
@@ -40,8 +43,7 @@
   Regularity: "A \<noteq> Empty \<longrightarrow> (? x. Elem x A & (! y. Elem y x \<longrightarrow> Not (Elem y A)))"
   Infinity: "Elem Empty Inf & (! x. Elem x Inf \<longrightarrow> Elem (SucNat x) Inf)"
 
-constdefs
-  Sep:: "ZF \<Rightarrow> (ZF \<Rightarrow> bool) \<Rightarrow> ZF"
+definition Sep :: "ZF \<Rightarrow> (ZF \<Rightarrow> bool) \<Rightarrow> ZF" where
   "Sep A p == (if (!x. Elem x A \<longrightarrow> Not (p x)) then Empty else 
   (let z = (\<some> x. Elem x A & p x) in
    let f = % x. (if p x then x else z) in Repl A f))" 
@@ -70,8 +72,7 @@
 lemma Singleton: "Elem x (Singleton y) = (x = y)"
   by (simp add: Singleton_def Upair)
 
-constdefs 
-  Opair:: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+definition Opair :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "Opair a b == Upair (Upair a a) (Upair a b)"
 
 lemma Upair_singleton: "(Upair a a = Upair c d) = (a = c & a = d)"
@@ -99,17 +100,16 @@
       done
   qed
 
-constdefs 
-  Replacement :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF option) \<Rightarrow> ZF"
+definition Replacement :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF option) \<Rightarrow> ZF" where
   "Replacement A f == Repl (Sep A (% a. f a \<noteq> None)) (the o f)"
 
 theorem Replacement: "Elem y (Replacement A f) = (? x. Elem x A & f x = Some y)"
   by (auto simp add: Replacement_def Repl Sep) 
 
-constdefs
-  Fst :: "ZF \<Rightarrow> ZF"
+definition Fst :: "ZF \<Rightarrow> ZF" where
   "Fst q == SOME x. ? y. q = Opair x y"
-  Snd :: "ZF \<Rightarrow> ZF"
+
+definition Snd :: "ZF \<Rightarrow> ZF" where
   "Snd q == SOME y. ? x. q = Opair x y"
 
 theorem Fst: "Fst (Opair x y) = x"
@@ -124,8 +124,7 @@
   apply (simp_all add: Opair)
   done
 
-constdefs 
-  isOpair :: "ZF \<Rightarrow> bool"
+definition isOpair :: "ZF \<Rightarrow> bool" where
   "isOpair q == ? x y. q = Opair x y"
 
 lemma isOpair: "isOpair (Opair x y) = True"
@@ -134,8 +133,7 @@
 lemma FstSnd: "isOpair x \<Longrightarrow> Opair (Fst x) (Snd x) = x"
   by (auto simp add: isOpair_def Fst Snd)
   
-constdefs
-  CartProd :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+definition CartProd :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "CartProd A B == Sum(Repl A (% a. Repl B (% b. Opair a b)))"
 
 lemma CartProd: "Elem x (CartProd A B) = (? a b. Elem a A & Elem b B & x = (Opair a b))"
@@ -144,8 +142,7 @@
   apply (auto simp add: Repl)
   done
 
-constdefs
-  explode :: "ZF \<Rightarrow> ZF set"
+definition explode :: "ZF \<Rightarrow> ZF set" where
   "explode z == { x. Elem x z }"
 
 lemma explode_Empty: "(explode x = {}) = (x = Empty)"
@@ -163,10 +160,10 @@
 lemma explode_Repl_eq: "explode (Repl A f) = image f (explode A)"
   by (simp add: explode_def Repl image_def)
 
-constdefs
-  Domain :: "ZF \<Rightarrow> ZF"
+definition Domain :: "ZF \<Rightarrow> ZF" where
   "Domain f == Replacement f (% p. if isOpair p then Some (Fst p) else None)"
-  Range :: "ZF \<Rightarrow> ZF"
+
+definition Range :: "ZF \<Rightarrow> ZF" where
   "Range f == Replacement f (% p. if isOpair p then Some (Snd p) else None)"
 
 theorem Domain: "Elem x (Domain f) = (? y. Elem (Opair x y) f)"
@@ -188,20 +185,16 @@
 theorem union: "Elem x (union A B) = (Elem x A | Elem x B)"
   by (auto simp add: union_def Sum Upair)
 
-constdefs
-  Field :: "ZF \<Rightarrow> ZF"
+definition Field :: "ZF \<Rightarrow> ZF" where
   "Field A == union (Domain A) (Range A)"
 
-constdefs
-  app :: "ZF \<Rightarrow> ZF => ZF"    (infixl "\<acute>" 90) --{*function application*} 
+definition app :: "ZF \<Rightarrow> ZF => ZF" (infixl "\<acute>" 90) --{*function application*} where
   "f \<acute> x == (THE y. Elem (Opair x y) f)"
 
-constdefs
-  isFun :: "ZF \<Rightarrow> bool"
+definition isFun :: "ZF \<Rightarrow> bool" where
   "isFun f == (! x y1 y2. Elem (Opair x y1) f & Elem (Opair x y2) f \<longrightarrow> y1 = y2)"
 
-constdefs
-  Lambda :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF) \<Rightarrow> ZF"
+definition Lambda :: "ZF \<Rightarrow> (ZF \<Rightarrow> ZF) \<Rightarrow> ZF" where
   "Lambda A f == Repl A (% x. Opair x (f x))"
 
 lemma Lambda_app: "Elem x A \<Longrightarrow> (Lambda A f)\<acute>x = f x"
@@ -238,10 +231,10 @@
     done
 qed
 
-constdefs 
-  PFun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+definition PFun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "PFun A B == Sep (Power (CartProd A B)) isFun"
-  Fun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF"
+
+definition Fun :: "ZF \<Rightarrow> ZF \<Rightarrow> ZF" where
   "Fun A B == Sep (PFun A B) (\<lambda> f. Domain f = A)"
 
 lemma Fun_Range: "Elem f (Fun U V) \<Longrightarrow> subset (Range f) V"
@@ -343,8 +336,7 @@
 qed    
 
 
-constdefs
-  is_Elem_of :: "(ZF * ZF) set"
+definition is_Elem_of :: "(ZF * ZF) set" where
   "is_Elem_of == { (a,b) | a b. Elem a b }"
 
 lemma cond_wf_Elem:
@@ -417,8 +409,7 @@
 nat2Nat_0[intro]:  "nat2Nat 0 = Empty"
 nat2Nat_Suc[intro]:  "nat2Nat (Suc n) = SucNat (nat2Nat n)"
 
-constdefs
-  Nat2nat :: "ZF \<Rightarrow> nat"
+definition Nat2nat :: "ZF \<Rightarrow> nat" where
   "Nat2nat == inv nat2Nat"
 
 lemma Elem_nat2Nat_inf[intro]: "Elem (nat2Nat n) Inf"
@@ -426,9 +417,8 @@
   apply (simp_all add: Infinity)
   done
 
-constdefs
-  Nat :: ZF
-  "Nat == Sep Inf (\<lambda> N. ? n. nat2Nat n = N)"
+definition Nat :: ZF
+ where  "Nat == Sep Inf (\<lambda> N. ? n. nat2Nat n = N)"
 
 lemma Elem_nat2Nat_Nat[intro]: "Elem (nat2Nat n) Nat"
   by (auto simp add: Nat_def Sep)
@@ -664,8 +654,7 @@
   qed
 qed
 
-constdefs 
-  SpecialR :: "(ZF * ZF) set"
+definition SpecialR :: "(ZF * ZF) set" where
   "SpecialR \<equiv> { (x, y) . x \<noteq> Empty \<and> y = Empty}"
 
 lemma "wf SpecialR"
@@ -673,8 +662,7 @@
   apply (auto simp add: SpecialR_def)
   done
 
-constdefs
-  Ext :: "('a * 'b) set \<Rightarrow> 'b \<Rightarrow> 'a set"
+definition Ext :: "('a * 'b) set \<Rightarrow> 'b \<Rightarrow> 'a set" where
   "Ext R y \<equiv> { x . (x, y) \<in> R }" 
 
 lemma Ext_Elem: "Ext is_Elem_of = explode"
@@ -690,8 +678,7 @@
   then show "False" by (simp add: UNIV_is_not_in_ZF)
 qed
 
-constdefs 
-  implode :: "ZF set \<Rightarrow> ZF"
+definition implode :: "ZF set \<Rightarrow> ZF" where
   "implode == inv explode"
 
 lemma inj_explode: "inj explode"
@@ -700,12 +687,13 @@
 lemma implode_explode[simp]: "implode (explode x) = x"
   by (simp add: implode_def inj_explode)
 
-constdefs
-  regular :: "(ZF * ZF) set \<Rightarrow> bool"
+definition regular :: "(ZF * ZF) set \<Rightarrow> bool" where
   "regular R == ! A. A \<noteq> Empty \<longrightarrow> (? x. Elem x A & (! y. (y, x) \<in> R \<longrightarrow> Not (Elem y A)))"
-  set_like :: "(ZF * ZF) set \<Rightarrow> bool"
+
+definition set_like :: "(ZF * ZF) set \<Rightarrow> bool" where
   "set_like R == ! y. Ext R y \<in> range explode"
-  wfzf :: "(ZF * ZF) set \<Rightarrow> bool"
+
+definition wfzf :: "(ZF * ZF) set \<Rightarrow> bool" where
   "wfzf R == regular R & set_like R"
 
 lemma regular_Elem: "regular is_Elem_of"
@@ -717,8 +705,7 @@
 lemma wfzf_is_Elem_of: "wfzf is_Elem_of"
   by (auto simp add: wfzf_def regular_Elem set_like_Elem)
 
-constdefs
-  SeqSum :: "(nat \<Rightarrow> ZF) \<Rightarrow> ZF"
+definition SeqSum :: "(nat \<Rightarrow> ZF) \<Rightarrow> ZF" where
   "SeqSum f == Sum (Repl Nat (f o Nat2nat))"
 
 lemma SeqSum: "Elem x (SeqSum f) = (? n. Elem x (f n))"
@@ -727,8 +714,7 @@
   apply auto
   done
 
-constdefs
-  Ext_ZF :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF"
+definition Ext_ZF :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF" where
   "Ext_ZF R s == implode (Ext R s)"
 
 lemma Elem_implode: "A \<in> range explode \<Longrightarrow> Elem x (implode A) = (x \<in> A)"
@@ -750,8 +736,7 @@
   "Ext_ZF_n R s 0 = Ext_ZF R s"
   "Ext_ZF_n R s (Suc n) = Sum (Repl (Ext_ZF_n R s n) (Ext_ZF R))"
 
-constdefs
-  Ext_ZF_hull :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF"
+definition Ext_ZF_hull :: "(ZF * ZF) set \<Rightarrow> ZF \<Rightarrow> ZF" where
   "Ext_ZF_hull R s == SeqSum (Ext_ZF_n R s)"
 
 lemma Elem_Ext_ZF_hull:
--- a/src/HOL/ZF/LProd.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ZF/LProd.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ZF/LProd.thy
-    ID:         $Id$
     Author:     Steven Obua
 
     Introduces the lprod relation.
@@ -95,10 +94,10 @@
   show ?thesis by (auto intro: lprod)
 qed
 
-constdefs
-  gprod_2_2 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set"
+definition gprod_2_2 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set" where
   "gprod_2_2 R \<equiv> { ((a,b), (c,d)) . (a = c \<and> (b,d) \<in> R) \<or> (b = d \<and> (a,c) \<in> R) }"
-  gprod_2_1 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set"
+
+definition gprod_2_1 :: "('a * 'a) set \<Rightarrow> (('a * 'a) * ('a * 'a)) set" where
   "gprod_2_1 R \<equiv>  { ((a,b), (c,d)) . (a = d \<and> (b,c) \<in> R) \<or> (b = c \<and> (a,d) \<in> R) }"
 
 lemma lprod_2_3: "(a, b) \<in> R \<Longrightarrow> ([a, c], [b, c]) \<in> lprod R"
@@ -170,8 +169,7 @@
   apply (simp add: z' lprod_2_4)
   done
 
-constdefs
-   perm :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> bool"
+definition perm :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> bool" where
    "perm f A \<equiv> inj_on f A \<and> f ` A = A"
 
 lemma "((as,bs) \<in> lprod R) = 
@@ -183,6 +181,4 @@
 lemma "trans R \<Longrightarrow> (ah@a#at, bh@b#bt) \<in> lprod R \<Longrightarrow> (b, a) \<in> R \<or> a = b \<Longrightarrow> (ah@at, bh@bt) \<in> lprod R" 
 oops
 
-
-
-end
\ No newline at end of file
+end
--- a/src/HOL/ZF/MainZF.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ZF/MainZF.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ZF/MainZF.thy
-    ID:         $Id$
     Author:     Steven Obua
 
     Starting point for using HOLZF.
@@ -9,4 +8,5 @@
 theory MainZF
 imports Zet LProd
 begin
-end
\ No newline at end of file
+
+end
--- a/src/HOL/ZF/Zet.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ZF/Zet.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ZF/Zet.thy
-    ID:         $Id$
     Author:     Steven Obua
 
     Introduces a type 'a zet of ZF representable sets.
@@ -13,15 +12,13 @@
 typedef 'a zet = "{A :: 'a set | A f z. inj_on f A \<and> f ` A \<subseteq> explode z}"
   by blast
 
-constdefs
-  zin :: "'a \<Rightarrow> 'a zet \<Rightarrow> bool"
+definition zin :: "'a \<Rightarrow> 'a zet \<Rightarrow> bool" where
   "zin x A == x \<in> (Rep_zet A)"
 
 lemma zet_ext_eq: "(A = B) = (! x. zin x A = zin x B)"
   by (auto simp add: Rep_zet_inject[symmetric] zin_def)
 
-constdefs
-  zimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a zet \<Rightarrow> 'b zet"
+definition zimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a zet \<Rightarrow> 'b zet" where
   "zimage f A == Abs_zet (image f (Rep_zet A))"
 
 lemma zet_def': "zet = {A :: 'a set | A f z. inj_on f A \<and> f ` A = explode z}"
@@ -74,10 +71,10 @@
 lemma zimage_iff: "zin y (zimage f A) = (? x. zin x A & y = f x)"
   by (auto simp add: zin_def Rep_zimage_eq)
 
-constdefs
-  zimplode :: "ZF zet \<Rightarrow> ZF"
+definition zimplode :: "ZF zet \<Rightarrow> ZF" where
   "zimplode A == implode (Rep_zet A)"
-  zexplode :: "ZF \<Rightarrow> ZF zet"
+
+definition zexplode :: "ZF \<Rightarrow> ZF zet" where
   "zexplode z == Abs_zet (explode z)"
 
 lemma Rep_zet_eq_explode: "? z. Rep_zet A = explode z"
@@ -114,10 +111,10 @@
   apply (simp_all add: comp_image_eq zet_image_mem Rep_zet)
   done
     
-constdefs
-  zunion :: "'a zet \<Rightarrow> 'a zet \<Rightarrow> 'a zet"
+definition zunion :: "'a zet \<Rightarrow> 'a zet \<Rightarrow> 'a zet" where
   "zunion a b \<equiv> Abs_zet ((Rep_zet a) \<union> (Rep_zet b))"
-  zsubset :: "'a zet \<Rightarrow> 'a zet \<Rightarrow> bool"
+
+definition zsubset :: "'a zet \<Rightarrow> 'a zet \<Rightarrow> bool" where
   "zsubset a b \<equiv> ! x. zin x a \<longrightarrow> zin x b"
 
 lemma explode_union: "explode (union a b) = (explode a) \<union> (explode b)"
@@ -181,8 +178,7 @@
   apply (simp_all add: zin_def Rep_zet range_explode_eq_zet)
   done
 
-constdefs
-  zempty :: "'a zet"
+definition zempty :: "'a zet" where
   "zempty \<equiv> Abs_zet {}"
 
 lemma zempty[simp]: "\<not> (zin x zempty)"
@@ -200,7 +196,7 @@
 lemma zimage_id[simp]: "zimage id A = A"
   by (simp add: zet_ext_eq zimage_iff)
 
-lemma zimage_cong[recdef_cong]: "\<lbrakk> M = N; !! x. zin x N \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> zimage f M = zimage g N"
+lemma zimage_cong[recdef_cong, fundef_cong]: "\<lbrakk> M = N; !! x. zin x N \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow> zimage f M = zimage g N"
   by (auto simp add: zet_ext_eq zimage_iff)
 
 end
--- a/src/HOL/ex/NormalForm.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/NormalForm.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -3,7 +3,7 @@
 header {* Testing implementation of normalization by evaluation *}
 
 theory NormalForm
-imports Main Rational
+imports Complex_Main
 begin
 
 lemma "True" by normalization
--- a/src/HOL/ex/Numeral.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/Numeral.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -327,7 +327,7 @@
       val k = int_of_num' n;
       val t' = Syntax.const @{syntax_const "_Numerals"} $ Syntax.free ("#" ^ string_of_int k);
     in case T
-     of Type ("fun", [_, T']) =>  (* FIXME @{type_syntax} *)
+     of Type (@{type_syntax fun}, [_, T']) =>
          if not (! show_types) andalso can Term.dest_Type T' then t'
          else Syntax.const Syntax.constrainC $ t' $ Syntax.term_of_typ show_sorts T'
       | T' => if T' = dummyT then t' else raise Match
--- a/src/HOL/ex/ReflectionEx.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/ReflectionEx.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -69,28 +69,29 @@
 oops
 
   text{* Let's perform NNF. This is a version that tends to generate disjunctions *}
-consts fmsize :: "fm \<Rightarrow> nat"
-primrec
+primrec fmsize :: "fm \<Rightarrow> nat" where
   "fmsize (At n) = 1"
-  "fmsize (NOT p) = 1 + fmsize p"
-  "fmsize (And p q) = 1 + fmsize p + fmsize q"
-  "fmsize (Or p q) = 1 + fmsize p + fmsize q"
-  "fmsize (Imp p q) = 2 + fmsize p + fmsize q"
-  "fmsize (Iff p q) = 2 + 2* fmsize p + 2* fmsize q"
+| "fmsize (NOT p) = 1 + fmsize p"
+| "fmsize (And p q) = 1 + fmsize p + fmsize q"
+| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
+| "fmsize (Imp p q) = 2 + fmsize p + fmsize q"
+| "fmsize (Iff p q) = 2 + 2* fmsize p + 2* fmsize q"
+
+lemma [measure_function]: "is_measure fmsize" ..
 
-consts nnf :: "fm \<Rightarrow> fm"
-recdef nnf "measure fmsize"
+fun nnf :: "fm \<Rightarrow> fm"
+where
   "nnf (At n) = At n"
-  "nnf (And p q) = And (nnf p) (nnf q)"
-  "nnf (Or p q) = Or (nnf p) (nnf q)"
-  "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)"
-  "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (NOT p)) (nnf (NOT q)))"
-  "nnf (NOT (And p q)) = Or (nnf (NOT p)) (nnf (NOT q))"
-  "nnf (NOT (Or p q)) = And (nnf (NOT p)) (nnf (NOT q))"
-  "nnf (NOT (Imp p q)) = And (nnf p) (nnf (NOT q))"
-  "nnf (NOT (Iff p q)) = Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q))"
-  "nnf (NOT (NOT p)) = nnf p"
-  "nnf (NOT p) = NOT p"
+| "nnf (And p q) = And (nnf p) (nnf q)"
+| "nnf (Or p q) = Or (nnf p) (nnf q)"
+| "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)"
+| "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (NOT p)) (nnf (NOT q)))"
+| "nnf (NOT (And p q)) = Or (nnf (NOT p)) (nnf (NOT q))"
+| "nnf (NOT (Or p q)) = And (nnf (NOT p)) (nnf (NOT q))"
+| "nnf (NOT (Imp p q)) = And (nnf p) (nnf (NOT q))"
+| "nnf (NOT (Iff p q)) = Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q))"
+| "nnf (NOT (NOT p)) = nnf p"
+| "nnf (NOT p) = NOT p"
 
   text{* The correctness theorem of nnf: it preserves the semantics of fm *}
 lemma nnf[reflection]: "Ifm (nnf p) vs = Ifm p vs"
@@ -113,22 +114,22 @@
 datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num
 
 text{* This is just technical to make recursive definitions easier. *}
-consts num_size :: "num \<Rightarrow> nat" 
-primrec 
+primrec num_size :: "num \<Rightarrow> nat" 
+where
   "num_size (C c) = 1"
-  "num_size (Var n) = 1"
-  "num_size (Add a b) = 1 + num_size a + num_size b"
-  "num_size (Mul c a) = 1 + num_size a"
-  "num_size (CN n c a) = 4 + num_size a "
+| "num_size (Var n) = 1"
+| "num_size (Add a b) = 1 + num_size a + num_size b"
+| "num_size (Mul c a) = 1 + num_size a"
+| "num_size (CN n c a) = 4 + num_size a "
 
   text{* The semantics of num *}
-consts Inum:: "num \<Rightarrow> nat list \<Rightarrow> nat"
-primrec 
+primrec Inum:: "num \<Rightarrow> nat list \<Rightarrow> nat"
+where
   Inum_C  : "Inum (C i) vs = i"
-  Inum_Var: "Inum (Var n) vs = vs!n"
-  Inum_Add: "Inum (Add s t) vs = Inum s vs + Inum t vs "
-  Inum_Mul: "Inum (Mul c t) vs = c * Inum t vs "
-  Inum_CN : "Inum (CN n c t) vs = c*(vs!n) + Inum t vs "
+| Inum_Var: "Inum (Var n) vs = vs!n"
+| Inum_Add: "Inum (Add s t) vs = Inum s vs + Inum t vs "
+| Inum_Mul: "Inum (Mul c t) vs = c * Inum t vs "
+| Inum_CN : "Inum (CN n c t) vs = c*(vs!n) + Inum t vs "
 
 
   text{* Let's reify some nat expressions \dots *}
@@ -168,39 +169,41 @@
   apply (reify Inum_eqs' ("1 * (2*x + (y::nat) + 0 + 1)"))
 oops
 text{* Okay, let's try reflection. Some simplifications on num follow. You can skim until the main theorem @{text linum} *}
-consts lin_add :: "num \<times> num \<Rightarrow> num"
-recdef lin_add "measure (\<lambda>(x,y). ((size x) + (size y)))"
-  "lin_add (CN n1 c1 r1,CN n2 c2 r2) =
+fun lin_add :: "num \<Rightarrow> num \<Rightarrow> num"
+where
+  "lin_add (CN n1 c1 r1) (CN n2 c2 r2) =
   (if n1=n2 then 
   (let c = c1 + c2
-  in (if c=0 then lin_add(r1,r2) else CN n1 c (lin_add (r1,r2))))
-  else if n1 \<le> n2 then (CN n1 c1 (lin_add (r1,CN n2 c2 r2))) 
-  else (CN n2 c2 (lin_add (CN n1 c1 r1,r2))))"
-  "lin_add (CN n1 c1 r1,t) = CN n1 c1 (lin_add (r1, t))"  
-  "lin_add (t,CN n2 c2 r2) = CN n2 c2 (lin_add (t,r2))" 
-  "lin_add (C b1, C b2) = C (b1+b2)"
-  "lin_add (a,b) = Add a b"
-lemma lin_add: "Inum (lin_add (t,s)) bs = Inum (Add t s) bs"
+  in (if c=0 then lin_add r1 r2 else CN n1 c (lin_add r1 r2)))
+  else if n1 \<le> n2 then (CN n1 c1 (lin_add r1 (CN n2 c2 r2))) 
+  else (CN n2 c2 (lin_add (CN n1 c1 r1) r2)))"
+| "lin_add (CN n1 c1 r1) t = CN n1 c1 (lin_add r1 t)"  
+| "lin_add t (CN n2 c2 r2) = CN n2 c2 (lin_add t r2)" 
+| "lin_add (C b1) (C b2) = C (b1+b2)"
+| "lin_add a b = Add a b"
+lemma lin_add: "Inum (lin_add t s) bs = Inum (Add t s) bs"
 apply (induct t s rule: lin_add.induct, simp_all add: Let_def)
 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
 by (case_tac "n1 = n2", simp_all add: algebra_simps)
 
-consts lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num"
-recdef lin_mul "measure size "
-  "lin_mul (C j) = (\<lambda> i. C (i*j))"
-  "lin_mul (CN n c a) = (\<lambda> i. if i=0 then (C 0) else CN n (i*c) (lin_mul a i))"
-  "lin_mul t = (\<lambda> i. Mul i t)"
+fun lin_mul :: "num \<Rightarrow> nat \<Rightarrow> num"
+where
+  "lin_mul (C j) i = C (i*j)"
+| "lin_mul (CN n c a) i = (if i=0 then (C 0) else CN n (i*c) (lin_mul a i))"
+| "lin_mul t i = (Mul i t)"
 
 lemma lin_mul: "Inum (lin_mul t i) bs = Inum (Mul i t) bs"
-by (induct t arbitrary: i rule: lin_mul.induct, auto simp add: algebra_simps)
+by (induct t i rule: lin_mul.induct, auto simp add: algebra_simps)
 
-consts linum:: "num \<Rightarrow> num"
-recdef linum "measure num_size"
+lemma [measure_function]: "is_measure num_size" ..
+
+fun linum:: "num \<Rightarrow> num"
+where
   "linum (C b) = C b"
-  "linum (Var n) = CN n 1 (C 0)"
-  "linum (Add t s) = lin_add (linum t, linum s)"
-  "linum (Mul c t) = lin_mul (linum t) c"
-  "linum (CN n c t) = lin_add (linum (Mul c (Var n)),linum t)"
+| "linum (Var n) = CN n 1 (C 0)"
+| "linum (Add t s) = lin_add (linum t) (linum s)"
+| "linum (Mul c t) = lin_mul (linum t) c"
+| "linum (CN n c t) = lin_add (linum (Mul c (Var n))) (linum t)"
 
 lemma linum[reflection] : "Inum (linum t) bs = Inum t bs"
 by (induct t rule: linum.induct, simp_all add: lin_mul lin_add)
@@ -214,30 +217,32 @@
 
 datatype aform = Lt num num  | Eq num num | Ge num num | NEq num num | 
   Conj aform aform | Disj aform aform | NEG aform | T | F
-consts linaformsize:: "aform \<Rightarrow> nat"
-recdef linaformsize "measure size"
+
+primrec linaformsize:: "aform \<Rightarrow> nat"
+where
   "linaformsize T = 1"
-  "linaformsize F = 1"
-  "linaformsize (Lt a b) = 1"
-  "linaformsize (Ge a b) = 1"
-  "linaformsize (Eq a b) = 1"
-  "linaformsize (NEq a b) = 1"
-  "linaformsize (NEG p) = 2 + linaformsize p"
-  "linaformsize (Conj p q) = 1 + linaformsize p + linaformsize q"
-  "linaformsize (Disj p q) = 1 + linaformsize p + linaformsize q"
+| "linaformsize F = 1"
+| "linaformsize (Lt a b) = 1"
+| "linaformsize (Ge a b) = 1"
+| "linaformsize (Eq a b) = 1"
+| "linaformsize (NEq a b) = 1"
+| "linaformsize (NEG p) = 2 + linaformsize p"
+| "linaformsize (Conj p q) = 1 + linaformsize p + linaformsize q"
+| "linaformsize (Disj p q) = 1 + linaformsize p + linaformsize q"
 
+lemma [measure_function]: "is_measure linaformsize" ..
 
-consts is_aform :: "aform => nat list => bool"
-primrec
+primrec is_aform :: "aform => nat list => bool"
+where
   "is_aform T vs = True"
-  "is_aform F vs = False"
-  "is_aform (Lt a b) vs = (Inum a vs < Inum b vs)"
-  "is_aform (Eq a b) vs = (Inum a vs = Inum b vs)"
-  "is_aform (Ge a b) vs = (Inum a vs \<ge> Inum b vs)"
-  "is_aform (NEq a b) vs = (Inum a vs \<noteq> Inum b vs)"
-  "is_aform (NEG p) vs = (\<not> (is_aform p vs))"
-  "is_aform (Conj p q) vs = (is_aform p vs \<and> is_aform q vs)"
-  "is_aform (Disj p q) vs = (is_aform p vs \<or> is_aform q vs)"
+| "is_aform F vs = False"
+| "is_aform (Lt a b) vs = (Inum a vs < Inum b vs)"
+| "is_aform (Eq a b) vs = (Inum a vs = Inum b vs)"
+| "is_aform (Ge a b) vs = (Inum a vs \<ge> Inum b vs)"
+| "is_aform (NEq a b) vs = (Inum a vs \<noteq> Inum b vs)"
+| "is_aform (NEG p) vs = (\<not> (is_aform p vs))"
+| "is_aform (Conj p q) vs = (is_aform p vs \<and> is_aform q vs)"
+| "is_aform (Disj p q) vs = (is_aform p vs \<or> is_aform q vs)"
 
   text{* Let's reify and do reflection *}
 lemma "(3::nat)*x + t < 0 \<and> (2 * x + y \<noteq> 17)"
@@ -250,24 +255,25 @@
 oops
 
   text{* For reflection we now define a simple transformation on aform: NNF + linum on atoms *}
-consts linaform:: "aform \<Rightarrow> aform"
-recdef linaform "measure linaformsize"
+
+fun linaform:: "aform \<Rightarrow> aform"
+where
   "linaform (Lt s t) = Lt (linum s) (linum t)"
-  "linaform (Eq s t) = Eq (linum s) (linum t)"
-  "linaform (Ge s t) = Ge (linum s) (linum t)"
-  "linaform (NEq s t) = NEq (linum s) (linum t)"
-  "linaform (Conj p q) = Conj (linaform p) (linaform q)"
-  "linaform (Disj p q) = Disj (linaform p) (linaform q)"
-  "linaform (NEG T) = F"
-  "linaform (NEG F) = T"
-  "linaform (NEG (Lt a b)) = Ge a b"
-  "linaform (NEG (Ge a b)) = Lt a b"
-  "linaform (NEG (Eq a b)) = NEq a b"
-  "linaform (NEG (NEq a b)) = Eq a b"
-  "linaform (NEG (NEG p)) = linaform p"
-  "linaform (NEG (Conj p q)) = Disj (linaform (NEG p)) (linaform (NEG q))"
-  "linaform (NEG (Disj p q)) = Conj (linaform (NEG p)) (linaform (NEG q))"
-  "linaform p = p"
+| "linaform (Eq s t) = Eq (linum s) (linum t)"
+| "linaform (Ge s t) = Ge (linum s) (linum t)"
+| "linaform (NEq s t) = NEq (linum s) (linum t)"
+| "linaform (Conj p q) = Conj (linaform p) (linaform q)"
+| "linaform (Disj p q) = Disj (linaform p) (linaform q)"
+| "linaform (NEG T) = F"
+| "linaform (NEG F) = T"
+| "linaform (NEG (Lt a b)) = Ge a b"
+| "linaform (NEG (Ge a b)) = Lt a b"
+| "linaform (NEG (Eq a b)) = NEq a b"
+| "linaform (NEG (NEq a b)) = Eq a b"
+| "linaform (NEG (NEG p)) = linaform p"
+| "linaform (NEG (Conj p q)) = Disj (linaform (NEG p)) (linaform (NEG q))"
+| "linaform (NEG (Disj p q)) = Conj (linaform (NEG p)) (linaform (NEG q))"
+| "linaform p = p"
 
 lemma linaform: "is_aform (linaform p) vs = is_aform p vs"
   by (induct p rule: linaform.induct) (auto simp add: linum)
@@ -283,11 +289,11 @@
 
 text{* We now give an example where Interpretaions have 0 or more than only one envornement of different types and show that automatic reification also deals with binding *}
 datatype rb = BC bool| BAnd rb rb | BOr rb rb
-consts Irb :: "rb \<Rightarrow> bool"
-primrec
+primrec Irb :: "rb \<Rightarrow> bool"
+where
   "Irb (BC p) = p"
-  "Irb (BAnd s t) = (Irb s \<and> Irb t)"
-  "Irb (BOr s t) = (Irb s \<or> Irb t)"
+| "Irb (BAnd s t) = (Irb s \<and> Irb t)"
+| "Irb (BOr s t) = (Irb s \<or> Irb t)"
 
 lemma "A \<and> (B \<or> D \<and> B) \<and> A \<and> (B \<or> D \<and> B) \<or> A \<and> (B \<or> D \<and> B) \<or> A \<and> (B \<or> D \<and> B)"
 apply (reify Irb.simps)
@@ -295,14 +301,14 @@
 
 
 datatype rint = IC int| IVar nat | IAdd rint rint | IMult rint rint | INeg rint | ISub rint rint
-consts Irint :: "rint \<Rightarrow> int list \<Rightarrow> int"
-primrec
-Irint_Var: "Irint (IVar n) vs = vs!n"
-Irint_Neg: "Irint (INeg t) vs = - Irint t vs"
-Irint_Add: "Irint (IAdd s t) vs = Irint s vs + Irint t vs"
-Irint_Sub: "Irint (ISub s t) vs = Irint s vs - Irint t vs"
-Irint_Mult: "Irint (IMult s t) vs = Irint s vs * Irint t vs"
-Irint_C: "Irint (IC i) vs = i"
+primrec Irint :: "rint \<Rightarrow> int list \<Rightarrow> int"
+where
+  Irint_Var: "Irint (IVar n) vs = vs!n"
+| Irint_Neg: "Irint (INeg t) vs = - Irint t vs"
+| Irint_Add: "Irint (IAdd s t) vs = Irint s vs + Irint t vs"
+| Irint_Sub: "Irint (ISub s t) vs = Irint s vs - Irint t vs"
+| Irint_Mult: "Irint (IMult s t) vs = Irint s vs * Irint t vs"
+| Irint_C: "Irint (IC i) vs = i"
 lemma Irint_C0: "Irint (IC 0) vs = 0"
   by simp
 lemma Irint_C1: "Irint (IC 1) vs = 1"
@@ -314,12 +320,12 @@
   apply (reify Irint_simps ("(3::int) * x + y*y - 9 + (- z)"))
   oops
 datatype rlist = LVar nat| LEmpty| LCons rint rlist | LAppend rlist rlist
-consts Irlist :: "rlist \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> (int list)"
-primrec
+primrec Irlist :: "rlist \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> (int list)"
+where
   "Irlist (LEmpty) is vs = []"
-  "Irlist (LVar n) is vs = vs!n"
-  "Irlist (LCons i t) is vs = ((Irint i is)#(Irlist t is vs))"
-  "Irlist (LAppend s t) is vs = (Irlist s is vs) @ (Irlist t is vs)"
+| "Irlist (LVar n) is vs = vs!n"
+| "Irlist (LCons i t) is vs = ((Irint i is)#(Irlist t is vs))"
+| "Irlist (LAppend s t) is vs = (Irlist s is vs) @ (Irlist t is vs)"
 lemma "[(1::int)] = []"
   apply (reify Irlist.simps Irint_simps ("[1]:: int list"))
   oops
@@ -329,16 +335,16 @@
   oops
 
 datatype rnat = NC nat| NVar nat| NSuc rnat | NAdd rnat rnat | NMult rnat rnat | NNeg rnat | NSub rnat rnat | Nlgth rlist
-consts Irnat :: "rnat \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> nat"
-primrec
-Irnat_Suc: "Irnat (NSuc t) is ls vs = Suc (Irnat t is ls vs)"
-Irnat_Var: "Irnat (NVar n) is ls vs = vs!n"
-Irnat_Neg: "Irnat (NNeg t) is ls vs = 0"
-Irnat_Add: "Irnat (NAdd s t) is ls vs = Irnat s is ls vs + Irnat t is ls vs"
-Irnat_Sub: "Irnat (NSub s t) is ls vs = Irnat s is ls vs - Irnat t is ls vs"
-Irnat_Mult: "Irnat (NMult s t) is ls vs = Irnat s is ls vs * Irnat t is ls vs"
-Irnat_lgth: "Irnat (Nlgth rxs) is ls vs = length (Irlist rxs is ls)"
-Irnat_C: "Irnat (NC i) is ls vs = i"
+primrec Irnat :: "rnat \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> nat"
+where
+  Irnat_Suc: "Irnat (NSuc t) is ls vs = Suc (Irnat t is ls vs)"
+| Irnat_Var: "Irnat (NVar n) is ls vs = vs!n"
+| Irnat_Neg: "Irnat (NNeg t) is ls vs = 0"
+| Irnat_Add: "Irnat (NAdd s t) is ls vs = Irnat s is ls vs + Irnat t is ls vs"
+| Irnat_Sub: "Irnat (NSub s t) is ls vs = Irnat s is ls vs - Irnat t is ls vs"
+| Irnat_Mult: "Irnat (NMult s t) is ls vs = Irnat s is ls vs * Irnat t is ls vs"
+| Irnat_lgth: "Irnat (Nlgth rxs) is ls vs = length (Irlist rxs is ls)"
+| Irnat_C: "Irnat (NC i) is ls vs = i"
 lemma Irnat_C0: "Irnat (NC 0) is ls vs = 0"
 by simp
 lemma Irnat_C1: "Irnat (NC 1) is ls vs = 1"
@@ -356,26 +362,26 @@
   | RNEX rifm | RIEX rifm| RLEX rifm | RNALL rifm | RIALL rifm| RLALL rifm
   | RBEX rifm | RBALL rifm
 
-consts Irifm :: "rifm \<Rightarrow> bool list \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> bool"
-primrec
-"Irifm RT ps is ls ns = True"
-"Irifm RF ps is ls ns = False"
-"Irifm (RVar n) ps is ls ns = ps!n"
-"Irifm (RNLT s t) ps is ls ns = (Irnat s is ls ns < Irnat t is ls ns)"
-"Irifm (RNILT s t) ps is ls ns = (int (Irnat s is ls ns) < Irint t is)"
-"Irifm (RNEQ s t) ps is ls ns = (Irnat s is ls ns = Irnat t is ls ns)"
-"Irifm (RAnd p q) ps is ls ns = (Irifm p ps is ls ns \<and> Irifm q ps is ls ns)"
-"Irifm (ROr p q) ps is ls ns = (Irifm p ps is ls ns \<or> Irifm q ps is ls ns)"
-"Irifm (RImp p q) ps is ls ns = (Irifm p ps is ls ns \<longrightarrow> Irifm q ps is ls ns)"
-"Irifm (RIff p q) ps is ls ns = (Irifm p ps is ls ns = Irifm q ps is ls ns)"
-"Irifm (RNEX p) ps is ls ns =  (\<exists>x. Irifm p ps is ls (x#ns))"
-"Irifm (RIEX p) ps is ls ns =  (\<exists>x. Irifm p ps (x#is) ls ns)"
-"Irifm (RLEX p) ps is ls ns =  (\<exists>x. Irifm p ps is (x#ls) ns)"
-"Irifm (RBEX p) ps is ls ns =  (\<exists>x. Irifm p (x#ps) is ls ns)"
-"Irifm (RNALL p) ps is ls ns = (\<forall>x. Irifm p ps is ls (x#ns))"
-"Irifm (RIALL p) ps is ls ns = (\<forall>x. Irifm p ps (x#is) ls ns)"
-"Irifm (RLALL p) ps is ls ns = (\<forall>x. Irifm p ps is (x#ls) ns)"
-"Irifm (RBALL p) ps is ls ns = (\<forall>x. Irifm p (x#ps) is ls ns)"
+primrec Irifm :: "rifm \<Rightarrow> bool list \<Rightarrow> int list \<Rightarrow> (int list) list \<Rightarrow> nat list \<Rightarrow> bool"
+where
+  "Irifm RT ps is ls ns = True"
+| "Irifm RF ps is ls ns = False"
+| "Irifm (RVar n) ps is ls ns = ps!n"
+| "Irifm (RNLT s t) ps is ls ns = (Irnat s is ls ns < Irnat t is ls ns)"
+| "Irifm (RNILT s t) ps is ls ns = (int (Irnat s is ls ns) < Irint t is)"
+| "Irifm (RNEQ s t) ps is ls ns = (Irnat s is ls ns = Irnat t is ls ns)"
+| "Irifm (RAnd p q) ps is ls ns = (Irifm p ps is ls ns \<and> Irifm q ps is ls ns)"
+| "Irifm (ROr p q) ps is ls ns = (Irifm p ps is ls ns \<or> Irifm q ps is ls ns)"
+| "Irifm (RImp p q) ps is ls ns = (Irifm p ps is ls ns \<longrightarrow> Irifm q ps is ls ns)"
+| "Irifm (RIff p q) ps is ls ns = (Irifm p ps is ls ns = Irifm q ps is ls ns)"
+| "Irifm (RNEX p) ps is ls ns =  (\<exists>x. Irifm p ps is ls (x#ns))"
+| "Irifm (RIEX p) ps is ls ns =  (\<exists>x. Irifm p ps (x#is) ls ns)"
+| "Irifm (RLEX p) ps is ls ns =  (\<exists>x. Irifm p ps is (x#ls) ns)"
+| "Irifm (RBEX p) ps is ls ns =  (\<exists>x. Irifm p (x#ps) is ls ns)"
+| "Irifm (RNALL p) ps is ls ns = (\<forall>x. Irifm p ps is ls (x#ns))"
+| "Irifm (RIALL p) ps is ls ns = (\<forall>x. Irifm p ps (x#is) ls ns)"
+| "Irifm (RLALL p) ps is ls ns = (\<forall>x. Irifm p ps is (x#ls) ns)"
+| "Irifm (RBALL p) ps is ls ns = (\<forall>x. Irifm p (x#ps) is ls ns)"
 
 lemma " \<forall>x. \<exists>n. ((Suc n) * length (([(3::int) * x + (f t)*y - 9 + (- z)] @ []) @ xs) = length xs) \<and> m < 5*n - length (xs @ [2,3,4,x*z + 8 - y]) \<longrightarrow> (\<exists>p. \<forall>q. p \<and> q \<longrightarrow> r)"
   apply (reify Irifm.simps Irnat_simps Irlist.simps Irint_simps)
@@ -385,28 +391,28 @@
 (* An example for equations containing type variables *)
 datatype prod = Zero | One | Var nat | Mul prod prod 
   | Pw prod nat | PNM nat nat prod
-consts Iprod :: " prod \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>'a" 
-primrec
+primrec Iprod :: " prod \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>'a" 
+where
   "Iprod Zero vs = 0"
-  "Iprod One vs = 1"
-  "Iprod (Var n) vs = vs!n"
-  "Iprod (Mul a b) vs = (Iprod a vs * Iprod b vs)"
-  "Iprod (Pw a n) vs = ((Iprod a vs) ^ n)"
-  "Iprod (PNM n k t) vs = (vs ! n)^k * Iprod t vs"
+| "Iprod One vs = 1"
+| "Iprod (Var n) vs = vs!n"
+| "Iprod (Mul a b) vs = (Iprod a vs * Iprod b vs)"
+| "Iprod (Pw a n) vs = ((Iprod a vs) ^ n)"
+| "Iprod (PNM n k t) vs = (vs ! n)^k * Iprod t vs"
 consts prodmul:: "prod \<times> prod \<Rightarrow> prod"
 datatype sgn = Pos prod | Neg prod | ZeroEq prod | NZeroEq prod | Tr | F 
   | Or sgn sgn | And sgn sgn
 
-consts Isgn :: " sgn \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>bool"
-primrec 
+primrec Isgn :: " sgn \<Rightarrow> ('a::{linordered_idom}) list \<Rightarrow>bool"
+where 
   "Isgn Tr vs = True"
-  "Isgn F vs = False"
-  "Isgn (ZeroEq t) vs = (Iprod t vs = 0)"
-  "Isgn (NZeroEq t) vs = (Iprod t vs \<noteq> 0)"
-  "Isgn (Pos t) vs = (Iprod t vs > 0)"
-  "Isgn (Neg t) vs = (Iprod t vs < 0)"
-  "Isgn (And p q) vs = (Isgn p vs \<and> Isgn q vs)"
-  "Isgn (Or p q) vs = (Isgn p vs \<or> Isgn q vs)"
+| "Isgn F vs = False"
+| "Isgn (ZeroEq t) vs = (Iprod t vs = 0)"
+| "Isgn (NZeroEq t) vs = (Iprod t vs \<noteq> 0)"
+| "Isgn (Pos t) vs = (Iprod t vs > 0)"
+| "Isgn (Neg t) vs = (Iprod t vs < 0)"
+| "Isgn (And p q) vs = (Isgn p vs \<and> Isgn q vs)"
+| "Isgn (Or p q) vs = (Isgn p vs \<or> Isgn q vs)"
 
 lemmas eqs = Isgn.simps Iprod.simps
 
--- a/src/HOL/ex/Refute_Examples.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/Refute_Examples.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -249,12 +249,13 @@
 
 text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *}
 
-constdefs
-  "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   "trans P == (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)"
-  "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+
+definition "subset" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   "subset P Q == (ALL x y. P x y \<longrightarrow> Q x y)"
-  "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+
+definition "trans_closure" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
   "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R \<longrightarrow> trans R \<longrightarrow> subset P R)"
 
 lemma "trans_closure T P \<longrightarrow> (\<exists>x y. T x y)"
--- a/src/HOL/ex/Sudoku.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/Sudoku.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -23,8 +23,7 @@
 
 datatype digit = A ("1") | B ("2") | C ("3") | D ("4") | E ("5") | F ("6") | G ("7") | H ("8") | I ("9")
 
-constdefs
-  valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool"
+definition valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where
 
   "valid x1 x2 x3 x4 x5 x6 x7 x8 x9 ==
     (x1 \<noteq> x2) \<and> (x1 \<noteq> x3) \<and> (x1 \<noteq> x4) \<and> (x1 \<noteq> x5) \<and> (x1 \<noteq> x6) \<and> (x1 \<noteq> x7) \<and> (x1 \<noteq> x8) \<and> (x1 \<noteq> x9)
@@ -36,8 +35,7 @@
     \<and> (x7 \<noteq> x8) \<and> (x7 \<noteq> x9)
     \<and> (x8 \<noteq> x9)"
 
-constdefs
-  sudoku :: "digit => digit => digit => digit => digit => digit => digit => digit => digit =>
+definition sudoku :: "digit => digit => digit => digit => digit => digit => digit => digit => digit =>
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
@@ -45,7 +43,7 @@
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
     digit => digit => digit => digit => digit => digit => digit => digit => digit =>
-    digit => digit => digit => digit => digit => digit => digit => digit => digit => bool"
+    digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where
 
   "sudoku x11 x12 x13 x14 x15 x16 x17 x18 x19
           x21 x22 x23 x24 x25 x26 x27 x28 x29
--- a/src/HOL/ex/ThreeDivides.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOL/ex/ThreeDivides.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -149,10 +149,10 @@
 The function @{text "nlen"} returns the number of digits in a natural
 number n. *}
 
-consts nlen :: "nat \<Rightarrow> nat"
-recdef nlen "measure id"
+fun nlen :: "nat \<Rightarrow> nat"
+where
   "nlen 0 = 0"
-  "nlen x = 1 + nlen (x div 10)"
+| "nlen x = 1 + nlen (x div 10)"
 
 text {* The function @{text "sumdig"} returns the sum of all digits in
 some number n. *}
--- a/src/HOLCF/IOA/meta_theory/Pred.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOLCF/IOA/meta_theory/Pred.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -24,25 +24,25 @@
 IMPLIES      ::"'a predicate => 'a predicate => 'a predicate"    (infixr ".-->" 25)
 
 
-syntax ("" output)
-  "NOT"     ::"'a predicate => 'a predicate" ("~ _" [40] 40)
-  "AND"     ::"'a predicate => 'a predicate => 'a predicate"    (infixr "&" 35)
-  "OR"      ::"'a predicate => 'a predicate => 'a predicate"    (infixr "|" 30)
-  "IMPLIES" ::"'a predicate => 'a predicate => 'a predicate"    (infixr "-->" 25)
+notation (output)
+  NOT  ("~ _" [40] 40) and
+  AND  (infixr "&" 35) and
+  OR  (infixr "|" 30) and
+  IMPLIES  (infixr "-->" 25)
 
-syntax (xsymbols output)
-  "NOT"    ::"'a predicate => 'a predicate" ("\<not> _" [40] 40)
-  "AND"    ::"'a predicate => 'a predicate => 'a predicate"    (infixr "\<and>" 35)
-  "OR"     ::"'a predicate => 'a predicate => 'a predicate"    (infixr "\<or>" 30)
-  "IMPLIES" ::"'a predicate => 'a predicate => 'a predicate"    (infixr "\<longrightarrow>" 25)
+notation (xsymbols output)
+  NOT  ("\<not> _" [40] 40) and
+  AND  (infixr "\<and>" 35) and
+  OR  (infixr "\<or>" 30) and
+  IMPLIES  (infixr "\<longrightarrow>" 25)
 
-syntax (xsymbols)
-  "satisfies"  ::"'a => 'a predicate => bool"    ("_ \<Turnstile> _" [100,9] 8)
+notation (xsymbols)
+  satisfies  ("_ \<Turnstile> _" [100,9] 8)
 
-syntax (HTML output)
-  "NOT"    ::"'a predicate => 'a predicate" ("\<not> _" [40] 40)
-  "AND"    ::"'a predicate => 'a predicate => 'a predicate"    (infixr "\<and>" 35)
-  "OR"     ::"'a predicate => 'a predicate => 'a predicate"    (infixr "\<or>" 30)
+notation (HTML output)
+  NOT  ("\<not> _" [40] 40) and
+  AND  (infixr "\<and>" 35) and
+  OR  (infixr "\<or>" 30)
 
 
 defs
--- a/src/HOLCF/Tools/Domain/domain_extender.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_extender.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -285,7 +285,7 @@
   || Scan.succeed [];
 
 val domain_decl =
-  (type_args' -- P.binding -- P.opt_infix) --
+  (type_args' -- P.binding -- P.opt_mixfix) --
     (P.$$$ "=" |-- P.enum1 "|" cons_decl);
 
 val domains_decl =
--- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -907,7 +907,7 @@
 val parse_domain_iso :
     (string list * binding * mixfix * string * (binding * binding) option)
       parser =
-  (P.type_args -- P.binding -- P.opt_infix -- (P.$$$ "=" |-- P.typ) --
+  (P.type_args -- P.binding -- P.opt_mixfix -- (P.$$$ "=" |-- P.typ) --
     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding)))
     >> (fn ((((vs, t), mx), rhs), morphs) => (vs, t, mx, rhs, morphs));
 
--- a/src/HOLCF/Tools/pcpodef.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOLCF/Tools/pcpodef.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -340,7 +340,7 @@
   Scan.optional (P.$$$ "(" |--
       ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
         --| P.$$$ ")") (true, NONE) --
-    (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
+    (P.type_args -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
 
 fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
--- a/src/HOLCF/Tools/repdef.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/HOLCF/Tools/repdef.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -167,7 +167,7 @@
   Scan.optional (P.$$$ "(" |--
       ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
         --| P.$$$ ")") (true, NONE) --
-    (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
+    (P.type_args -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
 
 fun mk_repdef ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
--- a/src/Provers/Arith/cancel_factor.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Provers/Arith/cancel_factor.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -35,7 +35,7 @@
       if t aconv u then (u, k + 1) :: uks
       else (t, 1) :: (u, k) :: uks;
 
-fun count_terms ts = foldr ins_term (sort TermOrd.term_ord ts, []);
+fun count_terms ts = foldr ins_term (sort Term_Ord.term_ord ts, []);
 
 
 (* divide sum *)
--- a/src/Provers/Arith/cancel_sums.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Provers/Arith/cancel_sums.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -37,11 +37,11 @@
 fun cons2 y (x, ys, z) = (x, y :: ys, z);
 fun cons12 x y (xs, ys, z) = (x :: xs, y :: ys, z);
 
-(*symmetric difference of multisets -- assumed to be sorted wrt. TermOrd.term_ord*)
+(*symmetric difference of multisets -- assumed to be sorted wrt. Term_Ord.term_ord*)
 fun cancel ts [] vs = (ts, [], vs)
   | cancel [] us vs = ([], us, vs)
   | cancel (t :: ts) (u :: us) vs =
-      (case TermOrd.term_ord (t, u) of
+      (case Term_Ord.term_ord (t, u) of
         EQUAL => cancel ts us (t :: vs)
       | LESS => cons1 t (cancel ts (u :: us) vs)
       | GREATER => cons2 u (cancel (t :: ts) us vs));
@@ -63,7 +63,7 @@
   | SOME bal =>
       let
         val thy = ProofContext.theory_of (Simplifier.the_context ss);
-        val (ts, us) = pairself (sort TermOrd.term_ord o Data.dest_sum) bal;
+        val (ts, us) = pairself (sort Term_Ord.term_ord o Data.dest_sum) bal;
         val (ts', us', vs) = cancel ts us [];
       in
         if null vs then NONE
--- a/src/Pure/Concurrent/par_list.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Concurrent/par_list.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -27,7 +27,7 @@
 struct
 
 fun raw_map f xs =
-  if Multithreading.enabled () then
+  if Multithreading.enabled () andalso not (Multithreading.self_critical ()) then
     let val group = Task_Queue.new_group (Future.worker_group ())
     in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end
   else map (Exn.capture f) xs;
--- a/src/Pure/General/graph.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/General/graph.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -48,7 +48,6 @@
   val topological_order: 'a T -> key list
   val add_edge_trans_acyclic: key * key -> 'a T -> 'a T               (*exception CYCLES*)
   val merge_trans_acyclic: ('a * 'a -> bool) -> 'a T * 'a T -> 'a T   (*exception CYCLES*)
-  val extend: (key -> 'a * key list) -> key -> 'a T -> 'a T
 end;
 
 functor Graph(Key: KEY): GRAPH =
@@ -279,26 +278,10 @@
   |> fold add_edge_trans_acyclic (diff_edges G2 G1);
 
 
-(* constructing graphs *)
-
-fun extend explore =
-  let
-    fun ext x G =
-      if can (get_entry G) x then G
-      else
-        let val (info, ys) = explore x in
-          G
-          |> new_node (x, info)
-          |> fold ext ys
-          |> fold (fn y => add_edge (x, y)) ys
-        end
-  in ext end;
-
-
 (*final declarations of this structure!*)
 val fold = fold_graph;
 
 end;
 
 structure Graph = Graph(type key = string val ord = fast_string_ord);
-structure IntGraph = Graph(type key = int val ord = int_ord);
+structure Int_Graph = Graph(type key = int val ord = int_ord);
--- a/src/Pure/Isar/args.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/args.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -54,9 +54,9 @@
   val term: term context_parser
   val term_abbrev: term context_parser
   val prop: term context_parser
-  val tyname: string context_parser
-  val const: string context_parser
-  val const_proper: string context_parser
+  val type_name: bool -> string context_parser
+  val const: bool -> string context_parser
+  val const_proper: bool -> string context_parser
   val bang_facts: thm list context_parser
   val goal_spec: ((int -> tactic) -> tactic) context_parser
   val parse: token list parser
@@ -209,13 +209,16 @@
 
 (* type and constant names *)
 
-val tyname = Scan.peek (named_typ o ProofContext.read_tyname o Context.proof_of)
+fun type_name strict =
+  Scan.peek (fn ctxt => named_typ (ProofContext.read_type_name (Context.proof_of ctxt) strict))
   >> (fn Type (c, _) => c | TFree (a, _) => a | _ => "");
 
-val const = Scan.peek (named_term o ProofContext.read_const o Context.proof_of)
+fun const strict =
+  Scan.peek (fn ctxt => named_term (ProofContext.read_const (Context.proof_of ctxt) strict))
   >> (fn Const (c, _) => c | Free (x, _) => x | _ => "");
 
-val const_proper = Scan.peek (named_term o ProofContext.read_const_proper o Context.proof_of)
+fun const_proper strict =
+  Scan.peek (fn ctxt => named_term (ProofContext.read_const_proper (Context.proof_of ctxt) strict))
   >> (fn Const (c, _) => c | _ => "");
 
 
--- a/src/Pure/Isar/code.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/code.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -420,10 +420,11 @@
      of Type ("fun", [ty, ty_abs]) => (ty, ty_abs)
       | _ => error ("Not a datatype abstractor:\n" ^ string_of_const thy abs
           ^ " :: " ^ string_of_typ thy ty');
-    val _ = Thm.cterm_of thy (Const (rep, ty_abs --> ty)) handle CTERM _ =>
+    val _ = Thm.cterm_of thy (Const (rep, ty_abs --> ty)) handle TYPE _ =>
       error ("Not a projection:\n" ^ string_of_const thy rep);
-    val cert = Logic.mk_equals (Const (abs, ty --> ty_abs) $ (Const (rep, ty_abs --> ty)
-      $ Free ("x", ty_abs)), Free ("x", ty_abs));
+    val param = Free ("x", ty_abs);
+    val cert = Logic.all param (Logic.mk_equals
+      (Const (abs, ty --> ty_abs) $ (Const (rep, ty_abs --> ty) $ param), param));
   in (tyco, (vs ~~ sorts, ((fst abs_ty, ty), (rep, cert)))) end;    
 
 fun get_type_entry thy tyco = case these (Symtab.lookup ((the_types o the_exec) thy) tyco)
@@ -839,7 +840,9 @@
 fun bare_thms_of_cert thy (cert as Equations _) =
       (map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
         o snd o equations_of_cert thy) cert
-  | bare_thms_of_cert thy _ = [];
+  | bare_thms_of_cert thy (Projection _) = []
+  | bare_thms_of_cert thy (Abstract (abs_thm, tyco)) =
+      [Thm.varifyT (snd (concretify_abs thy tyco abs_thm))];
 
 end;
 
--- a/src/Pure/Isar/isar_syn.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/isar_syn.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -104,13 +104,13 @@
 
 val _ =
   OuterSyntax.command "typedecl" "type declaration" K.thy_decl
-    (P.type_args -- P.binding -- P.opt_infix >> (fn ((args, a), mx) =>
+    (P.type_args -- P.binding -- P.opt_mixfix >> (fn ((args, a), mx) =>
       Toplevel.theory (ObjectLogic.typedecl (a, args, mx) #> snd)));
 
 val _ =
   OuterSyntax.command "types" "declare type abbreviations" K.thy_decl
     (Scan.repeat1
-      (P.type_args -- P.binding -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_infix')))
+      (P.type_args -- P.binding -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_mixfix')))
       >> (Toplevel.theory o Sign.add_tyabbrs o
         map (fn ((args, a), (T, mx)) => (a, args, T, mx))));
 
@@ -226,6 +226,16 @@
     >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
 
 val _ =
+  OuterSyntax.local_theory "type_notation" "add notation 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
+    (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
     (opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)
     >> (fn (mode, args) => Specification.notation_cmd true mode args));
--- a/src/Pure/Isar/local_syntax.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/local_syntax.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -13,11 +13,12 @@
   val structs_of: T -> string list
   val init: theory -> T
   val rebuild: theory -> T -> T
-  val add_syntax: theory -> (bool * (string * typ * mixfix)) list -> T -> T
+  datatype kind = Type | Const | Fixed
+  val add_syntax: theory -> (kind * (string * typ * mixfix)) list -> T -> T
   val set_mode: Syntax.mode -> T -> T
   val restore_mode: T -> T -> T
   val update_modesyntax: theory -> bool -> Syntax.mode ->
-    (bool * (string * typ * mixfix)) list -> T -> T
+    (kind * (string * typ * mixfix)) list -> T -> T
   val extern_term: T -> term -> term
 end;
 
@@ -27,8 +28,8 @@
 (* datatype T *)
 
 type local_mixfix =
-  (string * bool) *                                    (*name, fixed?*)
-  ((bool * Syntax.mode) * (string * typ * mixfix));    (*add?, mode, declaration*)
+  (string * bool) *  (*name, fixed?*)
+  ((bool * bool * Syntax.mode) * (string * typ * mixfix));  (*type?, add?, mode, declaration*)
 
 datatype T = Syntax of
  {thy_syntax: Syntax.syntax,
@@ -57,15 +58,16 @@
 fun build_syntax thy mode mixfixes (idents as (structs, _)) =
   let
     val thy_syntax = Sign.syn_of thy;
-    val is_logtype = Sign.is_logtype thy;
-    fun const_gram ((true, m), decls) = Syntax.update_const_gram is_logtype m decls
-      | const_gram ((false, m), decls) = Syntax.remove_const_gram is_logtype m decls;
+    fun update_gram ((true, add, m), decls) = Syntax.update_type_gram add m decls
+      | update_gram ((false, add, m), decls) =
+          Syntax.update_const_gram add (Sign.is_logtype thy) m decls;
+
     val (atrs, trs, trs', atrs') = Syntax.struct_trfuns structs;
     val local_syntax = thy_syntax
       |> Syntax.update_trfuns
           (map Syntax.mk_trfun atrs, map Syntax.mk_trfun trs,
            map Syntax.mk_trfun trs', map Syntax.mk_trfun atrs')
-      |> fold const_gram (AList.coalesce (op =) (rev (map snd mixfixes)));
+      |> fold update_gram (AList.coalesce (op =) (rev (map snd mixfixes)));
   in make_syntax (thy_syntax, local_syntax, mode, mixfixes, idents) end;
 
 fun init thy = build_syntax thy Syntax.mode_default [] ([], []);
@@ -77,16 +79,18 @@
 
 (* mixfix declarations *)
 
+datatype kind = Type | Const | Fixed;
+
 local
 
-fun prep_mixfix _ (_, (_, _, Structure)) = NONE
-  | prep_mixfix mode (fixed, (x, T, mx)) =
-      let val c = if fixed then Syntax.fixedN ^ x else x
-      in SOME ((x, fixed), (mode, (c, T, mx))) end;
+fun prep_mixfix _ _ (_, (_, _, Structure)) = NONE
+  | prep_mixfix add mode (Type, decl as (x, _, _)) = SOME ((x, false), ((true, add, mode), decl))
+  | prep_mixfix add mode (Const, decl as (x, _, _)) = SOME ((x, false), ((false, add, mode), decl))
+  | prep_mixfix add mode (Fixed, (x, T, mx)) =
+      SOME ((x, true), ((false, add, mode), (Syntax.mark_fixed x, T, mx)));
 
-fun prep_struct (fixed, (c, _, Structure)) =
-      if fixed then SOME c
-      else error ("Bad mixfix declaration for constant: " ^ quote c)
+fun prep_struct (Fixed, (c, _, Structure)) = SOME c
+  | prep_struct (_, (c, _, Structure)) = error ("Bad mixfix declaration for " ^ quote c)
   | prep_struct _ = NONE;
 
 in
@@ -97,7 +101,7 @@
     [] => syntax
   | decls =>
       let
-        val new_mixfixes = map_filter (prep_mixfix (add, mode)) decls;
+        val new_mixfixes = map_filter (prep_mixfix add mode) decls;
         val new_structs = map_filter prep_struct decls;
         val mixfixes' = rev new_mixfixes @ mixfixes;
         val structs' =
@@ -130,7 +134,7 @@
     fun map_free (t as Free (x, T)) =
           let val i = find_index (fn s => s = x) structs + 1 in
             if i = 0 andalso member (op =) fixes x then
-              Const (Syntax.fixedN ^ x, T)
+              Term.Const (Syntax.mark_fixed x, T)
             else if i = 1 andalso not (! show_structs) then
               Syntax.const "_struct" $ Syntax.const "_indexdefault"
             else t
--- a/src/Pure/Isar/local_theory.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/local_theory.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -42,6 +42,7 @@
   val type_syntax: bool -> declaration -> local_theory -> local_theory
   val term_syntax: bool -> declaration -> local_theory -> local_theory
   val declaration: bool -> declaration -> local_theory -> local_theory
+  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
   val init: serial option -> string -> operations -> Proof.context -> local_theory
   val restore: local_theory -> local_theory
@@ -198,6 +199,10 @@
 val notes = notes_kind "";
 fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single;
 
+fun type_notation add mode raw_args lthy =
+  let val args = map (apfst (Morphism.typ (target_morphism lthy))) raw_args
+  in type_syntax false (ProofContext.target_type_notation add mode args) lthy end;
+
 fun notation add mode raw_args lthy =
   let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args
   in term_syntax false (ProofContext.target_notation add mode args) lthy end;
--- a/src/Pure/Isar/outer_parse.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/outer_parse.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -72,8 +72,6 @@
   val typ: string parser
   val mixfix: mixfix parser
   val mixfix': mixfix parser
-  val opt_infix: mixfix parser
-  val opt_infix': mixfix parser
   val opt_mixfix: mixfix parser
   val opt_mixfix': mixfix parser
   val where_: string parser
@@ -279,8 +277,6 @@
 
 val mixfix = annotation !!! (mfix || binder || infxl || infxr || infx);
 val mixfix' = annotation I (mfix || binder || infxl || infxr || infx);
-val opt_infix = opt_annotation !!! (infxl || infxr || infx);
-val opt_infix' = opt_annotation I (infxl || infxr || infx);
 val opt_mixfix = opt_annotation !!! (mfix || binder || infxl || infxr || infx);
 val opt_mixfix' = opt_annotation I (mfix || binder || infxl || infxr || infx);
 
--- a/src/Pure/Isar/proof_context.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/proof_context.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -52,9 +52,9 @@
   val infer_type: Proof.context -> string -> typ
   val inferred_param: string -> Proof.context -> typ * Proof.context
   val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
-  val read_tyname: Proof.context -> string -> typ
-  val read_const_proper: Proof.context -> string -> term
-  val read_const: Proof.context -> string -> term
+  val read_type_name: Proof.context -> bool -> string -> typ
+  val read_const_proper: Proof.context -> bool -> string -> term
+  val read_const: Proof.context -> bool -> string -> term
   val allow_dummies: Proof.context -> Proof.context
   val decode_term: Proof.context -> term -> term
   val standard_infer_types: Proof.context -> term list -> term list
@@ -115,7 +115,10 @@
   val add_cases: bool -> (string * Rule_Cases.T option) list -> Proof.context -> Proof.context
   val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context
   val get_case: Proof.context -> string -> string option list -> Rule_Cases.T
+  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context
+  val target_type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> morphism ->
+    Context.generic -> Context.generic
   val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->
     Context.generic -> Context.generic
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
@@ -360,7 +363,7 @@
           (Pretty.str (setmp_CRITICAL show_question_marks true Term.string_of_vname (x', i))))
   | NONE => Pretty.mark Markup.var (Pretty.str s));
 
-fun class_markup _ c =    (* FIXME authentic name *)
+fun class_markup _ c =    (* FIXME authentic syntax *)
   Pretty.mark (Markup.tclassN, []) (Pretty.str c);
 
 fun plain_markup m _ s = Pretty.mark m (Pretty.str s);
@@ -405,19 +408,27 @@
 
 val token_content = Syntax.read_token #>> Symbol_Pos.content;
 
-fun prep_const_proper ctxt (c, pos) =
-  let val t as (Const (d, _)) =
-    (case Variable.lookup_const ctxt c of
-      SOME d => Const (d, Consts.type_scheme (consts_of ctxt) d handle TYPE (msg, _, _) => error msg)
-    | NONE => Consts.read_const (consts_of ctxt) c)
-  in Position.report (Markup.const d) pos; t end;
+fun prep_const_proper ctxt strict (c, pos) =
+  let
+    fun err msg = error (msg ^ Position.str_of pos);
+    val consts = consts_of ctxt;
+    val t as Const (d, _) =
+      (case Variable.lookup_const ctxt c of
+        SOME d =>
+          Const (d, Consts.type_scheme (consts_of ctxt) d handle TYPE (msg, _, _) => err msg)
+      | NONE => Consts.read_const consts c);
+    val _ =
+      if strict then ignore (Consts.the_type consts d) handle TYPE (msg, _, _) => err msg
+      else ();
+    val _ = Position.report (Markup.const d) pos;
+  in t end;
 
 in
 
-fun read_tyname ctxt str =
+fun read_type_name ctxt strict text =
   let
     val thy = theory_of ctxt;
-    val (c, pos) = token_content str;
+    val (c, pos) = token_content text;
   in
     if Syntax.is_tid c then
      (Position.report Markup.tfree pos;
@@ -425,20 +436,27 @@
     else
       let
         val d = Sign.intern_type thy c;
+        val decl = Sign.the_type_decl thy d;
         val _ = Position.report (Markup.tycon d) pos;
-      in Type (d, replicate (Sign.arity_number thy d) dummyT) end
+        fun err () = error ("Bad type name: " ^ quote d);
+        val args =
+          (case decl of
+            Type.LogicalType n => n
+          | Type.Abbreviation (vs, _, _) => if strict then err () else length vs
+          | Type.Nonterminal => if strict then err () else 0);
+      in Type (d, replicate args dummyT) end
   end;
 
-fun read_const_proper ctxt = prep_const_proper ctxt o token_content;
+fun read_const_proper ctxt strict = prep_const_proper ctxt strict o token_content;
 
-fun read_const ctxt str =
-  let val (c, pos) = token_content str in
+fun read_const ctxt strict 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))
-    | _ => prep_const_proper ctxt (c, pos))
+    | _ => prep_const_proper ctxt strict (c, pos))
   end;
 
 end;
@@ -567,7 +585,7 @@
   let
     val _ = no_skolem false x;
     val sko = lookup_skolem ctxt x;
-    val is_const = can (read_const_proper ctxt) x orelse Long_Name.is_qualified x;
+    val is_const = can (read_const_proper ctxt false) x orelse Long_Name.is_qualified x;
     val is_declared = is_some (def_type (x, ~1));
   in
     if Variable.is_const ctxt x then NONE
@@ -581,7 +599,7 @@
 fun term_context ctxt =
   let val thy = theory_of ctxt in
    {get_sort = get_sort thy (Variable.def_sort ctxt),
-    map_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt a)))
+    map_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt false a)))
       handle ERROR _ => (false, Consts.intern (consts_of ctxt) a)),
     map_free = intern_skolem ctxt (Variable.def_type ctxt false),
     map_type = Sign.intern_tycons thy,
@@ -987,7 +1005,7 @@
 
 fun const_ast_tr intern ctxt [Syntax.Variable c] =
       let
-        val Const (c', _) = read_const_proper ctxt c;
+        val Const (c', _) = read_const_proper ctxt false c;
         val d = if intern then Syntax.mark_const c' else c;
       in Syntax.Constant d end
   | const_ast_tr _ _ asts = raise Syntax.AST ("const_ast_tr", asts);
@@ -1014,18 +1032,34 @@
 
 local
 
-fun const_syntax _ (Free (x, T), mx) = SOME (true, (x, T, mx))
+fun type_syntax (Type (c, args), mx) =  (* FIXME authentic syntax *)
+      SOME (Local_Syntax.Type, (Long_Name.base_name c, Syntax.make_type (length args), mx))
+  | type_syntax _ = NONE;
+
+fun const_syntax _ (Free (x, T), mx) = SOME (Local_Syntax.Fixed, (x, T, mx))
   | const_syntax ctxt (Const (c, _), mx) =
       (case try (Consts.type_scheme (consts_of ctxt)) c of
-        SOME T => SOME (false, (Syntax.mark_const c, T, mx))
+        SOME T => SOME (Local_Syntax.Const, (Syntax.mark_const c, T, mx))
       | NONE => NONE)
   | const_syntax _ _ = NONE;
 
+fun gen_notation syntax add mode args ctxt =
+  ctxt |> map_syntax
+    (Local_Syntax.update_modesyntax (theory_of ctxt) add mode (map_filter (syntax ctxt) args));
+
 in
 
-fun notation add mode args ctxt =
-  ctxt |> map_syntax
-    (Local_Syntax.update_modesyntax (theory_of ctxt) add mode (map_filter (const_syntax ctxt) args));
+val type_notation = gen_notation (K type_syntax);
+val notation = gen_notation const_syntax;
+
+fun target_type_notation add  mode args phi =
+  let
+    val args' = args |> map_filter (fn (T, mx) =>
+      let
+        val T' = Morphism.typ phi T;
+        val similar = (case (T, T') of (Type (c, _), Type (c', _)) => c = c' | _ => false);
+      in if similar then SOME (T', mx) else NONE end);
+  in Context.mapping (Sign.type_notation add mode args') (type_notation add mode args') end;
 
 fun target_notation add mode args phi =
   let
@@ -1034,6 +1068,7 @@
       in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end);
   in Context.mapping (Sign.notation add mode args') (notation add mode args') end;
 
+
 end;
 
 
@@ -1071,7 +1106,7 @@
   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 (true, (x, T, mx));
+  else (Local_Syntax.Fixed, (x, T, mx));
 
 in
 
--- a/src/Pure/Isar/rule_cases.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/rule_cases.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -379,7 +379,7 @@
 local
 
 fun equal_cterms ts us =
-  is_equal (list_ord (TermOrd.fast_term_ord o pairself Thm.term_of) (ts, us));
+  is_equal (list_ord (Term_Ord.fast_term_ord o pairself Thm.term_of) (ts, us));
 
 fun prep_rule n th =
   let
--- a/src/Pure/Isar/specification.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Isar/specification.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -42,6 +42,8 @@
     local_theory -> local_theory
   val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option * string ->
     local_theory -> local_theory
+  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory
+  val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
   val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
   val theorems: string ->
@@ -255,11 +257,23 @@
 
 (* notation *)
 
+local
+
+fun gen_type_notation prep_type add mode args lthy =
+  lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args);
+
 fun gen_notation prep_const add mode args lthy =
   lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args);
 
+in
+
+val type_notation = gen_type_notation (K I);
+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 ProofContext.read_const;
+val notation_cmd = gen_notation (fn ctxt => ProofContext.read_const ctxt false);
+
+end;
 
 
 (* fact statements *)
--- a/src/Pure/ML/ml_antiquote.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/ML/ml_antiquote.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -73,9 +73,6 @@
 
 val _ = value "context" (Scan.succeed "ML_Context.the_local_context ()");
 
-val _ = inline "sort" (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
-  ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))));
-
 val _ = inline "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ));
 val _ = inline "term" (Args.term >> (ML_Syntax.atomic o ML_Syntax.print_term));
 val _ = inline "prop" (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term));
@@ -103,36 +100,67 @@
     "Thm.cterm_of (ML_Context.the_global_context ()) " ^ ML_Syntax.atomic (ML_Syntax.print_term t)));
 
 
-fun type_ syn = (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, c) =>
-    #1 (Term.dest_Type (ProofContext.read_tyname ctxt c))
-    |> syn ? Long_Name.base_name
-    |> ML_Syntax.print_string));
+(* type classes *)
 
-val _ = inline "type_name" (type_ false);
-val _ = inline "type_syntax" (type_ true);
+fun class syn = Args.theory -- Scan.lift Args.name_source >> (fn (thy, s) =>
+  Sign.read_class thy s
+  |> syn ? Long_Name.base_name   (* FIXME authentic syntax *)
+  |> ML_Syntax.print_string);
+
+val _ = inline "class" (class false);
+val _ = inline "class_syntax" (class true);
+
+val _ = inline "sort" (Args.context -- Scan.lift Args.name_source >> (fn (ctxt, s) =>
+  ML_Syntax.atomic (ML_Syntax.print_sort (Syntax.read_sort ctxt s))));
 
 
-fun const syn = Args.context -- Scan.lift Args.name_source >> (fn (ctxt, c) =>
-  #1 (Term.dest_Const (ProofContext.read_const_proper ctxt c))
-  |> syn ? Syntax.mark_const
-  |> ML_Syntax.print_string);
+(* type constructors *)
+
+fun type_name kind check = Args.context -- Scan.lift (OuterParse.position Args.name_source)
+  >> (fn (ctxt, (s, pos)) =>
+    let
+      val Type (c, _) = ProofContext.read_type_name ctxt false s;
+      val decl = Sign.the_type_decl (ProofContext.theory_of ctxt) c;
+      val res =
+        (case try check (c, decl) of
+          SOME res => res
+        | NONE => error ("Not a " ^ kind ^ ": " ^ quote c ^ Position.str_of pos));
+    in ML_Syntax.print_string res end);
+
+val _ = inline "type_name" (type_name "logical type" (fn (c, Type.LogicalType _) => c));
+val _ = inline "type_abbrev" (type_name "type abbreviation" (fn (c, Type.Abbreviation _) => c));
+val _ = inline "nonterminal" (type_name "nonterminal" (fn (c, Type.Nonterminal) => c));
+val _ = inline "type_syntax" (type_name "type" (fn (c, _) => Long_Name.base_name c));  (* FIXME authentic syntax *)
+
 
-val _ = inline "const_name" (const false);
-val _ = inline "const_syntax" (const true);
+(* constants *)
+
+fun const_name check = Args.context -- Scan.lift (OuterParse.position Args.name_source)
+  >> (fn (ctxt, (s, pos)) =>
+    let
+      val Const (c, _) = ProofContext.read_const_proper ctxt false s;
+      val res = check (ProofContext.consts_of ctxt, c)
+        handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos);
+    in ML_Syntax.print_string res end);
+
+val _ = inline "const_name" (const_name (fn (consts, c) => (Consts.the_type consts c; c)));
+val _ = inline "const_abbrev" (const_name (fn (consts, c) => (Consts.the_abbreviation consts c; c)));
+val _ = inline "const_syntax" (const_name (fn (_, c) => Syntax.mark_const c));
+
+
+val _ = inline "syntax_const"
+  (Args.context -- Scan.lift (OuterParse.position Args.name) >> (fn (ctxt, (c, pos)) =>
+    if Syntax.is_const (ProofContext.syn_of ctxt) c then ML_Syntax.print_string c
+    else error ("Unknown syntax const: " ^ quote c ^ Position.str_of pos)));
 
 val _ = inline "const"
   (Args.context -- Scan.lift Args.name_source -- Scan.optional
       (Scan.lift (Args.$$$ "(") |-- OuterParse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
     >> (fn ((ctxt, raw_c), Ts) =>
       let
-        val Const (c, _) = ProofContext.read_const_proper ctxt raw_c;
+        val Const (c, _) = ProofContext.read_const_proper ctxt true raw_c;
         val const = Const (c, Consts.instance (ProofContext.consts_of ctxt) (c, Ts));
       in ML_Syntax.atomic (ML_Syntax.print_term const) end));
 
-val _ = inline "syntax_const"
-  (Args.context -- Scan.lift Args.name >> (fn (ctxt, c) =>
-    if Syntax.is_const (ProofContext.syn_of ctxt) c then ML_Syntax.print_string c
-    else error ("Unknown syntax const: " ^ quote c)));
-
 end;
 
--- a/src/Pure/Syntax/mixfix.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Syntax/mixfix.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -21,17 +21,17 @@
 signature MIXFIX1 =
 sig
   include MIXFIX0
-  val literal: string -> mixfix
   val no_syn: 'a * 'b -> 'a * 'b * mixfix
   val pretty_mixfix: mixfix -> Pretty.T
   val mixfix_args: mixfix -> int
   val mixfixT: mixfix -> typ
+  val make_type: int -> typ
 end;
 
 signature MIXFIX =
 sig
   include MIXFIX1
-  val syn_ext_types: (string * int * mixfix) list -> SynExt.syn_ext
+  val syn_ext_types: (string * typ * mixfix) list -> SynExt.syn_ext
   val syn_ext_consts: (string -> bool) -> (string * typ * mixfix) list -> SynExt.syn_ext
 end;
 
@@ -51,8 +51,6 @@
   Binder of string * int * int |
   Structure;
 
-val literal = Delimfix o SynExt.escape_mfix;
-
 fun no_syn (x, y) = (x, y, NoSyn);
 
 
@@ -99,22 +97,29 @@
 
 (* syn_ext_types *)
 
+fun make_type n = replicate n SynExt.typeT ---> SynExt.typeT;
+
 fun syn_ext_types type_decls =
   let
-    fun mk_infix sy t p1 p2 p3 =
-      SynExt.Mfix ("(_ " ^ sy ^ "/ _)",
-        [SynExt.typeT, SynExt.typeT] ---> SynExt.typeT, t, [p1, p2], p3);
+    fun mk_infix sy ty t p1 p2 p3 = SynExt.Mfix ("(_ " ^ sy ^ "/ _)", ty, t, [p1, p2], p3);
 
     fun mfix_of (_, _, NoSyn) = NONE
-      | mfix_of (t, 2, Infix (sy, p)) = SOME (mk_infix sy t (p + 1) (p + 1) p)
-      | mfix_of (t, 2, Infixl (sy, p)) = SOME (mk_infix sy t p (p + 1) p)
-      | mfix_of (t, 2, Infixr (sy, p)) = SOME (mk_infix sy t (p + 1) p p)
-      | mfix_of (t, _, _) =
-          error ("Bad mixfix declaration for type: " ^ quote t);  (* FIXME printable!? *)
+      | mfix_of (t, ty, Mixfix (sy, ps, p)) = SOME (SynExt.Mfix (sy, ty, t, ps, p))
+      | mfix_of (t, ty, Delimfix sy) = SOME (SynExt.Mfix (sy, ty, t, [], SynExt.max_pri))
+      | mfix_of (t, ty, Infix (sy, p)) = SOME (mk_infix sy ty t (p + 1) (p + 1) p)
+      | mfix_of (t, ty, Infixl (sy, p)) = SOME (mk_infix sy ty t p (p + 1) p)
+      | mfix_of (t, ty, Infixr (sy, p)) = SOME (mk_infix sy ty t (p + 1) p p)
+      | mfix_of (t, _, _) = error ("Bad mixfix declaration for " ^ quote t);  (* FIXME printable t (!?) *)
 
-    val mfix = map_filter mfix_of type_decls;
+    fun check_args (_, ty, _) (SOME (mfix as SynExt.Mfix (sy, _, _, _, _))) =
+          if length (Term.binder_types ty) = SynExt.mfix_args sy then ()
+          else SynExt.err_in_mfix "Bad number of type constructor arguments" mfix
+      | check_args _ NONE = ();
+
+    val mfix = map mfix_of type_decls;
+    val _ = map2 check_args type_decls mfix;
     val xconsts = map #1 type_decls;
-  in SynExt.syn_ext mfix xconsts ([], [], [], []) [] ([], []) end;
+  in SynExt.syn_ext (map_filter I mfix) xconsts ([], [], [], []) [] ([], []) end;
 
 
 (* syn_ext_consts *)
@@ -140,7 +145,7 @@
       | mfix_of (c, ty, Infixr (sy, p)) = mk_infix sy ty c (p + 1) p p
       | mfix_of (c, ty, Binder (sy, p, q)) =
           [SynExt.Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, (binder_name c), [0, p], q)]
-      | mfix_of (c, _, _) = error ("Bad mixfix declaration for const: " ^ quote c);
+      | mfix_of (c, _, _) = error ("Bad mixfix declaration for " ^ quote c);   (* FIXME printable c (!?) *)
 
     fun binder (c, _, Binder _) = SOME (binder_name c, c)
       | binder _ = NONE;
--- a/src/Pure/Syntax/syn_ext.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Syntax/syn_ext.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -14,6 +14,7 @@
   val stamp_trfun: stamp -> string * 'a -> string * ('a * stamp)
   val mk_trfun: string * 'a -> string * ('a * stamp)
   val eq_trfun: ('a * stamp) * ('a * stamp) -> bool
+  val escape: string -> string
   val tokentrans_mode:
     string -> (string * (Proof.context -> string -> Pretty.T)) list ->
     (string * string * (Proof.context -> string -> Pretty.T)) list
@@ -28,6 +29,8 @@
   val cargs: string
   val any: string
   val sprop: string
+  datatype mfix = Mfix of string * typ * string * int list * int
+  val err_in_mfix: string -> mfix -> 'a
   val typ_to_nonterm: typ -> string
   datatype xsymb =
     Delim of string |
@@ -37,7 +40,6 @@
   datatype xprod = XProd of string * xsymb list * string * int
   val chain_pri: int
   val delims_of: xprod list -> string list list
-  datatype mfix = Mfix of string * typ * string * int list * int
   datatype syn_ext =
     SynExt of {
       xprods: xprod list,
@@ -53,7 +55,6 @@
       token_translation: (string * string * (Proof.context -> string -> Pretty.T)) list}
   val mfix_delims: string -> string list
   val mfix_args: string -> int
-  val escape_mfix: string -> string
   val syn_ext': bool -> (string -> bool) -> mfix list ->
     string list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *
     (string * ((Proof.context -> term list -> term) * stamp)) list *
@@ -228,7 +229,7 @@
 fun mfix_delims sy = fold_rev (fn Delim s => cons s | _ => I) (read_mfix sy) [];
 val mfix_args = length o filter is_argument o read_mfix;
 
-val escape_mfix = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode;
+val escape = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode;
 
 end;
 
--- a/src/Pure/Syntax/syntax.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Syntax/syntax.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -59,7 +59,6 @@
     Proof.context -> syntax -> bool -> term -> Pretty.T
   val standard_unparse_typ: Proof.context -> syntax -> typ -> Pretty.T
   val standard_unparse_sort: Proof.context -> syntax -> sort -> Pretty.T
-  val update_type_gram: (string * int * mixfix) list -> syntax -> syntax
   val update_consts: string list -> syntax -> syntax
   val update_trfuns:
     (string * ((ast list -> ast) * stamp)) list *
@@ -73,9 +72,8 @@
     (string * ((Proof.context -> ast list -> ast) * stamp)) list -> syntax -> syntax
   val extend_tokentrfuns: (string * string * (Proof.context -> string -> Pretty.T)) list ->
     syntax -> syntax
-  val update_const_gram: (string -> bool) ->
-    mode -> (string * typ * mixfix) list -> syntax -> syntax
-  val remove_const_gram: (string -> bool) ->
+  val update_type_gram: bool -> mode -> (string * typ * mixfix) list -> syntax -> syntax
+  val update_const_gram: bool -> (string -> bool) ->
     mode -> (string * typ * mixfix) list -> syntax -> syntax
   val update_trrules: Proof.context -> (string -> bool) -> syntax ->
     (string * string) trrule list -> syntax -> syntax
@@ -533,7 +531,7 @@
             \Retry with smaller Syntax.ambiguity_level for more information."
           else "";
 
-        val errs = map check ts;
+        val errs = Par_List.map check ts;
         val results = map_filter (fn (t, NONE) => SOME t | _ => NONE) (ts ~~ errs);
         val len = length results;
       in
@@ -669,17 +667,16 @@
 
 fun ext_syntax f decls = update_syntax mode_default (f decls);
 
-val update_type_gram       = ext_syntax Mixfix.syn_ext_types;
-val update_consts          = ext_syntax SynExt.syn_ext_const_names;
-val update_trfuns          = ext_syntax SynExt.syn_ext_trfuns;
+val update_consts = ext_syntax SynExt.syn_ext_const_names;
+val update_trfuns = ext_syntax SynExt.syn_ext_trfuns;
 val update_advanced_trfuns = ext_syntax SynExt.syn_ext_advanced_trfuns;
-val extend_tokentrfuns     = ext_syntax SynExt.syn_ext_tokentrfuns;
+val extend_tokentrfuns = ext_syntax SynExt.syn_ext_tokentrfuns;
 
-fun update_const_gram is_logtype prmode decls =
-  update_syntax prmode (Mixfix.syn_ext_consts is_logtype decls);
+fun update_type_gram add prmode decls =
+  (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_types decls);
 
-fun remove_const_gram is_logtype prmode decls =
-  remove_syntax prmode (Mixfix.syn_ext_consts is_logtype decls);
+fun update_const_gram add is_logtype prmode decls =
+  (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts is_logtype decls);
 
 fun update_trrules ctxt is_logtype syn =
   ext_syntax SynExt.syn_ext_rules o read_rules ctxt is_logtype syn;
--- a/src/Pure/Thy/thy_output.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Thy/thy_output.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -528,7 +528,7 @@
 val _ = basic_entity "term" (Term_Style.parse -- Args.term) pretty_term_style;
 val _ = basic_entity "term_type" (Term_Style.parse -- Args.term) pretty_term_typ;
 val _ = basic_entity "typeof" (Term_Style.parse -- Args.term) pretty_term_typeof;
-val _ = basic_entity "const" Args.const_proper pretty_const;
+val _ = basic_entity "const" (Args.const_proper false) pretty_const;
 val _ = basic_entity "abbrev" (Scan.lift Args.name_source) pretty_abbrev;
 val _ = basic_entity "typ" Args.typ_abbrev Syntax.pretty_typ;
 val _ = basic_entity "text" (Scan.lift Args.name) (K pretty_text);
--- a/src/Pure/Tools/find_theorems.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/Tools/find_theorems.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -364,7 +364,7 @@
 
 fun rem_thm_dups nicer xs =
   xs ~~ (1 upto length xs)
-  |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
+  |> sort (Term_Ord.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
   |> rem_cdups nicer
   |> sort (int_ord o pairself #2)
   |> map #1;
--- a/src/Pure/codegen.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/codegen.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -76,7 +76,7 @@
   val mk_term_of: codegr -> string -> bool -> typ -> Pretty.T
   val mk_gen: codegr -> string -> bool -> string list -> string -> typ -> Pretty.T
   val test_fn: (int -> term list option) Unsynchronized.ref
-  val test_term: Proof.context -> term -> int -> term list option
+  val test_term: Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
   val eval_result: (unit -> term) Unsynchronized.ref
   val eval_term: theory -> term -> term
   val evaluation_conv: cterm -> thm
@@ -871,7 +871,7 @@
 val test_fn : (int -> term list option) Unsynchronized.ref =
   Unsynchronized.ref (fn _ => NONE);
 
-fun test_term ctxt t =
+fun test_term ctxt report t =
   let
     val thy = ProofContext.theory_of ctxt;
     val (code, gr) = setmp_CRITICAL mode ["term_of", "test"]
@@ -897,7 +897,8 @@
           str ");"]) ^
       "\n\nend;\n";
     val _ = ML_Context.eval_in (SOME ctxt) false Position.none s;
-  in ! test_fn end;
+    val dummy_report = ([], false)
+  in (fn size => (! test_fn size, dummy_report)) end;
 
 
 
--- a/src/Pure/envir.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/envir.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -139,7 +139,7 @@
   (case t of
     Var (nT as (name', T)) =>
       if a = name' then env     (*cycle!*)
-      else if TermOrd.indexname_ord (a, name') = LESS then
+      else if Term_Ord.indexname_ord (a, name') = LESS then
         (case lookup (env, nT) of  (*if already assigned, chase*)
           NONE => update ((nT, Var (a, T)), env)
         | SOME u => vupdate ((aU, u), env))
--- a/src/Pure/facts.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/facts.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -171,7 +171,7 @@
 
 (* indexed props *)
 
-val prop_ord = TermOrd.term_ord o pairself Thm.full_prop_of;
+val prop_ord = Term_Ord.term_ord o pairself Thm.full_prop_of;
 
 fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props);
 fun could_unify (Facts {props, ...}) = Net.unify_term props;
--- a/src/Pure/meta_simplifier.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/meta_simplifier.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -755,7 +755,7 @@
   mk_eq_True = K NONE,
   reorient = default_reorient};
 
-val empty_ss = init_ss basic_mk_rews TermOrd.termless (K (K no_tac)) ([], []);
+val empty_ss = init_ss basic_mk_rews Term_Ord.termless (K (K no_tac)) ([], []);
 
 
 (* merge *)  (*NOTE: ignores some fields of 2nd simpset*)
--- a/src/Pure/more_thm.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/more_thm.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -151,12 +151,12 @@
     val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1;
     val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2;
   in
-    (case TermOrd.fast_term_ord (prop1, prop2) of
+    (case Term_Ord.fast_term_ord (prop1, prop2) of
       EQUAL =>
-        (case list_ord (prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord) (tpairs1, tpairs2) of
+        (case list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) (tpairs1, tpairs2) of
           EQUAL =>
-            (case list_ord TermOrd.fast_term_ord (hyps1, hyps2) of
-              EQUAL => list_ord TermOrd.sort_ord (shyps1, shyps2)
+            (case list_ord Term_Ord.fast_term_ord (hyps1, hyps2) of
+              EQUAL => list_ord Term_Ord.sort_ord (shyps1, shyps2)
             | ord => ord)
         | ord => ord)
     | ord => ord)
@@ -165,7 +165,7 @@
 
 (* tables and caches *)
 
-structure Ctermtab = Table(type key = cterm val ord = TermOrd.fast_term_ord o pairself Thm.term_of);
+structure Ctermtab = Table(type key = cterm val ord = Term_Ord.fast_term_ord o pairself Thm.term_of);
 structure Thmtab = Table(type key = thm val ord = thm_ord);
 
 fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f;
--- a/src/Pure/old_term.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/old_term.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -75,7 +75,7 @@
 
 (*Accumulates the Vars in the term, suppressing duplicates.*)
 fun add_term_vars (t, vars: term list) = case t of
-    Var   _ => OrdList.insert TermOrd.term_ord t vars
+    Var   _ => OrdList.insert Term_Ord.term_ord t vars
   | Abs (_,_,body) => add_term_vars(body,vars)
   | f$t =>  add_term_vars (f, add_term_vars(t, vars))
   | _ => vars;
@@ -84,7 +84,7 @@
 
 (*Accumulates the Frees in the term, suppressing duplicates.*)
 fun add_term_frees (t, frees: term list) = case t of
-    Free   _ => OrdList.insert TermOrd.term_ord t frees
+    Free   _ => OrdList.insert Term_Ord.term_ord t frees
   | Abs (_,_,body) => add_term_frees(body,frees)
   | f$t =>  add_term_frees (f, add_term_frees(t, frees))
   | _ => frees;
--- a/src/Pure/pattern.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/pattern.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -226,7 +226,7 @@
                      fun lam(is) = mkabs(binders,is,app(H,map (idx is) ks));
                  in Envir.update (((G, Gty), lam js), Envir.update (((F, Fty), lam is), env'))
                  end;
-  in if TermOrd.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
+  in if Term_Ord.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end
 
 fun unify_types thy (T, U) (env as Envir.Envir {maxidx, tenv, tyenv}) =
   if T = U then env
--- a/src/Pure/proofterm.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/proofterm.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -214,7 +214,7 @@
 
 (* proof body *)
 
-val oracle_ord = prod_ord fast_string_ord TermOrd.fast_term_ord;
+val oracle_ord = prod_ord fast_string_ord Term_Ord.fast_term_ord;
 fun thm_ord ((i, _): pthm, (j, _)) = int_ord (j, i);
 
 val merge_oracles = OrdList.union oracle_ord;
--- a/src/Pure/search.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/search.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -197,7 +197,7 @@
 structure Thm_Heap = Heap
 (
   type elem = int * thm;
-  val ord = prod_ord int_ord (TermOrd.term_ord o pairself Thm.prop_of);
+  val ord = prod_ord int_ord (Term_Ord.term_ord o pairself Thm.prop_of);
 );
 
 val trace_BEST_FIRST = Unsynchronized.ref false;
--- a/src/Pure/sign.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/sign.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -60,6 +60,7 @@
   val intern_term: theory -> term -> term
   val extern_term: theory -> term -> term
   val intern_tycons: theory -> typ -> typ
+  val the_type_decl: theory -> string -> Type.decl
   val arity_number: theory -> string -> int
   val arity_sorts: theory -> string -> sort -> sort list
   val certify_class: theory -> class -> class
@@ -88,6 +89,7 @@
   val add_modesyntax_i: Syntax.mode -> (string * typ * mixfix) list -> theory -> theory
   val del_modesyntax: Syntax.mode -> (string * string * mixfix) list -> theory -> theory
   val del_modesyntax_i: Syntax.mode -> (string * typ * mixfix) list -> theory -> theory
+  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> theory -> theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> theory -> theory
   val declare_const: (binding * typ) * mixfix -> theory -> term * theory
   val add_consts: (binding * string * mixfix) list -> theory -> theory
@@ -308,6 +310,7 @@
 
 (* certify wrt. type signature *)
 
+val the_type_decl = Type.the_decl o tsig_of;
 val arity_number = Type.arity_number o tsig_of;
 fun arity_sorts thy = Type.arity_sorts (Syntax.pp_global thy) (tsig_of thy);
 
@@ -402,8 +405,13 @@
 
 (* classes *)
 
-fun read_class thy c = certify_class thy (intern_class thy c)
-  handle TYPE (msg, _, _) => error msg;
+fun read_class thy text =
+  let
+    val (syms, pos) = Syntax.read_token text;
+    val c = certify_class thy (intern_class thy (Symbol_Pos.content syms))
+      handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos);
+    val _ = Position.report (Markup.tclass c) pos;
+  in c end;
 
 
 (* type arities *)
@@ -432,7 +440,9 @@
 
 fun add_types types thy = thy |> map_sign (fn (naming, syn, tsig, consts) =>
   let
-    val syn' = Syntax.update_type_gram (map (fn (a, n, mx) => (Name.of_binding a, n, mx)) types) syn;
+    val syn' =
+      Syntax.update_type_gram true Syntax.mode_default
+        (map (fn (a, n, mx) => (Name.of_binding a, Syntax.make_type n, mx)) types) syn;
     val decls = map (fn (a, n, _) => (a, n)) types;
     val tsig' = fold (Type.add_type naming) decls tsig;
   in (naming, syn', tsig', consts) end);
@@ -453,7 +463,9 @@
   thy |> map_sign (fn (naming, syn, tsig, consts) =>
     let
       val ctxt = ProofContext.init thy;
-      val syn' = Syntax.update_type_gram [(Name.of_binding b, length vs, mx)] syn;
+      val syn' =
+        Syntax.update_type_gram true Syntax.mode_default
+          [(Name.of_binding b, Syntax.make_type (length vs), mx)] syn;
       val abbr = (b, vs, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt rhs))
         handle ERROR msg => cat_error msg ("in type abbreviation " ^ quote (Binding.str_of b));
       val tsig' = Type.add_abbrev naming abbr tsig;
@@ -472,24 +484,30 @@
       handle ERROR msg => cat_error msg ("in syntax declaration " ^ quote c);
   in thy |> map_syn (change_gram (is_logtype thy) mode (map prep args)) end;
 
-fun gen_add_syntax x = gen_syntax Syntax.update_const_gram x;
+fun gen_add_syntax x = gen_syntax (Syntax.update_const_gram true) x;
 
 val add_modesyntax = gen_add_syntax Syntax.parse_typ;
 val add_modesyntax_i = gen_add_syntax (K I);
 val add_syntax = add_modesyntax Syntax.mode_default;
 val add_syntax_i = add_modesyntax_i Syntax.mode_default;
-val del_modesyntax = gen_syntax Syntax.remove_const_gram Syntax.parse_typ;
-val del_modesyntax_i = gen_syntax Syntax.remove_const_gram (K I);
+val del_modesyntax = gen_syntax (Syntax.update_const_gram false) Syntax.parse_typ;
+val del_modesyntax_i = gen_syntax (Syntax.update_const_gram false) (K I);
+
+fun type_notation add mode args =
+  let
+    fun type_syntax (Type (c, args), mx) =  (* FIXME authentic syntax *)
+          SOME (Long_Name.base_name c, Syntax.make_type (length args), mx)
+      | type_syntax _ = NONE;
+  in map_syn (Syntax.update_type_gram add mode (map_filter type_syntax args)) end;
 
 fun notation add mode args thy =
   let
-    val change_gram = if add then Syntax.update_const_gram else Syntax.remove_const_gram;
     fun const_syntax (Const (c, _), mx) =
           (case try (Consts.type_scheme (consts_of thy)) c of
             SOME T => SOME (Syntax.mark_const c, T, mx)
           | NONE => NONE)
       | const_syntax _ = NONE;
-  in gen_syntax change_gram (K I) mode (map_filter const_syntax args) thy end;
+  in gen_syntax (Syntax.update_const_gram add) (K I) mode (map_filter const_syntax args) thy end;
 
 
 (* add constants *)
--- a/src/Pure/sorts.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/sorts.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -71,13 +71,13 @@
 
 (** ordered lists of sorts **)
 
-val make = OrdList.make TermOrd.sort_ord;
-val subset = OrdList.subset TermOrd.sort_ord;
-val union = OrdList.union TermOrd.sort_ord;
-val subtract = OrdList.subtract TermOrd.sort_ord;
+val make = OrdList.make Term_Ord.sort_ord;
+val subset = OrdList.subset Term_Ord.sort_ord;
+val union = OrdList.union Term_Ord.sort_ord;
+val subtract = OrdList.subtract Term_Ord.sort_ord;
 
-val remove_sort = OrdList.remove TermOrd.sort_ord;
-val insert_sort = OrdList.insert TermOrd.sort_ord;
+val remove_sort = OrdList.remove Term_Ord.sort_ord;
+val insert_sort = OrdList.insert Term_Ord.sort_ord;
 
 fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss
   | insert_typ (TVar (_, S)) Ss = insert_sort S Ss
--- a/src/Pure/term_ord.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/term_ord.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -29,7 +29,7 @@
   val term_cache: (term -> 'a) -> term -> 'a
 end;
 
-structure TermOrd: TERM_ORD =
+structure Term_Ord: TERM_ORD =
 struct
 
 (* fast syntactic ordering -- tuned for inequalities *)
@@ -223,6 +223,11 @@
 
 end;
 
-structure Basic_Term_Ord: BASIC_TERM_ORD = TermOrd;
+structure Basic_Term_Ord: BASIC_TERM_ORD = Term_Ord;
 open Basic_Term_Ord;
 
+structure Var_Graph = Graph(type key = indexname val ord = Term_Ord.fast_indexname_ord);
+structure Sort_Graph = Graph(type key = sort val ord = Term_Ord.sort_ord);
+structure Typ_Graph = Graph(type key = typ val ord = Term_Ord.typ_ord);
+structure Term_Graph = Graph(type key = term val ord = Term_Ord.fast_term_ord);
+
--- a/src/Pure/term_subst.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/term_subst.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -198,9 +198,9 @@
 
 fun zero_var_indexes_inst ts =
   let
-    val tvars = sort TermOrd.tvar_ord (fold Term.add_tvars ts []);
+    val tvars = sort Term_Ord.tvar_ord (fold Term.add_tvars ts []);
     val instT = map (apsnd TVar) (zero_var_inst tvars);
-    val vars = sort TermOrd.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
+    val vars = sort Term_Ord.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
     val inst = map (apsnd Var) (zero_var_inst vars);
   in (instT, inst) end;
 
--- a/src/Pure/thm.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/thm.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -384,9 +384,9 @@
 
 fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop;
 
-val union_hyps = OrdList.union TermOrd.fast_term_ord;
-val insert_hyps = OrdList.insert TermOrd.fast_term_ord;
-val remove_hyps = OrdList.remove TermOrd.fast_term_ord;
+val union_hyps = OrdList.union Term_Ord.fast_term_ord;
+val insert_hyps = OrdList.insert Term_Ord.fast_term_ord;
+val remove_hyps = OrdList.remove Term_Ord.fast_term_ord;
 
 
 (* merge theories of cterms/thms -- trivial absorption only *)
--- a/src/Pure/type.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/type.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -35,6 +35,7 @@
   val get_mode: Proof.context -> mode
   val set_mode: mode -> Proof.context -> Proof.context
   val restore_mode: Proof.context -> Proof.context -> Proof.context
+  val the_decl: tsig -> string -> decl
   val cert_typ_mode: mode -> tsig -> typ -> typ
   val cert_typ: tsig -> typ -> typ
   val arity_number: tsig -> string -> int
@@ -163,6 +164,11 @@
 
 fun lookup_type (TSig {types = (_, types), ...}) = Symtab.lookup types;
 
+fun the_decl tsig c =
+  (case lookup_type tsig c of
+    NONE => error (undecl_type c)
+  | SOME decl => decl);
+
 
 (* certified types *)
 
@@ -189,15 +195,14 @@
             val Ts' = map cert Ts;
             fun nargs n = if length Ts <> n then err (bad_nargs c) else ();
           in
-            (case lookup_type tsig c of
-              SOME (LogicalType n) => (nargs n; Type (c, Ts'))
-            | SOME (Abbreviation (vs, U, syn)) =>
+            (case the_decl tsig c of
+              LogicalType n => (nargs n; Type (c, Ts'))
+            | Abbreviation (vs, U, syn) =>
                (nargs (length vs);
                 if syn then check_logical c else ();
                 if normalize then inst_typ (vs ~~ Ts') U
                 else Type (c, Ts'))
-            | SOME Nonterminal => (nargs 0; check_logical c; T)
-            | NONE => err (undecl_type c))
+            | Nonterminal => (nargs 0; check_logical c; T))
           end
       | cert (TFree (x, S)) = TFree (x, cert_sort tsig S)
       | cert (TVar (xi as (_, i), S)) =
--- a/src/Pure/unify.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Pure/unify.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -558,7 +558,7 @@
   if Term.eq_ix (v, w) then     (*...occur check would falsely return true!*)
       if T=U then (env, add_tpair (rbinder, (t,u), tpairs))
       else raise TERM ("add_ffpair: Var name confusion", [t,u])
-  else if TermOrd.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
+  else if Term_Ord.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)
        clean_ffpair thy ((rbinder, u, t), (env,tpairs))
         else clean_ffpair thy ((rbinder, t, u), (env,tpairs))
     | _ => raise TERM ("add_ffpair: Vars expected", [t,u])
--- a/src/Sequents/LK0.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Sequents/LK0.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -43,23 +43,23 @@
   not_equal  (infixl "~=" 50) where
   "x ~= y == ~ (x = y)"
 
-syntax (xsymbols)
-  Not           :: "o => o"               ("\<not> _" [40] 40)
-  conj          :: "[o, o] => o"          (infixr "\<and>" 35)
-  disj          :: "[o, o] => o"          (infixr "\<or>" 30)
-  imp           :: "[o, o] => o"          (infixr "\<longrightarrow>" 25)
-  iff           :: "[o, o] => o"          (infixr "\<longleftrightarrow>" 25)
-  All_binder    :: "[idts, o] => o"       ("(3\<forall>_./ _)" [0, 10] 10)
-  Ex_binder     :: "[idts, o] => o"       ("(3\<exists>_./ _)" [0, 10] 10)
-  not_equal     :: "['a, 'a] => o"        (infixl "\<noteq>" 50)
+notation (xsymbols)
+  Not  ("\<not> _" [40] 40) and
+  conj  (infixr "\<and>" 35) and
+  disj  (infixr "\<or>" 30) and
+  imp  (infixr "\<longrightarrow>" 25) and
+  iff  (infixr "\<longleftrightarrow>" 25) and
+  All  (binder "\<forall>" 10) and
+  Ex  (binder "\<exists>" 10) and
+  not_equal  (infixl "\<noteq>" 50)
 
-syntax (HTML output)
-  Not           :: "o => o"               ("\<not> _" [40] 40)
-  conj          :: "[o, o] => o"          (infixr "\<and>" 35)
-  disj          :: "[o, o] => o"          (infixr "\<or>" 30)
-  All_binder    :: "[idts, o] => o"       ("(3\<forall>_./ _)" [0, 10] 10)
-  Ex_binder     :: "[idts, o] => o"       ("(3\<exists>_./ _)" [0, 10] 10)
-  not_equal     :: "['a, 'a] => o"        (infixl "\<noteq>" 50)
+notation (HTML output)
+  Not  ("\<not> _" [40] 40) and
+  conj  (infixr "\<and>" 35) and
+  disj  (infixr "\<or>" 30) and
+  All  (binder "\<forall>" 10) and
+  Ex  (binder "\<exists>" 10) and
+  not_equal  (infixl "\<noteq>" 50)
 
 local
 
@@ -122,8 +122,7 @@
   The: "[| $H |- $E, P(a), $F;  !!x.$H, P(x) |- $E, x=a, $F |] ==>
           $H |- $E, P(THE x. P(x)), $F"
 
-constdefs
-  If :: "[o, 'a, 'a] => 'a"   ("(if (_)/ then (_)/ else (_))" 10)
+definition If :: "[o, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) where 
    "If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)"
 
 
--- a/src/Sequents/Sequents.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Sequents/Sequents.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -65,7 +65,7 @@
 
 (* parse translation for sequences *)
 
-fun abs_seq' t = Abs ("s", Type ("seq'", []), t);   (* FIXME @{type_syntax} *)
+fun abs_seq' t = Abs ("s", Type (@{type_syntax seq'}, []), t);
 
 fun seqobj_tr (Const (@{syntax_const "_SeqO"}, _) $ f) =
       Const (@{const_syntax SeqO'}, dummyT) $ f
--- a/src/Tools/Code/code_eval.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Tools/Code/code_eval.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -144,7 +144,7 @@
 
 val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq);
 val _ = ML_Context.add_antiq "code_datatype" (fn _ =>
-  (Args.tyname --| Scan.lift (Args.$$$ "=")
+  (Args.type_name true --| Scan.lift (Args.$$$ "=")
     -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term)))
       >> ml_code_datatype_antiq);
 
--- a/src/Tools/Compute_Oracle/linker.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Tools/Compute_Oracle/linker.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -50,7 +50,7 @@
   | constant_of _ = raise Link "constant_of"
 
 fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL)
-fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) TermOrd.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
+fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3))
 fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2))
 
 
@@ -70,7 +70,7 @@
     handle Type.TYPE_MATCH => NONE
 
 fun subst_ord (A:Type.tyenv, B:Type.tyenv) =
-    (list_ord (prod_ord TermOrd.fast_indexname_ord (prod_ord TermOrd.sort_ord TermOrd.typ_ord))) (Vartab.dest A, Vartab.dest B)
+    (list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B)
 
 structure Substtab = Table(type key = Type.tyenv val ord = subst_ord);
 
@@ -380,7 +380,7 @@
         ComputeThm (hyps, shyps, prop)
     end
 
-val cthm_ord' = prod_ord (prod_ord (list_ord TermOrd.term_ord) (list_ord TermOrd.sort_ord)) TermOrd.term_ord
+val cthm_ord' = prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord
 
 fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2))
 
--- a/src/Tools/induct.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Tools/induct.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -345,9 +345,9 @@
   Scan.lift (Args.$$$ k) >> K "";
 
 fun attrib add_type add_pred del =
-  spec typeN Args.tyname >> add_type ||
-  spec predN Args.const >> add_pred ||
-  spec setN Args.const >> add_pred ||
+  spec typeN (Args.type_name false) >> add_type ||
+  spec predN (Args.const false) >> add_pred ||
+  spec setN (Args.const false) >> add_pred ||
   Scan.lift Args.del >> K del;
 
 in
@@ -856,9 +856,9 @@
       | NONE => error ("No rule for " ^ k ^ " " ^ quote name))))));
 
 fun rule get_type get_pred =
-  named_rule typeN Args.tyname get_type ||
-  named_rule predN Args.const get_pred ||
-  named_rule setN Args.const get_pred ||
+  named_rule typeN (Args.type_name false) get_type ||
+  named_rule predN (Args.const false) get_pred ||
+  named_rule setN (Args.const false) get_pred ||
   Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms;
 
 val cases_rule = rule lookup_casesT lookup_casesP >> single_rule;
--- a/src/Tools/nbe.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Tools/nbe.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -164,6 +164,7 @@
   | same _ _ = false
 and sames xs ys = length xs = length ys andalso forall (uncurry same) (xs ~~ ys);
 
+
 (* constructor functions *)
 
 fun abss n f = Abs ((n, f), []);
@@ -213,6 +214,7 @@
         |> suffix "\n"
       end;
 
+
 (* nbe specific syntax and sandbox communication *)
 
 val univs_ref = Unsynchronized.ref (NONE : (unit -> Univ list -> Univ list) option);
@@ -255,6 +257,7 @@
 
 open Basic_Code_Thingol;
 
+
 (* code generation *)
 
 fun assemble_eqnss idx_of deps eqnss =
@@ -330,7 +333,7 @@
         val match_cont = if is_eval then NONE else SOME default_rhs;
         val assemble_arg = assemble_iterm
           (fn c => fn _ => fn ts => nbe_apps_constr idx_of c ts) NONE;
-        val assemble_rhs = assemble_iterm assemble_constapp match_cont ;
+        val assemble_rhs = assemble_iterm assemble_constapp match_cont;
         val (samepairs, args') = subst_nonlin_vars args;
         val s_args = map assemble_arg args';
         val s_rhs = if null samepairs then assemble_rhs rhs
@@ -357,6 +360,7 @@
     val deps_vars = ml_list (map (nbe_fun 0) deps);
   in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end;
 
+
 (* code compilation *)
 
 fun compile_eqnss _ gr raw_deps [] = []
@@ -457,6 +461,7 @@
     |> (fn t => apps t (rev dict_frees))
   end;
 
+
 (* reification *)
 
 fun typ_of_itype program vs (ityco `%% itys) =
@@ -480,9 +485,9 @@
       | is_dict (DFree _) = true
       | is_dict _ = false;
     fun const_of_idx idx = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx
-         of Code_Thingol.Fun (c, _) => c
-          | Code_Thingol.Datatypecons (c, _) => c
-          | Code_Thingol.Classparam (c, _) => c);
+     of Code_Thingol.Fun (c, _) => c
+      | Code_Thingol.Datatypecons (c, _) => c
+      | Code_Thingol.Classparam (c, _) => c);
     fun of_apps bounds (t, ts) =
       fold_map (of_univ bounds) ts
       #>> (fn ts' => list_comb (t, rev ts'))
@@ -503,6 +508,7 @@
           |-> (fn t' => pair (Term.Abs ("u", dummyT, t')))
   in of_univ 0 t 0 |> fst end;
 
+
 (* function store *)
 
 structure Nbe_Functions = Code_Data
@@ -511,6 +517,7 @@
   val empty = (Code_Thingol.empty_naming, (Graph.empty, (0, Inttab.empty)));
 );
 
+
 (* compilation, evaluation and reification *)
 
 fun compile_eval thy naming program vs_t deps =
@@ -524,6 +531,7 @@
     |> term_of_univ thy program idx_tab
   end;
 
+
 (* evaluation with type reconstruction *)
 
 fun normalize thy naming program ((vs0, (vs, ty)), t) deps =
@@ -593,6 +601,7 @@
 
 fun norm thy = lift_triv_classes_rew thy (no_frees_rew (Code_Thingol.eval thy I (normalize thy)));
 
+
 (* evaluation command *)
 
 fun norm_print_term ctxt modes t =
--- a/src/Tools/quickcheck.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/Tools/quickcheck.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -8,11 +8,16 @@
 sig
   val auto: bool Unsynchronized.ref
   val timing : bool Unsynchronized.ref
+  datatype report = Report of
+    { iterations : int, raised_match_errors : int,
+      satisfied_assms : int list, positive_concl_tests : int }
+  val gen_test_term: Proof.context -> bool -> bool -> string option -> int -> int -> term ->
+    (string * term) list option * ((string * int) list * ((int * report list) list) option)
   val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
     (string * term) list option
-  val timed_test_term: Proof.context -> bool -> string option -> int -> int -> term ->
-    ((string * term) list option * (string * int) list)
-  val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory
+  val add_generator:
+    string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))
+      -> theory -> theory
   val setup: theory -> theory
   val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option
 end;
@@ -33,31 +38,54 @@
       "auto-quickcheck"
       "Whether to run Quickcheck automatically.") ());
 
+(* quickcheck report *)
+
+datatype single_report = Run of bool list * bool | MatchExc
+
+datatype report = Report of
+  { iterations : int, raised_match_errors : int,
+    satisfied_assms : int list, positive_concl_tests : int }
+
+fun collect_single_report single_report
+    (Report {iterations = iterations, raised_match_errors = raised_match_errors,
+    satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) =
+  case single_report
+  of MatchExc =>
+    Report {iterations = iterations + 1, raised_match_errors = raised_match_errors + 1,
+      satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}
+   | Run (assms, concl) =>
+    Report {iterations = iterations + 1, raised_match_errors = raised_match_errors,
+      satisfied_assms =
+        map2 (fn b => fn s => if b then s + 1 else s) assms
+         (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms),
+      positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests}
 
 (* quickcheck configuration -- default parameters, test generators *)
 
 datatype test_params = Test_Params of
-  { size: int, iterations: int, default_type: typ option, no_assms: bool };
+  { size: int, iterations: int, default_type: typ option, no_assms: bool, report: bool, quiet : bool};
 
-fun dest_test_params (Test_Params { size, iterations, default_type, no_assms }) =
-  ((size, iterations), (default_type, no_assms));
-fun make_test_params ((size, iterations), (default_type, no_assms)) =
+fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
+  ((size, iterations), ((default_type, no_assms), (report, quiet)));
+fun make_test_params ((size, iterations), ((default_type, no_assms), (report, quiet))) =
   Test_Params { size = size, iterations = iterations, default_type = default_type,
-                no_assms = no_assms };
-fun map_test_params f (Test_Params { size, iterations, default_type, no_assms }) =
-  make_test_params (f ((size, iterations), (default_type, no_assms)));
+                no_assms = no_assms, report = report, quiet = quiet };
+fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, report, quiet }) =
+  make_test_params (f ((size, iterations), ((default_type, no_assms), (report, quiet))));
 fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
-                                     no_assms = no_assms1 },
+                                     no_assms = no_assms1, report = report1, quiet = quiet1 },
   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
-                no_assms = no_assms2 }) =
+                no_assms = no_assms2, report = report2, quiet = quiet2 }) =
   make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
-    (case default_type1 of NONE => default_type2 | _ => default_type1, no_assms1 orelse no_assms2));
+    ((case default_type1 of NONE => default_type2 | _ => default_type1, no_assms1 orelse no_assms2),
+    (report1 orelse report2, quiet1 orelse quiet2)));
 
 structure Data = Theory_Data
 (
-  type T = (string * (Proof.context -> term -> int -> term list option)) list
+  type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
     * test_params;
-  val empty = ([], Test_Params { size = 10, iterations = 100, default_type = NONE, no_assms = false });
+  val empty = ([], Test_Params
+    { size = 10, iterations = 100, default_type = NONE, no_assms = false, report = false, quiet = false});
   val extend = I;
   fun merge ((generators1, params1), (generators2, params2)) : T =
     (AList.merge (op =) (K true) (generators1, generators2),
@@ -66,7 +94,6 @@
 
 val add_generator = Data.map o apfst o AList.update (op =);
 
-
 (* generating tests *)
 
 fun mk_tester_select name ctxt =
@@ -74,14 +101,14 @@
    of NONE => error ("No such quickcheck generator: " ^ name)
     | SOME generator => generator ctxt;
 
-fun mk_testers ctxt t =
+fun mk_testers ctxt report t =
   (map snd o fst o Data.get o ProofContext.theory_of) ctxt
-  |> map_filter (fn generator => try (generator ctxt) t);
+  |> map_filter (fn generator => try (generator ctxt report) t);
 
-fun mk_testers_strict ctxt t =
+fun mk_testers_strict ctxt report t =
   let
     val generators = ((map snd o fst o Data.get o ProofContext.theory_of) ctxt)
-    val testers = map (fn generator => Exn.capture (generator ctxt) t) generators;
+    val testers = map (fn generator => Exn.capture (generator ctxt report) t) generators;
   in if forall (is_none o Exn.get_result) testers
     then [(Exn.release o snd o split_last) testers]
     else map_filter Exn.get_result testers
@@ -106,36 +133,45 @@
     val time = Time.toMilliseconds (#cpu (end_timing start))
   in (Exn.release result, (description, time)) end
 
-fun timed_test_term ctxt quiet generator_name size i t =
+fun gen_test_term ctxt quiet report generator_name size i t =
   let
     val (names, t') = prep_test_term t;
     val (testers, comp_time) = cpu_time "quickcheck compilation"
       (fn () => (case generator_name
-       of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
-        | SOME name => [mk_tester_select name ctxt t']));
-    fun iterate f 0 = NONE
-      | iterate f j = case f () handle Match => (if quiet then ()
-             else warning "Exception Match raised during quickcheck"; NONE)
-          of NONE => iterate f (j - 1) | SOME q => SOME q;
-    fun with_testers k [] = NONE
+       of NONE => if quiet then mk_testers ctxt report t' else mk_testers_strict ctxt report t'
+        | SOME name => [mk_tester_select name ctxt report t']));
+    fun iterate f 0 report = (NONE, report)
+      | iterate f j report =
+        let
+          val (test_result, single_report) = apsnd Run (f ()) handle Match => (if quiet then ()
+             else warning "Exception Match raised during quickcheck"; (NONE, MatchExc))
+          val report = collect_single_report single_report report
+        in
+          case test_result of NONE => iterate f (j - 1) report | SOME q => (SOME q, report)
+        end
+    val empty_report = Report { iterations = 0, raised_match_errors = 0,
+      satisfied_assms = [], positive_concl_tests = 0 }
+    fun with_testers k [] = (NONE, [])
       | with_testers k (tester :: testers) =
-          case iterate (fn () => tester (k - 1)) i
-           of NONE => with_testers k testers
-            | SOME q => SOME q;
-    fun with_size k = if k > size then NONE
+          case iterate (fn () => tester (k - 1)) i empty_report
+           of (NONE, report) => apsnd (cons report) (with_testers k testers)
+            | (SOME q, report) => (SOME q, [report]);
+    fun with_size k reports = if k > size then (NONE, reports)
       else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
-        case with_testers k testers
-         of NONE => with_size (k + 1) | SOME q => SOME q);
-    val (result, exec_time) = cpu_time "quickcheck execution"
-      (fn () => case with_size 1
-        of NONE => NONE
-        | SOME ts => SOME (names ~~ ts))
+        let
+          val (result, new_report) = with_testers k testers
+          val reports = ((k, new_report) :: reports)
+        in case result of NONE => with_size (k + 1) reports | SOME q => (SOME q, reports) end);
+    val ((result, reports), exec_time) = cpu_time "quickcheck execution"
+      (fn () => apfst
+         (fn result => case result of NONE => NONE
+        | SOME ts => SOME (names ~~ ts)) (with_size 1 []))
   in
-    (result, [exec_time, comp_time])
+    (result, ([exec_time, comp_time], if report then SOME reports else NONE))
   end;
 
 fun test_term ctxt quiet generator_name size i t =
-  fst (timed_test_term ctxt quiet generator_name size i t)
+  fst (gen_test_term ctxt quiet false generator_name size i t)
 
 fun monomorphic_term thy insts default_T = 
   let
@@ -152,7 +188,7 @@
       | subst T = T;
   in (map_types o map_atyps) subst end;
 
-fun test_goal quiet generator_name size iterations default_T no_assms insts i assms state =
+fun test_goal quiet report generator_name size iterations default_T no_assms insts i assms state =
   let
     val ctxt = Proof.context_of state;
     val thy = Proof.theory_of state;
@@ -164,7 +200,7 @@
                                   subst_bounds (frees, strip gi))
       |> monomorphic_term thy insts default_T
       |> ObjectLogic.atomize_term thy;
-  in test_term ctxt quiet generator_name size iterations gi' end;
+  in gen_test_term ctxt quiet report generator_name size iterations gi' end;
 
 fun pretty_counterex ctxt NONE = Pretty.str "Quickcheck found no counterexample."
   | pretty_counterex ctxt (SOME cex) =
@@ -172,6 +208,33 @@
         map (fn (s, t) =>
           Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex);
 
+fun pretty_report (Report {iterations = iterations, raised_match_errors = raised_match_errors,
+    satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) =
+  let
+    fun pretty_stat s i = Pretty.block ([Pretty.str (s ^ ": " ^ string_of_int i)])
+  in
+     ([pretty_stat "iterations" iterations,
+     pretty_stat "match exceptions" raised_match_errors]
+     @ map_index (fn (i, n) => pretty_stat ("satisfied " ^ string_of_int (i + 1) ^ ". assumption") n)
+       satisfied_assms
+     @ [pretty_stat "positive conclusion tests" positive_concl_tests])
+  end
+
+fun pretty_reports' [report] = [Pretty.chunks (pretty_report report)]
+  | pretty_reports' reports =
+  map_index (fn (i, report) =>
+    Pretty.chunks (Pretty.str (string_of_int (i + 1) ^ ". generator:\n") :: pretty_report report))
+    reports
+
+fun pretty_reports ctxt (SOME reports) =
+  Pretty.chunks (Pretty.str "Quickcheck report:" ::
+    maps (fn (size, reports) =>
+      Pretty.str ("size " ^ string_of_int size ^ ":") :: pretty_reports' reports @ [Pretty.brk 1])
+      (rev reports))
+  | pretty_reports ctxt NONE = Pretty.str ""
+
+fun pretty_counterex_and_reports ctxt (cex, (timing, reports)) =
+  Pretty.chunks [pretty_counterex ctxt cex, pretty_reports ctxt reports]
 
 (* automatic testing *)
 
@@ -182,15 +245,15 @@
     let
       val ctxt = Proof.context_of state;
       val assms = map term_of (Assumption.all_assms_of ctxt);
-      val Test_Params { size, iterations, default_type, no_assms } =
+      val Test_Params { size, iterations, default_type, no_assms, report, quiet } =
         (snd o Data.get o Proof.theory_of) state;
       val res =
-        try (test_goal true NONE size iterations default_type no_assms [] 1 assms) state;
+        try (test_goal true false NONE size iterations default_type no_assms [] 1 assms) state;
     in
       case res of
         NONE => (false, state)
-      | SOME NONE => (false, state)
-      | SOME cex => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "",
+      | SOME (NONE, report) => (false, state)
+      | SOME (cex, report) => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "",
           Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state)
     end
 
@@ -212,9 +275,13 @@
   | parse_test_param ctxt ("iterations", arg) =
       (apfst o apsnd o K) (read_nat arg)
   | parse_test_param ctxt ("default_type", arg) =
-      (apsnd o apfst o K o SOME) (ProofContext.read_typ ctxt arg)
+      (apsnd o apfst o apfst o K o SOME) (ProofContext.read_typ ctxt arg)
   | parse_test_param ctxt ("no_assms", arg) =
-      (apsnd o apsnd o K) (read_bool arg)
+      (apsnd o apfst o apsnd o K) (read_bool arg)
+  | parse_test_param ctxt ("report", arg) =
+      (apsnd o apsnd o apfst o K) (read_bool arg)
+  | parse_test_param ctxt ("quiet", arg) =
+      (apsnd o apsnd o apsnd o K) (read_bool arg)
   | parse_test_param ctxt (name, _) =
       error ("Unknown test parameter: " ^ name);
 
@@ -235,22 +302,24 @@
     |> (Data.map o apsnd o map_test_params) f
   end;
 
-fun quickcheck args i state =
+fun gen_quickcheck args i state =
   let
     val thy = Proof.theory_of state;
     val ctxt = Proof.context_of state;
     val assms = map term_of (Assumption.all_assms_of ctxt);
     val default_params = (dest_test_params o snd o Data.get) thy;
     val f = fold (parse_test_param_inst ctxt) args;
-    val (((size, iterations), (default_type, no_assms)), (generator_name, insts)) =
+    val (((size, iterations), ((default_type, no_assms), (report, quiet))), (generator_name, insts)) =
       f (default_params, (NONE, []));
   in
-    test_goal false generator_name size iterations default_type no_assms insts i assms state
+    test_goal quiet report generator_name size iterations default_type no_assms insts i assms state
   end;
 
+fun quickcheck args i state = fst (gen_quickcheck args i state)
+
 fun quickcheck_cmd args i state =
-  quickcheck args i (Toplevel.proof_of state)
-  |> Pretty.writeln o pretty_counterex (Toplevel.context_of state);
+  gen_quickcheck args i (Toplevel.proof_of state)
+  |> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state);
 
 local structure P = OuterParse and K = OuterKeyword in
 
--- a/src/ZF/Datatype_ZF.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Datatype_ZF.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -92,7 +92,7 @@
          else ();
        val goal = Logic.mk_equals (old, new)
        val thm = Goal.prove (Simplifier.the_context ss) [] [] goal
-         (fn _ => rtac iff_reflection 1 THEN
+         (fn _ => rtac @{thm iff_reflection} 1 THEN
            simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1)
          handle ERROR msg =>
          (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term_global sg goal);
--- a/src/ZF/Sum.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Sum.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -11,21 +11,20 @@
 
 global
 
-constdefs
-  sum     :: "[i,i]=>i"                     (infixr "+" 65)
+definition sum :: "[i,i]=>i" (infixr "+" 65) where
      "A+B == {0}*A Un {1}*B"
 
-  Inl     :: "i=>i"
+definition Inl :: "i=>i" where
      "Inl(a) == <0,a>"
 
-  Inr     :: "i=>i"
+definition Inr :: "i=>i" where
      "Inr(b) == <1,b>"
 
-  "case"  :: "[i=>i, i=>i, i]=>i"
+definition "case" :: "[i=>i, i=>i, i]=>i" where
      "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"
 
   (*operator for selecting out the various summands*)
-  Part    :: "[i,i=>i] => i"
+definition Part :: "[i,i=>i] => i" where
      "Part(A,h) == {x: A. EX z. x = h(z)}"
 
 local
--- a/src/ZF/Tools/datatype_package.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Tools/datatype_package.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -281,7 +281,7 @@
         list_comb (case_free, args)));
 
   val case_trans = hd con_defs RS @{thm def_trans}
-  and split_trans = Pr.split_eq RS meta_eq_to_obj_eq RS @{thm trans};
+  and split_trans = Pr.split_eq RS @{thm meta_eq_to_obj_eq} RS @{thm trans};
 
   fun prove_case_eqn (arg, con_def) =
     Goal.prove_global thy1 [] []
@@ -290,7 +290,9 @@
       (fn _ => EVERY
         [rewrite_goals_tac [con_def],
          rtac case_trans 1,
-         REPEAT (resolve_tac [refl, split_trans, Su.case_inl RS trans, Su.case_inr RS trans] 1)]);
+         REPEAT
+           (resolve_tac [@{thm refl}, split_trans,
+             Su.case_inl RS @{thm trans}, Su.case_inr RS @{thm trans}] 1)]);
 
   val free_iffs = map Drule.export_without_context (con_defs RL [@{thm def_swap_iff}]);
 
@@ -321,7 +323,7 @@
                          args)),
              subst_rec (Term.betapplys (recursor_case, args))));
 
-        val recursor_trans = recursor_def RS @{thm def_Vrecursor} RS trans;
+        val recursor_trans = recursor_def RS @{thm def_Vrecursor} RS @{thm trans};
 
         fun prove_recursor_eqn arg =
           Goal.prove_global thy1 [] []
--- a/src/ZF/Tools/induct_tacs.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Tools/induct_tacs.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -136,7 +136,7 @@
            rec_rewrites = recursor_eqns,
            case_rewrites = case_eqns,
            induct = induct,
-           mutual_induct = TrueI,  (*No need for mutual induction*)
+           mutual_induct = @{thm TrueI},  (*No need for mutual induction*)
            exhaustion = elim};
 
     val con_info =
--- a/src/ZF/Tools/inductive_package.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Tools/inductive_package.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -222,7 +222,7 @@
      rtac disjIn 2,
      (*Not ares_tac, since refl must be tried before equality assumptions;
        backtracking may occur if the premises have extra variables!*)
-     DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 2 APPEND assume_tac 2),
+     DEPTH_SOLVE_1 (resolve_tac [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac 2),
      (*Now solve the equations like Tcons(a,f) = Inl(?b4)*)
      rewrite_goals_tac con_defs,
      REPEAT (rtac @{thm refl} 2),
@@ -310,7 +310,7 @@
        Intro rules with extra Vars in premises still cause some backtracking *)
      fun ind_tac [] 0 = all_tac
        | ind_tac(prem::prems) i =
-             DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN ind_tac prems (i-1);
+             DEPTH_SOLVE_1 (ares_tac [prem, @{thm refl}] i) THEN ind_tac prems (i-1);
 
      val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT);
 
@@ -332,7 +332,7 @@
            setSolver (mk_solver "minimal"
                       (fn prems => resolve_tac (triv_rls@prems)
                                    ORELSE' assume_tac
-                                   ORELSE' etac FalseE));
+                                   ORELSE' etac @{thm FalseE}));
 
      val quant_induct =
        Goal.prove_global thy1 [] ind_prems
@@ -470,11 +470,11 @@
                       (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1
                    THEN
                    (*unpackage and use "prem" in the corresponding place*)
-                   REPEAT (rtac impI 1)  THEN
+                   REPEAT (rtac @{thm impI} 1)  THEN
                    rtac (rewrite_rule all_defs prem) 1  THEN
                    (*prem must not be REPEATed below: could loop!*)
-                   DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE'
-                                           eresolve_tac (conjE::mp::cmonos))))
+                   DEPTH_SOLVE (FIRSTGOAL (ares_tac [@{thm impI}] ORELSE'
+                                           eresolve_tac (@{thm conjE} :: @{thm mp} :: cmonos))))
                ) i)
             THEN mutual_ind_tac prems (i-1);
 
@@ -485,7 +485,7 @@
            (fn {prems, ...} => EVERY
              [rtac (quant_induct RS lemma) 1,
               mutual_ind_tac (rev prems) (length prems)])
-       else TrueI;
+       else @{thm TrueI};
 
      (** Uncurrying the predicate in the ordinary induction rule **)
 
@@ -498,7 +498,7 @@
                                       cterm_of thy1 elem_tuple)]);
 
      (*strip quantifier and the implication*)
-     val induct0 = inst (quant_induct RS spec RSN (2, @{thm rev_mp}));
+     val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
 
      val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
 
@@ -521,7 +521,7 @@
     else
       (thy1
       |> PureThy.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])]
-      |> apfst hd |> Library.swap, TrueI)
+      |> apfst hd |> Library.swap, @{thm TrueI})
   and defs = big_rec_def :: part_rec_defs
 
 
--- a/src/ZF/Tools/primrec_package.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Tools/primrec_package.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -175,7 +175,7 @@
     val eqn_thms =
       eqn_terms |> map (fn t =>
         Goal.prove_global thy1 [] [] (Ind_Syntax.traceIt "next primrec equation = " thy1 t)
-          (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1]));
+          (fn _ => EVERY [rewrite_goals_tac rewrites, rtac @{thm refl} 1]));
 
     val (eqn_thms', thy2) =
       thy1
--- a/src/ZF/Tools/typechk.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/Tools/typechk.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -98,7 +98,7 @@
 (*Instantiates variables in typing conditions.
   drawback: does not simplify conjunctions*)
 fun type_solver_tac ctxt hyps = SELECT_GOAL
-    (DEPTH_SOLVE (etac FalseE 1
+    (DEPTH_SOLVE (etac @{thm FalseE} 1
                   ORELSE basic_res_tac 1
                   ORELSE (ares_tac hyps 1
                           APPEND typecheck_step_tac (tcset_of ctxt))));
--- a/src/ZF/UNITY/Constrains.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/UNITY/Constrains.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -484,9 +484,9 @@
               REPEAT (force_tac css 2),
               full_simp_tac (ss addsimps (Program_Defs.get ctxt)) 1,
               ALLGOALS (clarify_tac cs),
-              REPEAT (FIRSTGOAL (etac disjE)),
+              REPEAT (FIRSTGOAL (etac @{thm disjE})),
               ALLGOALS (clarify_tac cs),
-              REPEAT (FIRSTGOAL (etac disjE)),
+              REPEAT (FIRSTGOAL (etac @{thm disjE})),
               ALLGOALS (clarify_tac cs),
               ALLGOALS (asm_full_simp_tac ss),
               ALLGOALS (clarify_tac cs)])
--- a/src/ZF/UNITY/SubstAx.thy	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/UNITY/SubstAx.thy	Tue Mar 02 04:35:44 2010 -0800
@@ -366,7 +366,7 @@
               (* proving the domain part *)
              clarify_tac cs 3, dtac @{thm swap} 3, force_tac css 4,
              rtac @{thm ReplaceI} 3, force_tac css 3, force_tac css 4,
-             asm_full_simp_tac ss 3, rtac conjI 3, simp_tac ss 4,
+             asm_full_simp_tac ss 3, rtac @{thm conjI} 3, simp_tac ss 4,
              REPEAT (rtac @{thm state_update_type} 3),
              constrains_tac ctxt 1,
              ALLGOALS (clarify_tac cs),
--- a/src/ZF/arith_data.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/arith_data.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -104,7 +104,7 @@
 
 (*Dummy version: the "coefficient" is always 1.
   In the result, the factors are sorted terms*)
-fun dest_coeff t = (1, mk_prod (sort TermOrd.term_ord (dest_prod t)));
+fun dest_coeff t = (1, mk_prod (sort Term_Ord.term_ord (dest_prod t)));
 
 (*Find first coefficient-term THAT MATCHES u*)
 fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
@@ -163,9 +163,9 @@
   val prove_conv = prove_conv "nateq_cancel_numerals"
   val mk_bal   = FOLogic.mk_eq
   val dest_bal = FOLogic.dest_eq
-  val bal_add1 = @{thm eq_add_iff} RS iff_trans
-  val bal_add2 = @{thm eq_add_iff} RS iff_trans
-  fun trans_tac _ = gen_trans_tac iff_trans
+  val bal_add1 = @{thm eq_add_iff} RS @{thm iff_trans}
+  val bal_add2 = @{thm eq_add_iff} RS @{thm iff_trans}
+  fun trans_tac _ = gen_trans_tac @{thm iff_trans}
   end;
 
 structure EqCancelNumerals = CancelNumeralsFun(EqCancelNumeralsData);
@@ -176,9 +176,9 @@
   val prove_conv = prove_conv "natless_cancel_numerals"
   val mk_bal   = FOLogic.mk_binrel "Ordinal.lt"
   val dest_bal = FOLogic.dest_bin "Ordinal.lt" iT
-  val bal_add1 = @{thm less_add_iff} RS iff_trans
-  val bal_add2 = @{thm less_add_iff} RS iff_trans
-  fun trans_tac _ = gen_trans_tac iff_trans
+  val bal_add1 = @{thm less_add_iff} RS @{thm iff_trans}
+  val bal_add2 = @{thm less_add_iff} RS @{thm iff_trans}
+  fun trans_tac _ = gen_trans_tac @{thm iff_trans}
   end;
 
 structure LessCancelNumerals = CancelNumeralsFun(LessCancelNumeralsData);
@@ -189,9 +189,9 @@
   val prove_conv = prove_conv "natdiff_cancel_numerals"
   val mk_bal   = FOLogic.mk_binop "Arith.diff"
   val dest_bal = FOLogic.dest_bin "Arith.diff" iT
-  val bal_add1 = @{thm diff_add_eq} RS trans
-  val bal_add2 = @{thm diff_add_eq} RS trans
-  fun trans_tac _ = gen_trans_tac trans
+  val bal_add1 = @{thm diff_add_eq} RS @{thm trans}
+  val bal_add2 = @{thm diff_add_eq} RS @{thm trans}
+  fun trans_tac _ = gen_trans_tac @{thm trans}
   end;
 
 structure DiffCancelNumerals = CancelNumeralsFun(DiffCancelNumeralsData);
--- a/src/ZF/int_arith.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/int_arith.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -95,7 +95,7 @@
 (*Express t as a product of (possibly) a numeral with other sorted terms*)
 fun dest_coeff sign (Const (@{const_name "zminus"}, _) $ t) = dest_coeff (~sign) t
   | dest_coeff sign t =
-    let val ts = sort TermOrd.term_ord (dest_prod t)
+    let val ts = sort Term_Ord.term_ord (dest_prod t)
         val (n, ts') = find_first_numeral [] ts
                           handle TERM _ => (1, ts)
     in (sign*n, mk_prod ts') end;
@@ -128,7 +128,7 @@
 
 (*To evaluate binary negations of coefficients*)
 val zminus_simps = @{thms NCons_simps} @
-                   [@{thm integ_of_minus} RS sym,
+                   [@{thm integ_of_minus} RS @{thm sym},
                     @{thm bin_minus_1}, @{thm bin_minus_0}, @{thm bin_minus_Pls}, @{thm bin_minus_Min},
                     @{thm bin_pred_1}, @{thm bin_pred_0}, @{thm bin_pred_Pls}, @{thm bin_pred_Min}];
 
@@ -144,7 +144,7 @@
 
 (*combine unary minus with numeric literals, however nested within a product*)
 val int_mult_minus_simps =
-    [@{thm zmult_assoc}, @{thm zmult_zminus} RS sym, int_minus_mult_eq_1_to_2];
+    [@{thm zmult_assoc}, @{thm zmult_zminus} RS @{thm sym}, int_minus_mult_eq_1_to_2];
 
 fun prep_simproc thy (name, pats, proc) =
   Simplifier.simproc thy name pats proc;
@@ -156,7 +156,7 @@
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
   val find_first_coeff  = find_first_coeff []
-  fun trans_tac _       = ArithData.gen_trans_tac iff_trans
+  fun trans_tac _       = ArithData.gen_trans_tac @{thm iff_trans}
 
   val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac}
   val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys
@@ -179,8 +179,8 @@
   val prove_conv = ArithData.prove_conv "inteq_cancel_numerals"
   val mk_bal   = FOLogic.mk_eq
   val dest_bal = FOLogic.dest_eq
-  val bal_add1 = @{thm eq_add_iff1} RS iff_trans
-  val bal_add2 = @{thm eq_add_iff2} RS iff_trans
+  val bal_add1 = @{thm eq_add_iff1} RS @{thm iff_trans}
+  val bal_add2 = @{thm eq_add_iff2} RS @{thm iff_trans}
 );
 
 structure LessCancelNumerals = CancelNumeralsFun
@@ -188,8 +188,8 @@
   val prove_conv = ArithData.prove_conv "intless_cancel_numerals"
   val mk_bal   = FOLogic.mk_binrel @{const_name "zless"}
   val dest_bal = FOLogic.dest_bin @{const_name "zless"} @{typ i}
-  val bal_add1 = @{thm less_add_iff1} RS iff_trans
-  val bal_add2 = @{thm less_add_iff2} RS iff_trans
+  val bal_add1 = @{thm less_add_iff1} RS @{thm iff_trans}
+  val bal_add2 = @{thm less_add_iff2} RS @{thm iff_trans}
 );
 
 structure LeCancelNumerals = CancelNumeralsFun
@@ -197,8 +197,8 @@
   val prove_conv = ArithData.prove_conv "intle_cancel_numerals"
   val mk_bal   = FOLogic.mk_binrel @{const_name "zle"}
   val dest_bal = FOLogic.dest_bin @{const_name "zle"} @{typ i}
-  val bal_add1 = @{thm le_add_iff1} RS iff_trans
-  val bal_add2 = @{thm le_add_iff2} RS iff_trans
+  val bal_add1 = @{thm le_add_iff1} RS @{thm iff_trans}
+  val bal_add2 = @{thm le_add_iff2} RS @{thm iff_trans}
 );
 
 val cancel_numerals =
@@ -232,9 +232,9 @@
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
-  val left_distrib      = @{thm left_zadd_zmult_distrib} RS trans
+  val left_distrib      = @{thm left_zadd_zmult_distrib} RS @{thm trans}
   val prove_conv        = prove_conv_nohyps "int_combine_numerals"
-  fun trans_tac _       = ArithData.gen_trans_tac trans
+  fun trans_tac _       = ArithData.gen_trans_tac @{thm trans}
 
   val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac} @ intifys
   val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys
@@ -274,14 +274,14 @@
   fun mk_coeff(k,t) = if t=one then mk_numeral k
                       else raise TERM("mk_coeff", [])
   fun dest_coeff t = (dest_numeral t, one)  (*We ONLY want pure numerals.*)
-  val left_distrib      = @{thm zmult_assoc} RS sym RS trans
+  val left_distrib      = @{thm zmult_assoc} RS @{thm sym} RS @{thm trans}
   val prove_conv        = prove_conv_nohyps "int_combine_numerals_prod"
-  fun trans_tac _       = ArithData.gen_trans_tac trans
+  fun trans_tac _       = ArithData.gen_trans_tac @{thm trans}
 
 
 
 val norm_ss1 = ZF_ss addsimps mult_1s @ diff_simps @ zminus_simps
-  val norm_ss2 = ZF_ss addsimps [@{thm zmult_zminus_right} RS sym] @
+  val norm_ss2 = ZF_ss addsimps [@{thm zmult_zminus_right} RS @{thm sym}] @
     bin_simps @ @{thms zmult_ac} @ tc_rules @ intifys
   fun norm_tac ss =
     ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1))
--- a/src/ZF/simpdata.ML	Tue Mar 02 04:31:50 2010 -0800
+++ b/src/ZF/simpdata.ML	Tue Mar 02 04:35:44 2010 -0800
@@ -32,9 +32,9 @@
 (*Analyse a rigid formula*)
 val ZF_conn_pairs =
   [("Ball",     [@{thm bspec}]),
-   ("All",      [spec]),
-   ("op -->",   [mp]),
-   ("op &",     [conjunct1,conjunct2])];
+   ("All",      [@{thm spec}]),
+   ("op -->",   [@{thm mp}]),
+   ("op &",     [@{thm conjunct1}, @{thm conjunct2}])];
 
 (*Analyse a:b, where b is rigid*)
 val ZF_mem_pairs =