modernized specifications;
authorwenzelm
Sun, 07 Oct 2007 21:19:31 +0200
changeset 24893 b8ef7afe3a6b
parent 24892 c663e675e177
child 24894 163c82d039cf
modernized specifications; removed legacy ML bindings;
src/ZF/AC.thy
src/ZF/AC/AC18_AC19.thy
src/ZF/AC/AC_Equiv.thy
src/ZF/AC/DC.thy
src/ZF/AC/Hartog.thy
src/ZF/AC/WO1_WO7.thy
src/ZF/AC/WO2_AC16.thy
src/ZF/AC/WO6_WO1.thy
src/ZF/Arith.thy
src/ZF/ArithSimp.thy
src/ZF/Bin.thy
src/ZF/Bool.thy
src/ZF/Cardinal.thy
src/ZF/CardinalArith.thy
src/ZF/Coind/ECR.thy
src/ZF/Coind/Map.thy
src/ZF/Coind/Static.thy
src/ZF/Coind/Values.thy
src/ZF/Datatype.thy
src/ZF/Epsilon.thy
src/ZF/EquivClass.thy
src/ZF/Finite.thy
src/ZF/Fixedpt.thy
src/ZF/Induct/Binary_Trees.thy
src/ZF/Induct/FoldSet.thy
src/ZF/Induct/Multiset.thy
src/ZF/Induct/Mutil.thy
src/ZF/Induct/Ntree.thy
src/ZF/Induct/Primrec.thy
src/ZF/Induct/PropLog.thy
src/ZF/Induct/Term.thy
src/ZF/Inductive.thy
src/ZF/Int.thy
src/ZF/IntDiv.thy
src/ZF/List.thy
src/ZF/Main.thy
src/ZF/Nat.thy
src/ZF/OrdQuant.thy
src/ZF/Order.thy
src/ZF/OrderArith.thy
src/ZF/OrderType.thy
src/ZF/Ordinal.thy
src/ZF/Perm.thy
src/ZF/QPair.thy
src/ZF/QUniv.thy
src/ZF/Resid/Confluence.thy
src/ZF/Resid/Residuals.thy
src/ZF/Resid/Substitution.thy
src/ZF/Sum.thy
src/ZF/Tools/datatype_package.ML
src/ZF/Tools/inductive_package.ML
src/ZF/Tools/typechk.ML
src/ZF/Trancl.thy
src/ZF/UNITY/AllocBase.thy
src/ZF/UNITY/AllocImpl.thy
src/ZF/UNITY/ClientImpl.thy
src/ZF/UNITY/Comp.thy
src/ZF/UNITY/Constrains.thy
src/ZF/UNITY/Distributor.thy
src/ZF/UNITY/FP.thy
src/ZF/UNITY/Follows.thy
src/ZF/UNITY/GenPrefix.thy
src/ZF/UNITY/Guar.thy
src/ZF/UNITY/Increasing.thy
src/ZF/UNITY/Merge.thy
src/ZF/UNITY/Monotonicity.thy
src/ZF/UNITY/MultisetSum.thy
src/ZF/UNITY/Mutex.thy
src/ZF/UNITY/State.thy
src/ZF/UNITY/SubstAx.thy
src/ZF/UNITY/UNITY.thy
src/ZF/UNITY/Union.thy
src/ZF/UNITY/WFair.thy
src/ZF/Univ.thy
src/ZF/WF.thy
src/ZF/ZF.thy
src/ZF/Zorn.thy
src/ZF/arith_data.ML
src/ZF/equalities.thy
src/ZF/ex/CoUnit.thy
src/ZF/ex/LList.thy
src/ZF/ex/Limit.thy
src/ZF/ex/Primes.thy
src/ZF/func.thy
src/ZF/ind_syntax.ML
src/ZF/int_arith.ML
src/ZF/pair.thy
src/ZF/simpdata.ML
src/ZF/upair.thy
--- a/src/ZF/AC.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -10,7 +10,8 @@
 theory AC imports Main begin
 
 text{*This definition comes from Halmos (1960), page 59.*}
-axioms AC: "[| a: A;  !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"
+axiomatization where
+  AC: "[| a: A;  !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"
 
 (*The same as AC, but no premise a \<in> A*)
 lemma AC_Pi: "[| !!x. x \<in> A ==> (\<exists>y. y \<in> B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
--- a/src/ZF/AC/AC18_AC19.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/AC18_AC19.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -7,8 +7,8 @@
 
 theory AC18_AC19 imports AC_Equiv begin
 
-constdefs
-  uu    :: "i => i"
+definition
+  uu    :: "i => i" where
     "uu(a) == {c Un {0}. c \<in> a}"
 
 
--- a/src/ZF/AC/AC_Equiv.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/AC_Equiv.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -14,107 +14,108 @@
 
 theory AC_Equiv imports Main begin (*obviously not Main_ZFC*)
 
-constdefs
-  
 (* Well Ordering Theorems *)
-  WO1 :: o
+
+definition  
     "WO1 == \<forall>A. \<exists>R. well_ord(A,R)"
 
-  WO2 :: o
+definition  
     "WO2 == \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
 
-  WO3 :: o
+definition  
     "WO3 == \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
 
-  WO4 :: "i => o"
+definition  
     "WO4(m) == \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &   
 		         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
 
-  WO5 :: o
+definition  
     "WO5 == \<exists>m \<in> nat. 1\<le>m & WO4(m)"
 
-  WO6 :: o
+definition  
     "WO6 == \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
 		               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
 
-  WO7 :: o
+definition  
     "WO7 == \<forall>A. Finite(A) <-> (\<forall>R. well_ord(A,R) --> well_ord(A,converse(R)))"
 
-  WO8 :: o
+definition  
     "WO8 == \<forall>A. (\<exists>f. f \<in> (\<Pi> X \<in> A. X)) --> (\<exists>R. well_ord(A,R))"
 
 
+definition
 (* Auxiliary concepts needed below *)
-  pairwise_disjoint :: "i => o"
+  pairwise_disjoint :: "i => o"  where
     "pairwise_disjoint(A) == \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 Int A2 \<noteq> 0 --> A1=A2"
 
-  sets_of_size_between :: "[i, i, i] => o"
+definition
+  sets_of_size_between :: "[i, i, i] => o"  where
     "sets_of_size_between(A,m,n) == \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
 
 
 (* Axioms of Choice *)  
-  AC0 :: o
+definition
     "AC0 == \<forall>A. \<exists>f. f \<in> (\<Pi> X \<in> Pow(A)-{0}. X)"
 
-  AC1 :: o
+definition
     "AC1 == \<forall>A. 0\<notin>A --> (\<exists>f. f \<in> (\<Pi> X \<in> A. X))"
 
-  AC2 :: o
+definition
     "AC2 == \<forall>A. 0\<notin>A & pairwise_disjoint(A)   
 		   --> (\<exists>C. \<forall>B \<in> A. \<exists>y. B Int C = {y})"
-  AC3 :: o
+definition
     "AC3 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Pi> x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
 
-  AC4 :: o
+definition
     "AC4 == \<forall>R A B. (R \<subseteq> A*B --> (\<exists>f. f \<in> (\<Pi> x \<in> domain(R). R``{x})))"
 
-  AC5 :: o
+definition
     "AC5 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g \<in> range(f)->A. \<forall>x \<in> domain(g). f`(g`x) = x"
 
-  AC6 :: o
+definition
     "AC6 == \<forall>A. 0\<notin>A --> (\<Pi> B \<in> A. B)\<noteq>0"
 
-  AC7 :: o
+definition
     "AC7 == \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) --> (\<Pi> B \<in> A. B) \<noteq> 0"
 
-  AC8 :: o
+definition
     "AC8 == \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)   
 		   --> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
 
-  AC9 :: o
+definition
     "AC9 == \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) -->   
 		   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
 
-  AC10 :: "i => o"
+definition
     "AC10(n) ==  \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->   
 		   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
 		   sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))"
 
-  AC11 :: o
+definition
     "AC11 == \<exists>n \<in> nat. 1\<le>n & AC10(n)"
 
-  AC12 :: o
+definition
     "AC12 == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->
   	         (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
 	              sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))"
 
-  AC13 :: "i => o"
+definition
     "AC13(m) == \<forall>A. 0\<notin>A --> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
 
-  AC14 :: o
+definition
     "AC14 == \<exists>m \<in> nat. 1\<le>m & AC13(m)"
 
-  AC15 :: o
+definition
     "AC15 == \<forall>A. 0\<notin>A --> 
                  (\<exists>m \<in> nat. 1\<le>m & (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m))"
 
-  AC16 :: "[i, i] => o"
+definition
     "AC16(n, k)  == 
        \<forall>A. ~Finite(A) -->   
 	   (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
 	   (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
 
-  AC17 :: o
+definition
     "AC17 == \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
 		   \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
 
@@ -124,8 +125,7 @@
       (\<Union>f \<in> \<Pi> a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a)))"
   --"AC18 cannot be expressed within the object-logic"
 
-constdefs
-  AC19 :: o
+definition
     "AC19 == \<forall>A. A\<noteq>0 & 0\<notin>A --> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
 		   (\<Union>f \<in> (\<Pi> B \<in> A. B). \<Inter>a \<in> A. f`a))"
 
--- a/src/ZF/AC/DC.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/DC.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -78,18 +78,19 @@
 by (fast elim!: not_emptyE)
 
 
-constdefs
-
-  DC  :: "i => o"
+definition
+  DC  :: "i => o"  where
     "DC(a) == \<forall>X R. R \<subseteq> Pow(X)*X  &
 		    (\<forall>Y \<in> Pow(X). Y \<prec> a --> (\<exists>x \<in> X. <Y,x> \<in> R)) 
 		    --> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
 
-  DC0 :: o
+definition
+  DC0 :: o  where
     "DC0 == \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R) 
                     --> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)"
 
-  ff  :: "[i, i, i, i] => i"
+definition
+  ff  :: "[i, i, i, i] => i"  where
     "ff(b, X, Q, R) ==
 	   transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
 
--- a/src/ZF/AC/Hartog.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/Hartog.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -7,8 +7,8 @@
 
 theory Hartog imports AC_Equiv begin
 
-constdefs
-  Hartog :: "i => i"
+definition
+  Hartog :: "i => i"  where
    "Hartog(X) == LEAST i. ~ i \<lesssim> X"
 
 lemma Ords_in_set: "\<forall>a. Ord(a) --> a \<in> X ==> P"
--- a/src/ZF/AC/WO1_WO7.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/WO1_WO7.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -11,8 +11,7 @@
 
 theory WO1_WO7 imports AC_Equiv begin
 
-constdefs
-  LEMMA :: o
+definition
     "LEMMA ==
      \<forall>X. ~Finite(X) --> (\<exists>R. well_ord(X,R) & ~well_ord(X,converse(R)))"
 
--- a/src/ZF/AC/WO2_AC16.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/WO2_AC16.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -18,8 +18,8 @@
 
 (**** A recursive definition used in the proof of WO2 ==> AC16 ****)
 
-constdefs
-  recfunAC16 :: "[i,i,i,i] => i"
+definition
+  recfunAC16 :: "[i,i,i,i] => i"  where
     "recfunAC16(f,h,i,a) == 
          transrec2(i, 0, 
               %g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
--- a/src/ZF/AC/WO6_WO1.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/AC/WO6_WO1.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -13,17 +13,20 @@
 
 theory WO6_WO1 imports Cardinal_aux begin
 
-constdefs
 (* Auxiliary definitions used in proof *)
-  NN  :: "i => i"
+definition
+  NN  :: "i => i"  where
     "NN(y) == {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a  &  
                         (\<Union>b<a. f`b) = y & (\<forall>b<a. f`b \<lesssim> m)}"
   
-  uu  :: "[i, i, i, i] => i"
+definition
+  uu  :: "[i, i, i, i] => i"  where
     "uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta"
 
-  (** Definitions for case 1 **)
-  vv1 :: "[i, i, i] => i"
+
+(** Definitions for case 1 **)
+definition
+  vv1 :: "[i, i, i] => i"  where
      "vv1(f,m,b) ==                                                
            let g = LEAST g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 & 
                                  domain(uu(f,b,g,d)) \<lesssim> m));      
@@ -31,21 +34,27 @@
                             domain(uu(f,b,g,d)) \<lesssim> m         
            in  if f`b \<noteq> 0 then domain(uu(f,b,g,d)) else 0"
 
-  ww1 :: "[i, i, i] => i"
+definition
+  ww1 :: "[i, i, i] => i"  where
      "ww1(f,m,b) == f`b - vv1(f,m,b)"
 
-  gg1 :: "[i, i, i] => i"
+definition
+  gg1 :: "[i, i, i] => i"  where
      "gg1(f,a,m) == \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
 
-  (** Definitions for case 2 **)
-  vv2 :: "[i, i, i, i] => i"
+
+(** Definitions for case 2 **)
+definition
+  vv2 :: "[i, i, i, i] => i"  where
      "vv2(f,b,g,s) ==   
            if f`g \<noteq> 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \<noteq> 0)`s} else 0"
 
-  ww2 :: "[i, i, i, i] => i"
+definition
+  ww2 :: "[i, i, i, i] => i"  where
      "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)"
 
-  gg2 :: "[i, i, i, i] => i"
+definition
+  gg2 :: "[i, i, i, i] => i"  where
      "gg2(f,a,b,s) ==
 	      \<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
 
--- a/src/ZF/Arith.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Arith.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -17,11 +17,12 @@
 
 text{*Proofs about elementary arithmetic: addition, multiplication, etc.*}
 
-constdefs
-  pred   :: "i=>i"    (*inverse of succ*)
+definition
+  pred   :: "i=>i"    (*inverse of succ*)  where
     "pred(y) == nat_case(0, %x. x, y)"
 
-  natify :: "i=>i"    (*coerces non-nats to nats*)
+definition
+  natify :: "i=>i"    (*coerces non-nats to nats*)  where
     "natify == Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a))
                                                     else 0)"
 
@@ -43,36 +44,41 @@
   "raw_mult(0, n) = 0"
   "raw_mult(succ(m), n) = raw_add (n, raw_mult(m, n))"
 
-constdefs
-  add :: "[i,i]=>i"                    (infixl "#+" 65)
+definition
+  add :: "[i,i]=>i"                    (infixl "#+" 65)  where
     "m #+ n == raw_add (natify(m), natify(n))"
 
-  diff :: "[i,i]=>i"                    (infixl "#-" 65)
+definition
+  diff :: "[i,i]=>i"                    (infixl "#-" 65)  where
     "m #- n == raw_diff (natify(m), natify(n))"
 
-  mult :: "[i,i]=>i"                    (infixl "#*" 70)
+definition
+  mult :: "[i,i]=>i"                    (infixl "#*" 70)  where
     "m #* n == raw_mult (natify(m), natify(n))"
 
-  raw_div  :: "[i,i]=>i"
+definition
+  raw_div  :: "[i,i]=>i"  where
     "raw_div (m, n) ==
        transrec(m, %j f. if j<n | n=0 then 0 else succ(f`(j#-n)))"
 
-  raw_mod  :: "[i,i]=>i"
+definition
+  raw_mod  :: "[i,i]=>i"  where
     "raw_mod (m, n) ==
        transrec(m, %j f. if j<n | n=0 then j else f`(j#-n))"
 
-  div  :: "[i,i]=>i"                    (infixl "div" 70)
+definition
+  div  :: "[i,i]=>i"                    (infixl "div" 70)  where
     "m div n == raw_div (natify(m), natify(n))"
 
-  mod  :: "[i,i]=>i"                    (infixl "mod" 70)
+definition
+  mod  :: "[i,i]=>i"                    (infixl "mod" 70)  where
     "m mod n == raw_mod (natify(m), natify(n))"
 
-syntax (xsymbols)
-  "mult"      :: "[i,i] => i"               (infixr "#\<times>" 70)
+notation (xsymbols)
+  mult  (infixr "#\<times>" 70)
 
-syntax (HTML output)
-  "mult"      :: "[i, i] => i"               (infixr "#\<times>" 70)
-
+notation (HTML output)
+  mult  (infixr "#\<times>" 70)
 
 declare rec_type [simp]
         nat_0_le [simp]
@@ -548,94 +554,4 @@
 
 lemmas nat_typechecks = rec_type nat_0I nat_1I nat_succI Ord_nat
 
-ML
-{*
-val pred_def = thm "pred_def";
-val raw_div_def = thm "raw_div_def";
-val raw_mod_def = thm "raw_mod_def";
-val div_def = thm "div_def";
-val mod_def = thm "mod_def";
-
-val zero_lt_natE = thm "zero_lt_natE";
-val pred_succ_eq = thm "pred_succ_eq";
-val natify_succ = thm "natify_succ";
-val natify_0 = thm "natify_0";
-val natify_non_succ = thm "natify_non_succ";
-val natify_in_nat = thm "natify_in_nat";
-val natify_ident = thm "natify_ident";
-val natify_eqE = thm "natify_eqE";
-val natify_idem = thm "natify_idem";
-val add_natify1 = thm "add_natify1";
-val add_natify2 = thm "add_natify2";
-val mult_natify1 = thm "mult_natify1";
-val mult_natify2 = thm "mult_natify2";
-val diff_natify1 = thm "diff_natify1";
-val diff_natify2 = thm "diff_natify2";
-val mod_natify1 = thm "mod_natify1";
-val mod_natify2 = thm "mod_natify2";
-val div_natify1 = thm "div_natify1";
-val div_natify2 = thm "div_natify2";
-val raw_add_type = thm "raw_add_type";
-val add_type = thm "add_type";
-val raw_mult_type = thm "raw_mult_type";
-val mult_type = thm "mult_type";
-val raw_diff_type = thm "raw_diff_type";
-val diff_type = thm "diff_type";
-val diff_0_eq_0 = thm "diff_0_eq_0";
-val diff_succ_succ = thm "diff_succ_succ";
-val diff_0 = thm "diff_0";
-val diff_le_self = thm "diff_le_self";
-val add_0_natify = thm "add_0_natify";
-val add_succ = thm "add_succ";
-val add_0 = thm "add_0";
-val add_assoc = thm "add_assoc";
-val add_0_right_natify = thm "add_0_right_natify";
-val add_succ_right = thm "add_succ_right";
-val add_0_right = thm "add_0_right";
-val add_commute = thm "add_commute";
-val add_left_commute = thm "add_left_commute";
-val raw_add_left_cancel = thm "raw_add_left_cancel";
-val add_left_cancel_natify = thm "add_left_cancel_natify";
-val add_left_cancel = thm "add_left_cancel";
-val add_le_elim1_natify = thm "add_le_elim1_natify";
-val add_le_elim1 = thm "add_le_elim1";
-val add_lt_elim1_natify = thm "add_lt_elim1_natify";
-val add_lt_elim1 = thm "add_lt_elim1";
-val add_lt_mono1 = thm "add_lt_mono1";
-val add_lt_mono2 = thm "add_lt_mono2";
-val add_lt_mono = thm "add_lt_mono";
-val Ord_lt_mono_imp_le_mono = thm "Ord_lt_mono_imp_le_mono";
-val add_le_mono1 = thm "add_le_mono1";
-val add_le_mono = thm "add_le_mono";
-val diff_add_inverse = thm "diff_add_inverse";
-val diff_add_inverse2 = thm "diff_add_inverse2";
-val diff_cancel = thm "diff_cancel";
-val diff_cancel2 = thm "diff_cancel2";
-val diff_add_0 = thm "diff_add_0";
-val eq_add_iff = thm "eq_add_iff";
-val less_add_iff = thm "less_add_iff";
-val diff_add_eq = thm "diff_add_eq";
-val eq_cong2 = thm "eq_cong2";
-val iff_cong2 = thm "iff_cong2";
-val mult_0 = thm "mult_0";
-val mult_succ = thm "mult_succ";
-val mult_0_right = thm "mult_0_right";
-val mult_succ_right = thm "mult_succ_right";
-val mult_1_natify = thm "mult_1_natify";
-val mult_1_right_natify = thm "mult_1_right_natify";
-val mult_1 = thm "mult_1";
-val mult_1_right = thm "mult_1_right";
-val mult_commute = thm "mult_commute";
-val add_mult_distrib = thm "add_mult_distrib";
-val add_mult_distrib_left = thm "add_mult_distrib_left";
-val mult_assoc = thm "mult_assoc";
-val mult_left_commute = thm "mult_left_commute";
-val lt_succ_eq_0_disj = thm "lt_succ_eq_0_disj";
-val less_diff_conv = thm "less_diff_conv";
-
-val add_ac = thms "add_ac";
-val mult_ac = thms "mult_ac";
-val nat_typechecks = thms "nat_typechecks";
-*}
-
 end
--- a/src/ZF/ArithSimp.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ArithSimp.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -12,7 +12,6 @@
 uses "~~/src/Provers/Arith/cancel_numerals.ML"
       "~~/src/Provers/Arith/combine_numerals.ML"
       "arith_data.ML"
-
 begin
 
 subsection{*Difference*}
@@ -567,81 +566,4 @@
 apply (blast intro: lt_imp_diff_lt diff_lt_imp_lt) 
 done
 
-
-ML
-{*
-val diff_self_eq_0 = thm "diff_self_eq_0";
-val add_diff_inverse = thm "add_diff_inverse";
-val add_diff_inverse2 = thm "add_diff_inverse2";
-val diff_succ = thm "diff_succ";
-val zero_less_diff = thm "zero_less_diff";
-val diff_mult_distrib = thm "diff_mult_distrib";
-val diff_mult_distrib2 = thm "diff_mult_distrib2";
-val div_termination = thm "div_termination";
-val raw_mod_type = thm "raw_mod_type";
-val mod_type = thm "mod_type";
-val DIVISION_BY_ZERO_DIV = thm "DIVISION_BY_ZERO_DIV";
-val DIVISION_BY_ZERO_MOD = thm "DIVISION_BY_ZERO_MOD";
-val raw_mod_less = thm "raw_mod_less";
-val mod_less = thm "mod_less";
-val raw_mod_geq = thm "raw_mod_geq";
-val mod_geq = thm "mod_geq";
-val raw_div_type = thm "raw_div_type";
-val div_type = thm "div_type";
-val raw_div_less = thm "raw_div_less";
-val div_less = thm "div_less";
-val raw_div_geq = thm "raw_div_geq";
-val div_geq = thm "div_geq";
-val mod_div_equality_natify = thm "mod_div_equality_natify";
-val mod_div_equality = thm "mod_div_equality";
-val mod_succ = thm "mod_succ";
-val mod_less_divisor = thm "mod_less_divisor";
-val mod_1_eq = thm "mod_1_eq";
-val mod2_cases = thm "mod2_cases";
-val mod2_succ_succ = thm "mod2_succ_succ";
-val mod2_add_more = thm "mod2_add_more";
-val mod2_add_self = thm "mod2_add_self";
-val add_le_self = thm "add_le_self";
-val add_le_self2 = thm "add_le_self2";
-val mult_le_mono1 = thm "mult_le_mono1";
-val mult_le_mono = thm "mult_le_mono";
-val mult_lt_mono2 = thm "mult_lt_mono2";
-val mult_lt_mono1 = thm "mult_lt_mono1";
-val add_eq_0_iff = thm "add_eq_0_iff";
-val zero_lt_mult_iff = thm "zero_lt_mult_iff";
-val mult_eq_1_iff = thm "mult_eq_1_iff";
-val mult_is_zero = thm "mult_is_zero";
-val mult_is_zero_natify = thm "mult_is_zero_natify";
-val mult_less_cancel2 = thm "mult_less_cancel2";
-val mult_less_cancel1 = thm "mult_less_cancel1";
-val mult_le_cancel2 = thm "mult_le_cancel2";
-val mult_le_cancel1 = thm "mult_le_cancel1";
-val mult_le_cancel_le1 = thm "mult_le_cancel_le1";
-val Ord_eq_iff_le = thm "Ord_eq_iff_le";
-val mult_cancel2 = thm "mult_cancel2";
-val mult_cancel1 = thm "mult_cancel1";
-val div_cancel_raw = thm "div_cancel_raw";
-val div_cancel = thm "div_cancel";
-val mult_mod_distrib_raw = thm "mult_mod_distrib_raw";
-val mod_mult_distrib2 = thm "mod_mult_distrib2";
-val mult_mod_distrib = thm "mult_mod_distrib";
-val mod_add_self2_raw = thm "mod_add_self2_raw";
-val mod_add_self2 = thm "mod_add_self2";
-val mod_add_self1 = thm "mod_add_self1";
-val mod_mult_self1_raw = thm "mod_mult_self1_raw";
-val mod_mult_self1 = thm "mod_mult_self1";
-val mod_mult_self2 = thm "mod_mult_self2";
-val mult_eq_self_implies_10 = thm "mult_eq_self_implies_10";
-val less_imp_succ_add = thm "less_imp_succ_add";
-val less_iff_succ_add = thm "less_iff_succ_add";
-val diff_is_0_iff = thm "diff_is_0_iff";
-val nat_lt_imp_diff_eq_0 = thm "nat_lt_imp_diff_eq_0";
-val nat_diff_split = thm "nat_diff_split";
-
-val add_lt_elim2 = thm "add_lt_elim2";
-val add_le_elim2 = thm "add_le_elim2";
-
-val diff_lt_iff_lt = thm "diff_lt_iff_lt";
-*}
-
 end
--- a/src/ZF/Bin.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Bin.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -91,8 +91,8 @@
 				x xor y)"
 *)
 
-constdefs
-  bin_add   :: "[i,i]=>i"
+definition
+  bin_add   :: "[i,i]=>i"  where
     "bin_add(v,w) == bin_adder(v)`w"
 
 
@@ -595,98 +595,4 @@
      "(integ_of(w) = x $* y) <-> (x $* y = integ_of(w))"
 by auto
 
-ML
-{*
-val bin_pred_Pls = thm "bin_pred_Pls";
-val bin_pred_Min = thm "bin_pred_Min";
-val bin_minus_Pls = thm "bin_minus_Pls";
-val bin_minus_Min = thm "bin_minus_Min";
-
-val NCons_Pls_0 = thm "NCons_Pls_0";
-val NCons_Pls_1 = thm "NCons_Pls_1";
-val NCons_Min_0 = thm "NCons_Min_0";
-val NCons_Min_1 = thm "NCons_Min_1";
-val NCons_BIT = thm "NCons_BIT";
-val NCons_simps = thms "NCons_simps";
-val integ_of_type = thm "integ_of_type";
-val NCons_type = thm "NCons_type";
-val bin_succ_type = thm "bin_succ_type";
-val bin_pred_type = thm "bin_pred_type";
-val bin_minus_type = thm "bin_minus_type";
-val bin_add_type = thm "bin_add_type";
-val bin_mult_type = thm "bin_mult_type";
-val integ_of_NCons = thm "integ_of_NCons";
-val integ_of_succ = thm "integ_of_succ";
-val integ_of_pred = thm "integ_of_pred";
-val integ_of_minus = thm "integ_of_minus";
-val bin_add_Pls = thm "bin_add_Pls";
-val bin_add_Pls_right = thm "bin_add_Pls_right";
-val bin_add_Min = thm "bin_add_Min";
-val bin_add_Min_right = thm "bin_add_Min_right";
-val bin_add_BIT_Pls = thm "bin_add_BIT_Pls";
-val bin_add_BIT_Min = thm "bin_add_BIT_Min";
-val bin_add_BIT_BIT = thm "bin_add_BIT_BIT";
-val integ_of_add = thm "integ_of_add";
-val diff_integ_of_eq = thm "diff_integ_of_eq";
-val integ_of_mult = thm "integ_of_mult";
-val bin_succ_1 = thm "bin_succ_1";
-val bin_succ_0 = thm "bin_succ_0";
-val bin_pred_1 = thm "bin_pred_1";
-val bin_pred_0 = thm "bin_pred_0";
-val bin_minus_1 = thm "bin_minus_1";
-val bin_minus_0 = thm "bin_minus_0";
-val bin_add_BIT_11 = thm "bin_add_BIT_11";
-val bin_add_BIT_10 = thm "bin_add_BIT_10";
-val bin_add_BIT_0 = thm "bin_add_BIT_0";
-val bin_mult_1 = thm "bin_mult_1";
-val bin_mult_0 = thm "bin_mult_0";
-val int_of_0 = thm "int_of_0";
-val int_of_succ = thm "int_of_succ";
-val zminus_0 = thm "zminus_0";
-val zadd_0_intify = thm "zadd_0_intify";
-val zadd_0_right_intify = thm "zadd_0_right_intify";
-val zmult_1_intify = thm "zmult_1_intify";
-val zmult_1_right_intify = thm "zmult_1_right_intify";
-val zmult_0 = thm "zmult_0";
-val zmult_0_right = thm "zmult_0_right";
-val zmult_minus1 = thm "zmult_minus1";
-val zmult_minus1_right = thm "zmult_minus1_right";
-val eq_integ_of_eq = thm "eq_integ_of_eq";
-val iszero_integ_of_Pls = thm "iszero_integ_of_Pls";
-val nonzero_integ_of_Min = thm "nonzero_integ_of_Min";
-val iszero_integ_of_BIT = thm "iszero_integ_of_BIT";
-val iszero_integ_of_0 = thm "iszero_integ_of_0";
-val iszero_integ_of_1 = thm "iszero_integ_of_1";
-val less_integ_of_eq_neg = thm "less_integ_of_eq_neg";
-val not_neg_integ_of_Pls = thm "not_neg_integ_of_Pls";
-val neg_integ_of_Min = thm "neg_integ_of_Min";
-val neg_integ_of_BIT = thm "neg_integ_of_BIT";
-val le_integ_of_eq_not_less = thm "le_integ_of_eq_not_less";
-val bin_arith_extra_simps = thms "bin_arith_extra_simps";
-val bin_arith_simps = thms "bin_arith_simps";
-val bin_rel_simps = thms "bin_rel_simps";
-val add_integ_of_left = thm "add_integ_of_left";
-val mult_integ_of_left = thm "mult_integ_of_left";
-val add_integ_of_diff1 = thm "add_integ_of_diff1";
-val add_integ_of_diff2 = thm "add_integ_of_diff2";
-val zdiff0 = thm "zdiff0";
-val zdiff0_right = thm "zdiff0_right";
-val zdiff_self = thm "zdiff_self";
-val znegative_iff_zless_0 = thm "znegative_iff_zless_0";
-val zero_zless_imp_znegative_zminus = thm "zero_zless_imp_znegative_zminus";
-val zero_zle_int_of = thm "zero_zle_int_of";
-val nat_of_0 = thm "nat_of_0";
-val nat_le_int0 = thm "nat_le_int0";
-val int_of_eq_0_imp_natify_eq_0 = thm "int_of_eq_0_imp_natify_eq_0";
-val nat_of_zminus_int_of = thm "nat_of_zminus_int_of";
-val int_of_nat_of = thm "int_of_nat_of";
-val int_of_nat_of_if = thm "int_of_nat_of_if";
-val zless_nat_iff_int_zless = thm "zless_nat_iff_int_zless";
-val zless_nat_conj = thm "zless_nat_conj";
-val integ_of_minus_reorient = thm "integ_of_minus_reorient";
-val integ_of_add_reorient = thm "integ_of_add_reorient";
-val integ_of_diff_reorient = thm "integ_of_diff_reorient";
-val integ_of_mult_reorient = thm "integ_of_mult_reorient";
-*}
-
 end
--- a/src/ZF/Bool.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Bool.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -19,23 +19,22 @@
 
 text{*2 is equal to bool, but is used as a number rather than a type.*}
 
-constdefs
-  bool        :: i
-    "bool == {0,1}"
+definition "bool == {0,1}"
+
+definition "cond(b,c,d) == if(b=1,c,d)"
 
-  cond        :: "[i,i,i]=>i"
-    "cond(b,c,d) == if(b=1,c,d)"
+definition "not(b) == cond(b,0,1)"
 
-  not         :: "i=>i"
-    "not(b) == cond(b,0,1)"
-
-  "and"       :: "[i,i]=>i"      (infixl "and" 70)
+definition
+  "and"       :: "[i,i]=>i"      (infixl "and" 70)  where
     "a and b == cond(a,b,0)"
 
-  or          :: "[i,i]=>i"      (infixl "or" 65)
+definition
+  or          :: "[i,i]=>i"      (infixl "or" 65)  where
     "a or b == cond(a,1,b)"
 
-  xor         :: "[i,i]=>i"      (infixl "xor" 65)
+definition
+  xor         :: "[i,i]=>i"      (infixl "xor" 65) where
     "a xor b == cond(a,not(b),b)"
 
 
@@ -154,7 +153,8 @@
 by (elim boolE, auto)
 
 
-constdefs bool_of_o :: "o=>i"
+definition
+  bool_of_o :: "o=>i" where
    "bool_of_o(P) == (if P then 1 else 0)"
 
 lemma [simp]: "bool_of_o(True) = 1"
--- a/src/ZF/Cardinal.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Cardinal.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,28 +9,33 @@
 
 theory Cardinal imports OrderType Finite Nat Sum begin
 
-constdefs
-
+definition
   (*least ordinal operator*)
-   Least    :: "(i=>o) => i"    (binder "LEAST " 10)
+   Least    :: "(i=>o) => i"    (binder "LEAST " 10)  where
      "Least(P) == THE i. Ord(i) & P(i) & (ALL j. j<i --> ~P(j))"
 
-  eqpoll   :: "[i,i] => o"     (infixl "eqpoll" 50)
+definition
+  eqpoll   :: "[i,i] => o"     (infixl "eqpoll" 50)  where
     "A eqpoll B == EX f. f: bij(A,B)"
 
-  lepoll   :: "[i,i] => o"     (infixl "lepoll" 50)
+definition
+  lepoll   :: "[i,i] => o"     (infixl "lepoll" 50)  where
     "A lepoll B == EX f. f: inj(A,B)"
 
-  lesspoll :: "[i,i] => o"     (infixl "lesspoll" 50)
+definition
+  lesspoll :: "[i,i] => o"     (infixl "lesspoll" 50)  where
     "A lesspoll B == A lepoll B & ~(A eqpoll B)"
 
-  cardinal :: "i=>i"           ("|_|")
+definition
+  cardinal :: "i=>i"           ("|_|")  where
     "|A| == LEAST i. i eqpoll A"
 
-  Finite   :: "i=>o"
+definition
+  Finite   :: "i=>o"  where
     "Finite(A) == EX n:nat. A eqpoll n"
 
-  Card     :: "i=>o"
+definition
+  Card     :: "i=>o"  where
     "Card(i) == (i = |i|)"
 
 notation (xsymbols)
--- a/src/ZF/CardinalArith.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/CardinalArith.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,40 +9,45 @@
 
 theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin
 
-constdefs
-
-  InfCard       :: "i=>o"
+definition
+  InfCard       :: "i=>o"  where
     "InfCard(i) == Card(i) & nat le i"
 
-  cmult         :: "[i,i]=>i"       (infixl "|*|" 70)
+definition
+  cmult         :: "[i,i]=>i"       (infixl "|*|" 70)  where
     "i |*| j == |i*j|"
   
-  cadd          :: "[i,i]=>i"       (infixl "|+|" 65)
+definition
+  cadd          :: "[i,i]=>i"       (infixl "|+|" 65)  where
     "i |+| j == |i+j|"
 
-  csquare_rel   :: "i=>i"
+definition
+  csquare_rel   :: "i=>i"  where
     "csquare_rel(K) ==   
 	  rvimage(K*K,   
 		  lam <x,y>:K*K. <x Un y, x, y>, 
 		  rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))"
 
-  jump_cardinal :: "i=>i"
+definition
+  jump_cardinal :: "i=>i"  where
     --{*This def is more complex than Kunen's but it more easily proved to
         be a cardinal*}
     "jump_cardinal(K) ==   
          \<Union>X\<in>Pow(K). {z. r: Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
   
-  csucc         :: "i=>i"
+definition
+  csucc         :: "i=>i"  where
     --{*needed because @{term "jump_cardinal(K)"} might not be the successor
         of @{term K}*}
     "csucc(K) == LEAST L. Card(L) & K<L"
 
-syntax (xsymbols)
-  "op |+|"     :: "[i,i] => i"          (infixl "\<oplus>" 65)
-  "op |*|"     :: "[i,i] => i"          (infixl "\<otimes>" 70)
-syntax (HTML output)
-  "op |+|"     :: "[i,i] => i"          (infixl "\<oplus>" 65)
-  "op |*|"     :: "[i,i] => i"          (infixl "\<otimes>" 70)
+notation (xsymbols output)
+  cadd  (infixl "\<oplus>" 65) and
+  cmult  (infixl "\<otimes>" 70)
+
+notation (HTML output)
+  cadd  (infixl "\<oplus>" 65) and
+  cmult  (infixl "\<otimes>" 70)
 
 
 lemma Card_Union [simp,intro,TC]: "(ALL x:A. Card(x)) ==> Card(Union(A))"
--- a/src/ZF/Coind/ECR.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Coind/ECR.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -26,8 +26,8 @@
   
 (* Pointwise extension to environments *)
  
-constdefs
-  hastyenv :: "[i,i] => o"
+definition
+  hastyenv :: "[i,i] => o"  where
     "hastyenv(ve,te) ==                        
 	 ve_dom(ve) = te_dom(te) &          
 	 (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
--- a/src/ZF/Coind/Map.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Coind/Map.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -10,22 +10,26 @@
 
 theory Map imports Main begin
 
-constdefs
-  TMap :: "[i,i] => i"
+definition
+  TMap :: "[i,i] => i"  where
    "TMap(A,B) == {m \<in> Pow(A*Union(B)).\<forall>a \<in> A. m``{a} \<in> B}"
 
-  PMap :: "[i,i] => i"
+definition
+  PMap :: "[i,i] => i"  where
    "PMap(A,B) == TMap(A,cons(0,B))"
 
 (* Note: 0 \<in> B ==> TMap(A,B) = PMap(A,B) *)
   
-  map_emp :: i
+definition
+  map_emp :: i  where
    "map_emp == 0"
 
-  map_owr :: "[i,i,i]=>i"
+definition
+  map_owr :: "[i,i,i]=>i"  where
    "map_owr(m,a,b) == \<Sigma> x \<in> {a} Un domain(m). if x=a then b else m``{x}"
 
-  map_app :: "[i,i]=>i"
+definition
+  map_app :: "[i,i]=>i"  where
    "map_app(m,a) == m``{a}"
   
 lemma "{0,1} \<subseteq> {1} Un TMap(I, {0,1})"
--- a/src/ZF/Coind/Static.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Coind/Static.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -18,8 +18,8 @@
 
 (*Its extension to environments*)
 
-constdefs
-  isofenv :: "[i,i] => o"
+definition
+  isofenv :: "[i,i] => o"  where
    "isofenv(ve,te) ==                
       ve_dom(ve) = te_dom(te) &            
       (\<forall>x \<in> ve_dom(ve).                          
@@ -64,10 +64,6 @@
 declare ElabRel.dom_subset [THEN subsetD, dest]
 
 end
-  
-
-
-
 
 
 
@@ -75,14 +71,3 @@
 
 
 
-
-
-
-
-
-
-
-
-
-
-
--- a/src/ZF/Coind/Values.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Coind/Values.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -33,8 +33,8 @@
 
 primrec "ve_app(ve_mk(m), a) = map_app(m,a)"
 
-constdefs
-  ve_emp :: i
+definition
+  ve_emp :: i  where
    "ve_emp == ve_mk(map_emp)"
 
 
--- a/src/ZF/Datatype.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Datatype.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -17,13 +17,13 @@
 structure Data_Arg =
   struct
   val intrs = 
-      [SigmaI, InlI, InrI,
-       Pair_in_univ, Inl_in_univ, Inr_in_univ, 
-       zero_in_univ, A_into_univ, nat_into_univ, UnCI];
+      [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
+       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, 
+       @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
 
 
-  val elims = [make_elim InlD, make_elim InrD,   (*for mutual recursion*)
-               SigmaE, sumE];                    (*allows * and + in spec*)
+  val elims = [make_elim @{thm InlD}, make_elim @{thm InrD},   (*for mutual recursion*)
+               @{thm SigmaE}, @{thm sumE}];                    (*allows * and + in spec*)
   end;
 
 
@@ -40,12 +40,12 @@
 structure CoData_Arg =
   struct
   val intrs = 
-      [QSigmaI, QInlI, QInrI,
-       QPair_in_quniv, QInl_in_quniv, QInr_in_quniv, 
-       zero_in_quniv, A_into_quniv, nat_into_quniv, UnCI];
+      [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
+       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, 
+       @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
 
-  val elims = [make_elim QInlD, make_elim QInrD,   (*for mutual recursion*)
-               QSigmaE, qsumE];                    (*allows * and + in spec*)
+  val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
+               @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
   end;
 
 structure CoData_Package = 
--- a/src/ZF/Epsilon.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Epsilon.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,17 +9,20 @@
 
 theory Epsilon imports Nat begin
 
-constdefs
-  eclose    :: "i=>i"
+definition
+  eclose    :: "i=>i"  where
     "eclose(A) == \<Union>n\<in>nat. nat_rec(n, A, %m r. Union(r))"
 
-  transrec  :: "[i, [i,i]=>i] =>i"
+definition
+  transrec  :: "[i, [i,i]=>i] =>i"  where
     "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)"
  
-  rank      :: "i=>i"
+definition
+  rank      :: "i=>i"  where
     "rank(a) == transrec(a, %x f. \<Union>y\<in>x. succ(f`y))"
 
-  transrec2 :: "[i, i, [i,i]=>i] =>i"
+definition
+  transrec2 :: "[i, i, [i,i]=>i] =>i"  where
     "transrec2(k, a, b) ==                     
        transrec(k, 
                 %i r. if(i=0, a, 
@@ -27,10 +30,12 @@
                            b(THE j. i=succ(j), r`(THE j. i=succ(j))),   
                            \<Union>j<i. r`j)))"
 
-  recursor  :: "[i, [i,i]=>i, i]=>i"
+definition
+  recursor  :: "[i, [i,i]=>i, i]=>i"  where
     "recursor(a,b,k) ==  transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))"
 
-  rec  :: "[i, i, [i,i]=>i]=>i"
+definition
+  rec  :: "[i, i, [i,i]=>i]=>i"  where
     "rec(k,a,b) == recursor(a,b,k)"
 
 
--- a/src/ZF/EquivClass.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/EquivClass.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,15 +9,16 @@
 
 theory EquivClass imports Trancl Perm begin
 
-constdefs
-
-  quotient   :: "[i,i]=>i"    (infixl "'/'/" 90)  (*set of equiv classes*)
+definition
+  quotient   :: "[i,i]=>i"    (infixl "'/'/" 90)  (*set of equiv classes*)  where
       "A//r == {r``{x} . x:A}"
 
-  congruent  :: "[i,i=>i]=>o"
+definition
+  congruent  :: "[i,i=>i]=>o"  where
       "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)"
 
-  congruent2 :: "[i,i,[i,i]=>i]=>o"
+definition
+  congruent2 :: "[i,i,[i,i]=>i]=>o"  where
       "congruent2(r1,r2,b) == ALL y1 z1 y2 z2.
            <y1,z1>:r1 --> <y2,z2>:r2 --> b(y1,y2) = b(z1,z2)"
 
@@ -231,36 +232,4 @@
 apply (simp add: congruent_def) 
 done
 
-ML
-{*
-val sym_trans_comp_subset = thm "sym_trans_comp_subset";
-val refl_comp_subset = thm "refl_comp_subset";
-val equiv_comp_eq = thm "equiv_comp_eq";
-val comp_equivI = thm "comp_equivI";
-val equiv_class_subset = thm "equiv_class_subset";
-val equiv_class_eq = thm "equiv_class_eq";
-val equiv_class_self = thm "equiv_class_self";
-val subset_equiv_class = thm "subset_equiv_class";
-val eq_equiv_class = thm "eq_equiv_class";
-val equiv_class_nondisjoint = thm "equiv_class_nondisjoint";
-val equiv_type = thm "equiv_type";
-val equiv_class_eq_iff = thm "equiv_class_eq_iff";
-val eq_equiv_class_iff = thm "eq_equiv_class_iff";
-val quotientI = thm "quotientI";
-val quotientE = thm "quotientE";
-val Union_quotient = thm "Union_quotient";
-val quotient_disj = thm "quotient_disj";
-val UN_equiv_class = thm "UN_equiv_class";
-val UN_equiv_class_type = thm "UN_equiv_class_type";
-val UN_equiv_class_inject = thm "UN_equiv_class_inject";
-val congruent2_implies_congruent = thm "congruent2_implies_congruent";
-val congruent2_implies_congruent_UN = thm "congruent2_implies_congruent_UN";
-val congruent_commuteI = thm "congruent_commuteI";
-val UN_equiv_class2 = thm "UN_equiv_class2";
-val UN_equiv_class_type2 = thm "UN_equiv_class_type2";
-val congruent2I = thm "congruent2I";
-val congruent2_commuteI = thm "congruent2_commuteI";
-val congruent_commuteI = thm "congruent_commuteI";
-*}
-
 end
--- a/src/ZF/Finite.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Finite.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -200,37 +200,11 @@
 
 subsection{*The Contents of a Singleton Set*}
 
-constdefs
-  contents :: "i=>i"
+definition
+  contents :: "i=>i"  where
    "contents(X) == THE x. X = {x}"
 
 lemma contents_eq [simp]: "contents ({x}) = x"
 by (simp add: contents_def)
 
-
-ML
-{*
-val Fin_intros = thms "Fin.intros";
-
-val Fin_mono = thm "Fin_mono";
-val FinD = thm "FinD";
-val Fin_induct = thm "Fin_induct";
-val Fin_UnI = thm "Fin_UnI";
-val Fin_UnionI = thm "Fin_UnionI";
-val Fin_subset = thm "Fin_subset";
-val Fin_IntI1 = thm "Fin_IntI1";
-val Fin_IntI2 = thm "Fin_IntI2";
-val Fin_0_induct = thm "Fin_0_induct";
-val nat_fun_subset_Fin = thm "nat_fun_subset_Fin";
-val FiniteFun_mono = thm "FiniteFun_mono";
-val FiniteFun_mono1 = thm "FiniteFun_mono1";
-val FiniteFun_is_fun = thm "FiniteFun_is_fun";
-val FiniteFun_domain_Fin = thm "FiniteFun_domain_Fin";
-val FiniteFun_apply_type = thm "FiniteFun_apply_type";
-val FiniteFun_subset = thm "FiniteFun_subset";
-val fun_FiniteFunI = thm "fun_FiniteFunI";
-val lam_FiniteFun = thm "lam_FiniteFun";
-val FiniteFun_Collect_iff = thm "FiniteFun_Collect_iff";
-*}
-
 end
--- a/src/ZF/Fixedpt.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Fixedpt.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -8,16 +8,17 @@
 
 theory Fixedpt imports equalities begin
 
-constdefs
-  
+definition 
   (*monotone operator from Pow(D) to itself*)
-  bnd_mono :: "[i,i=>i]=>o"
+  bnd_mono :: "[i,i=>i]=>o"  where
      "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))"
 
-  lfp      :: "[i,i=>i]=>i"
+definition 
+  lfp      :: "[i,i=>i]=>i"  where
      "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})"
 
-  gfp      :: "[i,i=>i]=>i"
+definition 
+  gfp      :: "[i,i=>i]=>i"  where
      "gfp(D,h) == Union({X: Pow(D). X <= h(X)})"
 
 text{*The theorem is proved in the lattice of subsets of @{term D}, 
@@ -301,42 +302,4 @@
 apply (blast del: subsetI intro: subset_trans gfp_subset) 
 done
 
-ML
-{*
-val bnd_mono_def = thm "bnd_mono_def";
-val lfp_def = thm "lfp_def";
-val gfp_def = thm "gfp_def";
-
-val bnd_monoI = thm "bnd_monoI";
-val bnd_monoD1 = thm "bnd_monoD1";
-val bnd_monoD2 = thm "bnd_monoD2";
-val bnd_mono_subset = thm "bnd_mono_subset";
-val bnd_mono_Un = thm "bnd_mono_Un";
-val bnd_mono_Int = thm "bnd_mono_Int";
-val lfp_lowerbound = thm "lfp_lowerbound";
-val lfp_subset = thm "lfp_subset";
-val def_lfp_subset = thm "def_lfp_subset";
-val lfp_greatest = thm "lfp_greatest";
-val lfp_unfold = thm "lfp_unfold";
-val def_lfp_unfold = thm "def_lfp_unfold";
-val Collect_is_pre_fixedpt = thm "Collect_is_pre_fixedpt";
-val induct = thm "induct";
-val def_induct = thm "def_induct";
-val lfp_Int_lowerbound = thm "lfp_Int_lowerbound";
-val lfp_mono = thm "lfp_mono";
-val lfp_mono2 = thm "lfp_mono2";
-val gfp_upperbound = thm "gfp_upperbound";
-val gfp_subset = thm "gfp_subset";
-val def_gfp_subset = thm "def_gfp_subset";
-val gfp_least = thm "gfp_least";
-val gfp_unfold = thm "gfp_unfold";
-val def_gfp_unfold = thm "def_gfp_unfold";
-val weak_coinduct = thm "weak_coinduct";
-val coinduct = thm "coinduct";
-val def_coinduct = thm "def_coinduct";
-val def_Collect_coinduct = thm "def_Collect_coinduct";
-val gfp_mono = thm "gfp_mono";
-*}
-
-
 end
--- a/src/ZF/Induct/Binary_Trees.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Binary_Trees.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -88,8 +88,8 @@
   apply (atomize, simp)
   done
 
-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/src/ZF/Induct/FoldSet.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/FoldSet.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -20,12 +20,12 @@
 		==>  <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
   type_intros Fin.intros
   
-constdefs
-  
-  fold :: "[i, [i,i]=>i, i, i] => i"  ("fold[_]'(_,_,_')")
+definition
+  fold :: "[i, [i,i]=>i, i, i] => i"  ("fold[_]'(_,_,_')")  where
    "fold[B](f,e, A) == THE x. <A, x>\<in>fold_set(A, B, f,e)"
 
-   setsum :: "[i=>i, i] => i"
+definition
+   setsum :: "[i=>i, i] => i"  where
    "setsum(g, C) == if Finite(C) then
                      fold[int](%x y. g(x) $+ y, #0, C) else #0"
 
@@ -395,32 +395,4 @@
 apply (simp_all add: Diff_cons_eq Finite_Diff) 
 done
 
-ML
-{*
-val fold_set_mono = thm "fold_set_mono";
-val Diff1_fold_set = thm "Diff1_fold_set";
-val Diff_sing_imp = thm "Diff_sing_imp";
-val fold_0 = thm "fold_0";
-val setsum_0 = thm "setsum_0";
-val setsum_cons = thm "setsum_cons";
-val setsum_K0 = thm "setsum_K0";
-val setsum_Un_Int = thm "setsum_Un_Int";
-val setsum_type = thm "setsum_type";
-val setsum_Un_disjoint = thm "setsum_Un_disjoint";
-val Finite_RepFun = thm "Finite_RepFun";
-val setsum_UN_disjoint = thm "setsum_UN_disjoint";
-val setsum_addf = thm "setsum_addf";
-val fold_set_cong = thm "fold_set_cong";
-val fold_cong = thm "fold_cong";
-val setsum_cong = thm "setsum_cong";
-val setsum_Un = thm "setsum_Un";
-val setsum_zneg_or_0 = thm "setsum_zneg_or_0";
-val setsum_succD = thm "setsum_succD";
-val g_zpos_imp_setsum_zpos = thm "g_zpos_imp_setsum_zpos";
-val g_zpos_imp_setsum_zpos2 = thm "g_zpos_imp_setsum_zpos2";
-val g_zspos_imp_setsum_zspos = thm "g_zspos_imp_setsum_zspos";
-val setsum_Diff = thm "setsum_Diff";
-*}
-
-
 end
\ No newline at end of file
--- a/src/ZF/Induct/Multiset.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Multiset.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -17,50 +17,57 @@
   Mult :: "i=>i" where
   "Mult(A) == A -||> nat-{0}"
 
-constdefs
-
+definition
   (* This is the original "restrict" from ZF.thy.
      Restricts the function f to the domain A
      FIXME: adapt Multiset to the new "restrict". *)
-
-  funrestrict :: "[i,i] => i"
+  funrestrict :: "[i,i] => i"  where
   "funrestrict(f,A) == \<lambda>x \<in> A. f`x"
 
+definition
   (* M is a multiset *)
-  multiset :: "i => o"
+  multiset :: "i => o"  where
   "multiset(M) == \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
 
-  mset_of :: "i=>i"
+definition
+  mset_of :: "i=>i"  where
   "mset_of(M) == domain(M)"
 
-  munion    :: "[i, i] => i" (infixl "+#" 65)
+definition
+  munion    :: "[i, i] => i" (infixl "+#" 65)  where
   "M +# N == \<lambda>x \<in> mset_of(M) Un mset_of(N).
      if x \<in> mset_of(M) Int mset_of(N) then  (M`x) #+ (N`x)
      else (if x \<in> mset_of(M) then M`x else N`x)"
 
+definition
   (*convert a function to a multiset by eliminating 0*)
-  normalize :: "i => i"
+  normalize :: "i => i"  where
   "normalize(f) ==
        if (\<exists>A. f \<in> A -> nat & Finite(A)) then
             funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
        else 0"
 
-  mdiff  :: "[i, i] => i" (infixl "-#" 65)
+definition
+  mdiff  :: "[i, i] => i" (infixl "-#" 65)  where
   "M -# N ==  normalize(\<lambda>x \<in> mset_of(M).
 			if x \<in> mset_of(N) then M`x #- N`x else M`x)"
 
+definition
   (* set of elements of a multiset *)
-  msingle :: "i => i"    ("{#_#}")
+  msingle :: "i => i"    ("{#_#}")  where
   "{#a#} == {<a, 1>}"
 
-  MCollect :: "[i, i=>o] => i"  (*comprehension*)
+definition
+  MCollect :: "[i, i=>o] => i"  (*comprehension*)  where
   "MCollect(M, P) == funrestrict(M, {x \<in> mset_of(M). P(x)})"
 
+definition
   (* Counts the number of occurences of an element in a multiset *)
-  mcount :: "[i, i] => i"
+  mcount :: "[i, i] => i"  where
   "mcount(M, a) == if a \<in> mset_of(M) then  M`a else 0"
 
-  msize :: "i => i"
+definition
+  msize :: "i => i"  where
   "msize(M) == setsum(%a. $# mcount(M,a), mset_of(M))"
 
 abbreviation
@@ -72,32 +79,35 @@
 syntax (xsymbols)
   "@MColl" :: "[pttrn, i, o] => i" ("(1{# _ \<in> _./ _#})")
 translations
-  "{#x \<in> M. P#}" == "MCollect(M, %x. P)"
+  "{#x \<in> M. P#}" == "CONST MCollect(M, %x. P)"
 
   (* multiset orderings *)
 
-constdefs
+definition
    (* multirel1 has to be a set (not a predicate) so that we can form
       its transitive closure and reason about wf(.) and acc(.) *)
-
-  multirel1 :: "[i,i]=>i"
+  multirel1 :: "[i,i]=>i"  where
   "multirel1(A, r) ==
      {<M, N> \<in> Mult(A)*Mult(A).
       \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).
       N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
 
-  multirel :: "[i, i] => i"
+definition
+  multirel :: "[i, i] => i"  where
   "multirel(A, r) == multirel1(A, r)^+" 		
 
   (* ordinal multiset orderings *)
 
-  omultiset :: "i => o"
+definition
+  omultiset :: "i => o"  where
   "omultiset(M) == \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"
 
-  mless :: "[i, i] => o" (infixl "<#" 50)
+definition
+  mless :: "[i, i] => o" (infixl "<#" 50)  where
   "M <# N ==  \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
 
-  mle  :: "[i, i] => o"  (infixl "<#=" 50)
+definition
+  mle  :: "[i, i] => o"  (infixl "<#=" 50)  where
   "M <#= N == (omultiset(M) & M = N) | M <# N"
 
 
@@ -1283,147 +1293,4 @@
 apply (rule_tac [2] mle_mono, auto)
 done
 
-ML
-{*
-val munion_ac = thms "munion_ac";
-val funrestrict_subset = thm "funrestrict_subset";
-val funrestrict_type = thm "funrestrict_type";
-val funrestrict_type2 = thm "funrestrict_type2";
-val funrestrict = thm "funrestrict";
-val funrestrict_empty = thm "funrestrict_empty";
-val domain_funrestrict = thm "domain_funrestrict";
-val fun_cons_funrestrict_eq = thm "fun_cons_funrestrict_eq";
-val multiset_fun_iff = thm "multiset_fun_iff";
-val multiset_into_Mult = thm "multiset_into_Mult";
-val Mult_into_multiset = thm "Mult_into_multiset";
-val Mult_iff_multiset = thm "Mult_iff_multiset";
-val multiset_iff_Mult_mset_of = thm "multiset_iff_Mult_mset_of";
-val multiset_0 = thm "multiset_0";
-val multiset_set_of_Finite = thm "multiset_set_of_Finite";
-val mset_of_0 = thm "mset_of_0";
-val mset_is_0_iff = thm "mset_is_0_iff";
-val mset_of_single = thm "mset_of_single";
-val mset_of_union = thm "mset_of_union";
-val mset_of_diff = thm "mset_of_diff";
-val msingle_not_0 = thm "msingle_not_0";
-val msingle_eq_iff = thm "msingle_eq_iff";
-val msingle_multiset = thm "msingle_multiset";
-val Collect_Finite = thms "Collect_Finite";
-val normalize_idem = thm "normalize_idem";
-val normalize_multiset = thm "normalize_multiset";
-val multiset_normalize = thm "multiset_normalize";
-val munion_multiset = thm "munion_multiset";
-val mdiff_multiset = thm "mdiff_multiset";
-val munion_0 = thm "munion_0";
-val munion_commute = thm "munion_commute";
-val munion_assoc = thm "munion_assoc";
-val munion_lcommute = thm "munion_lcommute";
-val mdiff_self_eq_0 = thm "mdiff_self_eq_0";
-val mdiff_0 = thm "mdiff_0";
-val mdiff_0_right = thm "mdiff_0_right";
-val mdiff_union_inverse2 = thm "mdiff_union_inverse2";
-val mcount_type = thm "mcount_type";
-val mcount_0 = thm "mcount_0";
-val mcount_single = thm "mcount_single";
-val mcount_union = thm "mcount_union";
-val mcount_diff = thm "mcount_diff";
-val mcount_elem = thm "mcount_elem";
-val msize_0 = thm "msize_0";
-val msize_single = thm "msize_single";
-val msize_type = thm "msize_type";
-val msize_zpositive = thm "msize_zpositive";
-val msize_int_of_nat = thm "msize_int_of_nat";
-val not_empty_multiset_imp_exist = thm "not_empty_multiset_imp_exist";
-val msize_eq_0_iff = thm "msize_eq_0_iff";
-val setsum_mcount_Int = thm "setsum_mcount_Int";
-val msize_union = thm "msize_union";
-val msize_eq_succ_imp_elem = thm "msize_eq_succ_imp_elem";
-val multiset_equality = thm "multiset_equality";
-val munion_eq_0_iff = thm "munion_eq_0_iff";
-val empty_eq_munion_iff = thm "empty_eq_munion_iff";
-val munion_right_cancel = thm "munion_right_cancel";
-val munion_left_cancel = thm "munion_left_cancel";
-val nat_add_eq_1_cases = thm "nat_add_eq_1_cases";
-val munion_is_single = thm "munion_is_single";
-val msingle_is_union = thm "msingle_is_union";
-val setsum_decr = thm "setsum_decr";
-val setsum_decr2 = thm "setsum_decr2";
-val setsum_decr3 = thm "setsum_decr3";
-val nat_le_1_cases = thm "nat_le_1_cases";
-val succ_pred_eq_self = thm "succ_pred_eq_self";
-val multiset_funrestict = thm "multiset_funrestict";
-val multiset_induct_aux = thm "multiset_induct_aux";
-val multiset_induct2 = thm "multiset_induct2";
-val munion_single_case1 = thm "munion_single_case1";
-val munion_single_case2 = thm "munion_single_case2";
-val multiset_induct = thm "multiset_induct";
-val MCollect_multiset = thm "MCollect_multiset";
-val mset_of_MCollect = thm "mset_of_MCollect";
-val MCollect_mem_iff = thm "MCollect_mem_iff";
-val mcount_MCollect = thm "mcount_MCollect";
-val multiset_partition = thm "multiset_partition";
-val natify_elem_is_self = thm "natify_elem_is_self";
-val munion_eq_conv_diff = thm "munion_eq_conv_diff";
-val melem_diff_single = thm "melem_diff_single";
-val munion_eq_conv_exist = thm "munion_eq_conv_exist";
-val multirel1_type = thm "multirel1_type";
-val multirel1_0 = thm "multirel1_0";
-val multirel1_iff = thm "multirel1_iff";
-val multirel1_mono1 = thm "multirel1_mono1";
-val multirel1_mono2 = thm "multirel1_mono2";
-val multirel1_mono = thm "multirel1_mono";
-val not_less_0 = thm "not_less_0";
-val less_munion = thm "less_munion";
-val multirel1_base = thm "multirel1_base";
-val acc_0 = thm "acc_0";
-val all_accessible = thm "all_accessible";
-val wf_on_multirel1 = thm "wf_on_multirel1";
-val wf_multirel1 = thm "wf_multirel1";
-val multirel_type = thm "multirel_type";
-val multirel_mono = thm "multirel_mono";
-val add_diff_eq = thm "add_diff_eq";
-val mdiff_union_single_conv = thm "mdiff_union_single_conv";
-val diff_add_commute = thm "diff_add_commute";
-val multirel_implies_one_step = thm "multirel_implies_one_step";
-val melem_imp_eq_diff_union = thm "melem_imp_eq_diff_union";
-val msize_eq_succ_imp_eq_union = thm "msize_eq_succ_imp_eq_union";
-val one_step_implies_multirel = thm "one_step_implies_multirel";
-val irrefl_on_multirel = thm "irrefl_on_multirel";
-val trans_on_multirel = thm "trans_on_multirel";
-val multirel_trans = thm "multirel_trans";
-val trans_multirel = thm "trans_multirel";
-val part_ord_multirel = thm "part_ord_multirel";
-val munion_multirel1_mono = thm "munion_multirel1_mono";
-val munion_multirel_mono2 = thm "munion_multirel_mono2";
-val munion_multirel_mono1 = thm "munion_multirel_mono1";
-val munion_multirel_mono = thm "munion_multirel_mono";
-val field_Memrel_mono = thms "field_Memrel_mono";
-val multirel_Memrel_mono = thms "multirel_Memrel_mono";
-val omultiset_is_multiset = thm "omultiset_is_multiset";
-val munion_omultiset = thm "munion_omultiset";
-val mdiff_omultiset = thm "mdiff_omultiset";
-val irrefl_Memrel = thm "irrefl_Memrel";
-val trans_iff_trans_on = thm "trans_iff_trans_on";
-val part_ord_Memrel = thm "part_ord_Memrel";
-val part_ord_mless = thms "part_ord_mless";
-val mless_not_refl = thm "mless_not_refl";
-val mless_irrefl = thms "mless_irrefl";
-val mless_trans = thm "mless_trans";
-val mless_not_sym = thm "mless_not_sym";
-val mless_asym = thm "mless_asym";
-val mle_refl = thm "mle_refl";
-val mle_antisym = thm "mle_antisym";
-val mle_trans = thm "mle_trans";
-val mless_le_iff = thm "mless_le_iff";
-val munion_less_mono2 = thm "munion_less_mono2";
-val munion_less_mono1 = thm "munion_less_mono1";
-val mless_imp_omultiset = thm "mless_imp_omultiset";
-val munion_less_mono = thm "munion_less_mono";
-val mle_imp_omultiset = thm "mle_imp_omultiset";
-val mle_mono = thm "mle_mono";
-val omultiset_0 = thm "omultiset_0";
-val empty_leI = thm "empty_leI";
-val munion_upper1 = thm "munion_upper1";
-*}
-
 end
--- a/src/ZF/Induct/Mutil.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Mutil.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -32,8 +32,8 @@
   type_intros empty_subsetI Union_upper Un_least PowI
   type_elims PowD [elim_format]
 
-constdefs
-  evnodd :: "[i, i] => i"
+definition
+  evnodd :: "[i, i] => i"  where
   "evnodd(A,b) == {z \<in> A. \<exists>i j. z = <i,j> \<and> (i #+ j) mod 2 = b}"
 
 
--- a/src/ZF/Induct/Ntree.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Ntree.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -31,13 +31,13 @@
   monos FiniteFun_mono [OF subset_refl]
   type_intros FiniteFun_in_univ'
 
-constdefs
-  ntree_rec :: "[[i, i, i] => i, i] => i"
+definition
+  ntree_rec :: "[[i, i, i] => i, i] => i"  where
   "ntree_rec(b) ==
     Vrecursor(\<lambda>pr. ntree_case(\<lambda>x h. b(x, h, \<lambda>i \<in> domain(h). pr`(h`i))))"
 
-constdefs
-  ntree_copy :: "i => i"
+definition
+  ntree_copy :: "i => i"  where
   "ntree_copy(z) == ntree_rec(\<lambda>x h r. Branch(x,r), z)"
 
 
--- a/src/ZF/Induct/Primrec.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Primrec.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -17,20 +17,24 @@
 
 subsection {* Basic definitions *}
 
-constdefs
-  SC :: "i"
+definition
+  SC :: "i"  where
   "SC == \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. succ(x), l)"
 
-  CONSTANT :: "i=>i"
+definition
+  CONSTANT :: "i=>i"  where
   "CONSTANT(k) == \<lambda>l \<in> list(nat). k"
 
-  PROJ :: "i=>i"
+definition
+  PROJ :: "i=>i"  where
   "PROJ(i) == \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. x, drop(i,l))"
 
-  COMP :: "[i,i]=>i"
+definition
+  COMP :: "[i,i]=>i"  where
   "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` List.map(\<lambda>f. f`l, fs)"
 
-  PREC :: "[i,i]=>i"
+definition
+  PREC :: "[i,i]=>i"  where
   "PREC(f,g) ==
      \<lambda>l \<in> list(nat). list_case(0,
                       \<lambda>x xs. rec(x, f`xs, \<lambda>y r. g ` Cons(r, Cons(y, xs))), l)"
--- a/src/ZF/Induct/PropLog.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/PropLog.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -62,8 +62,8 @@
   "is_true_fun(Var(v), t) = (if v \<in> t then 1 else 0)"
   "is_true_fun(p=>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)"
 
-constdefs
-  is_true :: "[i,i] => o"
+definition
+  is_true :: "[i,i] => o"  where
   "is_true(p,t) == is_true_fun(p,t) = 1"
   -- {* this definition is required since predicates can't be recursive *}
 
@@ -84,8 +84,8 @@
   is @{text p}.
 *}
 
-constdefs
-  logcon :: "[i,i] => o"    (infixl "|=" 50)
+definition
+  logcon :: "[i,i] => o"    (infixl "|=" 50)  where
   "H |= p == \<forall>t. (\<forall>q \<in> H. is_true(q,t)) --> is_true(p,t)"
 
 
--- a/src/ZF/Induct/Term.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Induct/Term.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -22,24 +22,29 @@
 
 declare Apply [TC]
 
-constdefs
-  term_rec :: "[i, [i, i, i] => i] => i"
+definition
+  term_rec :: "[i, [i, i, i] => i] => i"  where
   "term_rec(t,d) ==
     Vrec(t, \<lambda>t g. term_case(\<lambda>x zs. d(x, zs, map(\<lambda>z. g`z, zs)), t))"
 
-  term_map :: "[i => i, i] => i"
+definition
+  term_map :: "[i => i, i] => i"  where
   "term_map(f,t) == term_rec(t, \<lambda>x zs rs. Apply(f(x), rs))"
 
-  term_size :: "i => i"
+definition
+  term_size :: "i => i"  where
   "term_size(t) == term_rec(t, \<lambda>x zs rs. succ(list_add(rs)))"
 
-  reflect :: "i => i"
+definition
+  reflect :: "i => i"  where
   "reflect(t) == term_rec(t, \<lambda>x zs rs. Apply(x, rev(rs)))"
 
-  preorder :: "i => i"
+definition
+  preorder :: "i => i"  where
   "preorder(t) == term_rec(t, \<lambda>x zs rs. Cons(x, flat(rs)))"
 
-  postorder :: "i => i"
+definition
+  postorder :: "i => i"  where
   "postorder(t) == term_rec(t, \<lambda>x zs rs. flat(rs) @ [x])"
 
 lemma term_unfold: "term(A) = A * list(term(A))"
--- a/src/ZF/Inductive.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Inductive.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -32,10 +32,10 @@
   struct
   val oper      = Const("Fixedpt.lfp", [iT,iT-->iT]--->iT)
   val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
-  val bnd_monoI = bnd_monoI
-  val subs      = def_lfp_subset
-  val Tarski    = def_lfp_unfold
-  val induct    = def_induct
+  val bnd_monoI = @{thm bnd_monoI}
+  val subs      = @{thm def_lfp_subset}
+  val Tarski    = @{thm def_lfp_unfold}
+  val induct    = @{thm def_induct}
   end;
 
 structure Standard_Prod =
@@ -43,11 +43,11 @@
   val sigma     = Const("Sigma", [iT, iT-->iT]--->iT)
   val pair      = Const("Pair", [iT,iT]--->iT)
   val split_name = "split"
-  val pair_iff  = Pair_iff
-  val split_eq  = split
-  val fsplitI   = splitI
-  val fsplitD   = splitD
-  val fsplitE   = splitE
+  val pair_iff  = @{thm Pair_iff}
+  val split_eq  = @{thm split}
+  val fsplitI   = @{thm splitI}
+  val fsplitD   = @{thm splitD}
+  val fsplitE   = @{thm splitE}
   end;
 
 structure Standard_CP = CartProd_Fun (Standard_Prod);
@@ -58,12 +58,12 @@
   val inl       = Const("Inl", iT-->iT)
   val inr       = Const("Inr", iT-->iT)
   val elim      = Const("case", [iT-->iT, iT-->iT, iT]--->iT)
-  val case_inl  = case_Inl
-  val case_inr  = case_Inr
-  val inl_iff   = Inl_iff
-  val inr_iff   = Inr_iff
-  val distinct  = Inl_Inr_iff
-  val distinct' = Inr_Inl_iff
+  val case_inl  = @{thm case_Inl}
+  val case_inr  = @{thm case_Inr}
+  val inl_iff   = @{thm Inl_iff}
+  val inr_iff   = @{thm Inr_iff}
+  val distinct  = @{thm Inl_Inr_iff}
+  val distinct' = @{thm Inr_Inl_iff}
   val free_SEs  = Ind_Syntax.mk_free_SEs
             [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff]
   end;
@@ -79,10 +79,10 @@
   struct
   val oper      = Const("Fixedpt.gfp", [iT,iT-->iT]--->iT)
   val bnd_mono  = Const("Fixedpt.bnd_mono", [iT,iT-->iT]--->oT)
-  val bnd_monoI = bnd_monoI
-  val subs      = def_gfp_subset
-  val Tarski    = def_gfp_unfold
-  val induct    = def_Collect_coinduct
+  val bnd_monoI = @{thm bnd_monoI}
+  val subs      = @{thm def_gfp_subset}
+  val Tarski    = @{thm def_gfp_unfold}
+  val induct    = @{thm def_Collect_coinduct}
   end;
 
 structure Quine_Prod =
@@ -90,11 +90,11 @@
   val sigma     = Const("QPair.QSigma", [iT, iT-->iT]--->iT)
   val pair      = Const("QPair.QPair", [iT,iT]--->iT)
   val split_name = "QPair.qsplit"
-  val pair_iff  = QPair_iff
-  val split_eq  = qsplit
-  val fsplitI   = qsplitI
-  val fsplitD   = qsplitD
-  val fsplitE   = qsplitE
+  val pair_iff  = @{thm QPair_iff}
+  val split_eq  = @{thm qsplit}
+  val fsplitI   = @{thm qsplitI}
+  val fsplitD   = @{thm qsplitD}
+  val fsplitE   = @{thm qsplitE}
   end;
 
 structure Quine_CP = CartProd_Fun (Quine_Prod);
@@ -105,12 +105,12 @@
   val inl       = Const("QPair.QInl", iT-->iT)
   val inr       = Const("QPair.QInr", iT-->iT)
   val elim      = Const("QPair.qcase", [iT-->iT, iT-->iT, iT]--->iT)
-  val case_inl  = qcase_QInl
-  val case_inr  = qcase_QInr
-  val inl_iff   = QInl_iff
-  val inr_iff   = QInr_iff
-  val distinct  = QInl_QInr_iff
-  val distinct' = QInr_QInl_iff
+  val case_inl  = @{thm qcase_QInl}
+  val case_inr  = @{thm qcase_QInr}
+  val inl_iff   = @{thm QInl_iff}
+  val inr_iff   = @{thm QInr_iff}
+  val distinct  = @{thm QInl_QInr_iff}
+  val distinct' = @{thm QInr_QInl_iff}
   val free_SEs  = Ind_Syntax.mk_free_SEs
             [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff]
   end;
--- a/src/ZF/Int.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Int.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,45 +9,56 @@
 
 theory Int imports EquivClass ArithSimp begin
 
-constdefs
-  intrel :: i
+definition
+  intrel :: i  where
     "intrel == {p : (nat*nat)*(nat*nat).                 
                 \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
 
-  int :: i
+definition
+  int :: i  where
     "int == (nat*nat)//intrel"  
 
-  int_of :: "i=>i" --{*coercion from nat to int*}    ("$# _" [80] 80)
+definition
+  int_of :: "i=>i" --{*coercion from nat to int*}    ("$# _" [80] 80)  where
     "$# m == intrel `` {<natify(m), 0>}"
 
-  intify :: "i=>i" --{*coercion from ANYTHING to int*}
+definition
+  intify :: "i=>i" --{*coercion from ANYTHING to int*}  where
     "intify(m) == if m : int then m else $#0"
 
-  raw_zminus :: "i=>i"
+definition
+  raw_zminus :: "i=>i"  where
     "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}"
 
-  zminus :: "i=>i"                                 ("$- _" [80] 80)
+definition
+  zminus :: "i=>i"                                 ("$- _" [80] 80)  where
     "$- z == raw_zminus (intify(z))"
 
-  znegative   ::      "i=>o"
+definition
+  znegative   ::      "i=>o"  where
     "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
 
-  iszero      ::      "i=>o"
+definition
+  iszero      ::      "i=>o"  where
     "iszero(z) == z = $# 0"
     
-  raw_nat_of  :: "i=>i"
+definition
+  raw_nat_of  :: "i=>i"  where
   "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)"
 
-  nat_of  :: "i=>i"
+definition
+  nat_of  :: "i=>i"  where
   "nat_of(z) == raw_nat_of (intify(z))"
 
-  zmagnitude  ::      "i=>i"
+definition
+  zmagnitude  ::      "i=>i"  where
   --{*could be replaced by an absolute value function from int to int?*}
     "zmagnitude(z) ==
      THE m. m\<in>nat & ((~ znegative(z) & z = $# m) |
 		       (znegative(z) & $- z = $# m))"
 
-  raw_zmult   ::      "[i,i]=>i"
+definition
+  raw_zmult   ::      "[i,i]=>i"  where
     (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
       Perhaps a "curried" or even polymorphic congruent predicate would be
       better.*)
@@ -55,34 +66,40 @@
        \<Union>p1\<in>z1. \<Union>p2\<in>z2.  split(%x1 y1. split(%x2 y2.        
                    intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
 
-  zmult       ::      "[i,i]=>i"      (infixl "$*" 70)
+definition
+  zmult       ::      "[i,i]=>i"      (infixl "$*" 70)  where
      "z1 $* z2 == raw_zmult (intify(z1),intify(z2))"
 
-  raw_zadd    ::      "[i,i]=>i"
+definition
+  raw_zadd    ::      "[i,i]=>i"  where
      "raw_zadd (z1, z2) == 
        \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2                 
                            in intrel``{<x1#+x2, y1#+y2>}"
 
-  zadd        ::      "[i,i]=>i"      (infixl "$+" 65)
+definition
+  zadd        ::      "[i,i]=>i"      (infixl "$+" 65)  where
      "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))"
 
-  zdiff        ::      "[i,i]=>i"      (infixl "$-" 65)
+definition
+  zdiff        ::      "[i,i]=>i"      (infixl "$-" 65)  where
      "z1 $- z2 == z1 $+ zminus(z2)"
 
-  zless        ::      "[i,i]=>o"      (infixl "$<" 50)
+definition
+  zless        ::      "[i,i]=>o"      (infixl "$<" 50)  where
      "z1 $< z2 == znegative(z1 $- z2)"
   
-  zle          ::      "[i,i]=>o"      (infixl "$<=" 50)
+definition
+  zle          ::      "[i,i]=>o"      (infixl "$<=" 50)  where
      "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)"
   
 
-syntax (xsymbols)
-  zmult :: "[i,i]=>i"          (infixl "$\<times>" 70)
-  zle   :: "[i,i]=>o"          (infixl "$\<le>" 50)  --{*less than or equals*}
+notation (xsymbols)
+  zmult  (infixl "$\<times>" 70) and
+  zle  (infixl "$\<le>" 50)  --{*less than or equals*}
 
-syntax (HTML output)
-  zmult :: "[i,i]=>i"          (infixl "$\<times>" 70)
-  zle   :: "[i,i]=>o"          (infixl "$\<le>" 50)
+notation (HTML output)
+  zmult  (infixl "$\<times>" 70) and
+  zle  (infixl "$\<le>" 50)
 
 
 declare quotientE [elim!]
@@ -911,147 +928,4 @@
 lemma zminus_zle: "($- x $<= y) <-> ($- y $<= x)"
 by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus)
 
-ML
-{*
-val zdiff_def = thm "zdiff_def";
-val int_of_type = thm "int_of_type";
-val int_of_eq = thm "int_of_eq";
-val int_of_inject = thm "int_of_inject";
-val intify_in_int = thm "intify_in_int";
-val intify_ident = thm "intify_ident";
-val intify_idem = thm "intify_idem";
-val int_of_natify = thm "int_of_natify";
-val zminus_intify = thm "zminus_intify";
-val zadd_intify1 = thm "zadd_intify1";
-val zadd_intify2 = thm "zadd_intify2";
-val zdiff_intify1 = thm "zdiff_intify1";
-val zdiff_intify2 = thm "zdiff_intify2";
-val zmult_intify1 = thm "zmult_intify1";
-val zmult_intify2 = thm "zmult_intify2";
-val zless_intify1 = thm "zless_intify1";
-val zless_intify2 = thm "zless_intify2";
-val zle_intify1 = thm "zle_intify1";
-val zle_intify2 = thm "zle_intify2";
-val zminus_congruent = thm "zminus_congruent";
-val zminus_type = thm "zminus_type";
-val zminus_inject_intify = thm "zminus_inject_intify";
-val zminus_inject = thm "zminus_inject";
-val zminus = thm "zminus";
-val zminus_zminus_intify = thm "zminus_zminus_intify";
-val zminus_int0 = thm "zminus_int0";
-val zminus_zminus = thm "zminus_zminus";
-val not_znegative_int_of = thm "not_znegative_int_of";
-val znegative_zminus_int_of = thm "znegative_zminus_int_of";
-val not_znegative_imp_zero = thm "not_znegative_imp_zero";
-val nat_of_intify = thm "nat_of_intify";
-val nat_of_int_of = thm "nat_of_int_of";
-val nat_of_type = thm "nat_of_type";
-val zmagnitude_int_of = thm "zmagnitude_int_of";
-val natify_int_of_eq = thm "natify_int_of_eq";
-val zmagnitude_zminus_int_of = thm "zmagnitude_zminus_int_of";
-val zmagnitude_type = thm "zmagnitude_type";
-val not_zneg_int_of = thm "not_zneg_int_of";
-val not_zneg_mag = thm "not_zneg_mag";
-val zneg_int_of = thm "zneg_int_of";
-val zneg_mag = thm "zneg_mag";
-val int_cases = thm "int_cases";
-val not_zneg_nat_of_intify = thm "not_zneg_nat_of_intify";
-val not_zneg_nat_of = thm "not_zneg_nat_of";
-val zneg_nat_of = thm "zneg_nat_of";
-val zadd_congruent2 = thm "zadd_congruent2";
-val zadd_type = thm "zadd_type";
-val zadd = thm "zadd";
-val zadd_int0_intify = thm "zadd_int0_intify";
-val zadd_int0 = thm "zadd_int0";
-val zminus_zadd_distrib = thm "zminus_zadd_distrib";
-val zadd_commute = thm "zadd_commute";
-val zadd_assoc = thm "zadd_assoc";
-val zadd_left_commute = thm "zadd_left_commute";
-val zadd_ac = thms "zadd_ac";
-val int_of_add = thm "int_of_add";
-val int_succ_int_1 = thm "int_succ_int_1";
-val int_of_diff = thm "int_of_diff";
-val zadd_zminus_inverse = thm "zadd_zminus_inverse";
-val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2";
-val zadd_int0_right_intify = thm "zadd_int0_right_intify";
-val zadd_int0_right = thm "zadd_int0_right";
-val zmult_congruent2 = thm "zmult_congruent2";
-val zmult_type = thm "zmult_type";
-val zmult = thm "zmult";
-val zmult_int0 = thm "zmult_int0";
-val zmult_int1_intify = thm "zmult_int1_intify";
-val zmult_int1 = thm "zmult_int1";
-val zmult_commute = thm "zmult_commute";
-val zmult_zminus = thm "zmult_zminus";
-val zmult_zminus_right = thm "zmult_zminus_right";
-val zmult_assoc = thm "zmult_assoc";
-val zmult_left_commute = thm "zmult_left_commute";
-val zmult_ac = thms "zmult_ac";
-val zadd_zmult_distrib = thm "zadd_zmult_distrib";
-val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2";
-val int_typechecks = thms "int_typechecks";
-val zdiff_type = thm "zdiff_type";
-val zminus_zdiff_eq = thm "zminus_zdiff_eq";
-val zdiff_zmult_distrib = thm "zdiff_zmult_distrib";
-val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2";
-val zadd_zdiff_eq = thm "zadd_zdiff_eq";
-val zdiff_zadd_eq = thm "zdiff_zadd_eq";
-val zless_linear = thm "zless_linear";
-val zless_not_refl = thm "zless_not_refl";
-val neq_iff_zless = thm "neq_iff_zless";
-val zless_imp_intify_neq = thm "zless_imp_intify_neq";
-val zless_imp_succ_zadd = thm "zless_imp_succ_zadd";
-val zless_succ_zadd = thm "zless_succ_zadd";
-val zless_iff_succ_zadd = thm "zless_iff_succ_zadd";
-val zless_int_of = thm "zless_int_of";
-val zless_trans = thm "zless_trans";
-val zless_not_sym = thm "zless_not_sym";
-val zless_asym = thm "zless_asym";
-val zless_imp_zle = thm "zless_imp_zle";
-val zle_linear = thm "zle_linear";
-val zle_refl = thm "zle_refl";
-val zle_eq_refl = thm "zle_eq_refl";
-val zle_anti_sym_intify = thm "zle_anti_sym_intify";
-val zle_anti_sym = thm "zle_anti_sym";
-val zle_trans = thm "zle_trans";
-val zle_zless_trans = thm "zle_zless_trans";
-val zless_zle_trans = thm "zless_zle_trans";
-val not_zless_iff_zle = thm "not_zless_iff_zle";
-val not_zle_iff_zless = thm "not_zle_iff_zless";
-val zdiff_zdiff_eq = thm "zdiff_zdiff_eq";
-val zdiff_zdiff_eq2 = thm "zdiff_zdiff_eq2";
-val zdiff_zless_iff = thm "zdiff_zless_iff";
-val zless_zdiff_iff = thm "zless_zdiff_iff";
-val zdiff_eq_iff = thm "zdiff_eq_iff";
-val eq_zdiff_iff = thm "eq_zdiff_iff";
-val zdiff_zle_iff = thm "zdiff_zle_iff";
-val zle_zdiff_iff = thm "zle_zdiff_iff";
-val zcompare_rls = thms "zcompare_rls";
-val zadd_left_cancel = thm "zadd_left_cancel";
-val zadd_left_cancel_intify = thm "zadd_left_cancel_intify";
-val zadd_right_cancel = thm "zadd_right_cancel";
-val zadd_right_cancel_intify = thm "zadd_right_cancel_intify";
-val zadd_right_cancel_zless = thm "zadd_right_cancel_zless";
-val zadd_left_cancel_zless = thm "zadd_left_cancel_zless";
-val zadd_right_cancel_zle = thm "zadd_right_cancel_zle";
-val zadd_left_cancel_zle = thm "zadd_left_cancel_zle";
-val zadd_zless_mono1 = thm "zadd_zless_mono1";
-val zadd_zless_mono2 = thm "zadd_zless_mono2";
-val zadd_zle_mono1 = thm "zadd_zle_mono1";
-val zadd_zle_mono2 = thm "zadd_zle_mono2";
-val zadd_zle_mono = thm "zadd_zle_mono";
-val zadd_zless_mono = thm "zadd_zless_mono";
-val zminus_zless_zminus = thm "zminus_zless_zminus";
-val zminus_zle_zminus = thm "zminus_zle_zminus";
-val equation_zminus = thm "equation_zminus";
-val zminus_equation = thm "zminus_equation";
-val equation_zminus_intify = thm "equation_zminus_intify";
-val zminus_equation_intify = thm "zminus_equation_intify";
-val zless_zminus = thm "zless_zminus";
-val zminus_zless = thm "zminus_zless";
-val zle_zminus = thm "zle_zminus";
-val zminus_zle = thm "zminus_zle";
-*}
-
-
 end
--- a/src/ZF/IntDiv.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/IntDiv.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -33,20 +33,22 @@
 
 theory IntDiv imports IntArith OrderArith begin
 
-constdefs
-  quorem :: "[i,i] => o"
+definition
+  quorem :: "[i,i] => o"  where
     "quorem == %<a,b> <q,r>.
                       a = b$*q $+ r &
                       (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)"
 
-  adjust :: "[i,i] => i"
+definition
+  adjust :: "[i,i] => i"  where
     "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b>
                           else <#2$*q,r>"
 
 
 (** the division algorithm **)
 
-constdefs posDivAlg :: "i => i"
+definition
+  posDivAlg :: "i => i"  where
 (*for the case a>=0, b>0*)
 (*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
     "posDivAlg(ab) ==
@@ -57,7 +59,8 @@
 
 
 (*for the case a<0, b>0*)
-constdefs negDivAlg :: "i => i"
+definition
+  negDivAlg :: "i => i"  where
 (*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
     "negDivAlg(ab) ==
        wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
@@ -67,13 +70,14 @@
 
 (*for the general case b\<noteq>0*)
 
-constdefs
-  negateSnd :: "i => i"
+definition
+  negateSnd :: "i => i"  where
     "negateSnd == %<q,r>. <q, $-r>"
 
   (*The full division algorithm considers all possible signs for a, b
     including the special case a=0, b<0, because negDivAlg requires a<0*)
-  divAlg :: "i => i"
+definition
+  divAlg :: "i => i"  where
     "divAlg ==
        %<a,b>. if #0 $<= a then
                   if #0 $<= b then posDivAlg (<a,b>)
@@ -83,10 +87,12 @@
                   if #0$<b then negDivAlg (<a,b>)
                   else         negateSnd (posDivAlg (<$-a,$-b>))"
 
-  zdiv  :: "[i,i]=>i"                    (infixl "zdiv" 70) 
+definition
+  zdiv  :: "[i,i]=>i"                    (infixl "zdiv" 70)  where
     "a zdiv b == fst (divAlg (<intify(a), intify(b)>))"
 
-  zmod  :: "[i,i]=>i"                    (infixl "zmod" 70)
+definition
+  zmod  :: "[i,i]=>i"                    (infixl "zmod" 70)  where
     "a zmod b == snd (divAlg (<intify(a), intify(b)>))"
 
 
@@ -1779,147 +1785,5 @@
  declare zmod_integ_of_BIT [simp]
 *)
 
-ML{*
-val zspos_add_zspos_imp_zspos = thm "zspos_add_zspos_imp_zspos";
-val zpos_add_zpos_imp_zpos = thm "zpos_add_zpos_imp_zpos";
-val zneg_add_zneg_imp_zneg = thm "zneg_add_zneg_imp_zneg";
-val zneg_or_0_add_zneg_or_0_imp_zneg_or_0 = thm "zneg_or_0_add_zneg_or_0_imp_zneg_or_0";
-val zero_lt_zmagnitude = thm "zero_lt_zmagnitude";
-val zless_add_succ_iff = thm "zless_add_succ_iff";
-val zadd_succ_zle_iff = thm "zadd_succ_zle_iff";
-val zless_add1_iff_zle = thm "zless_add1_iff_zle";
-val add1_zle_iff = thm "add1_zle_iff";
-val add1_left_zle_iff = thm "add1_left_zle_iff";
-val zmult_zle_mono1 = thm "zmult_zle_mono1";
-val zmult_zle_mono1_neg = thm "zmult_zle_mono1_neg";
-val zmult_zle_mono2 = thm "zmult_zle_mono2";
-val zmult_zle_mono2_neg = thm "zmult_zle_mono2_neg";
-val zmult_zle_mono = thm "zmult_zle_mono";
-val zmult_zless_mono2 = thm "zmult_zless_mono2";
-val zmult_zless_mono1 = thm "zmult_zless_mono1";
-val zmult_zless_mono = thm "zmult_zless_mono";
-val zmult_zless_mono1_neg = thm "zmult_zless_mono1_neg";
-val zmult_zless_mono2_neg = thm "zmult_zless_mono2_neg";
-val zmult_eq_0_iff = thm "zmult_eq_0_iff";
-val zmult_zless_cancel2 = thm "zmult_zless_cancel2";
-val zmult_zless_cancel1 = thm "zmult_zless_cancel1";
-val zmult_zle_cancel2 = thm "zmult_zle_cancel2";
-val zmult_zle_cancel1 = thm "zmult_zle_cancel1";
-val int_eq_iff_zle = thm "int_eq_iff_zle";
-val zmult_cancel2 = thm "zmult_cancel2";
-val zmult_cancel1 = thm "zmult_cancel1";
-val unique_quotient = thm "unique_quotient";
-val unique_remainder = thm "unique_remainder";
-val adjust_eq = thm "adjust_eq";
-val posDivAlg_termination = thm "posDivAlg_termination";
-val posDivAlg_unfold = thm "posDivAlg_unfold";
-val posDivAlg_eqn = thm "posDivAlg_eqn";
-val posDivAlg_induct = thm "posDivAlg_induct";
-val intify_eq_0_iff_zle = thm "intify_eq_0_iff_zle";
-val zmult_pos = thm "zmult_pos";
-val zmult_neg = thm "zmult_neg";
-val zmult_pos_neg = thm "zmult_pos_neg";
-val int_0_less_mult_iff = thm "int_0_less_mult_iff";
-val int_0_le_mult_iff = thm "int_0_le_mult_iff";
-val zmult_less_0_iff = thm "zmult_less_0_iff";
-val zmult_le_0_iff = thm "zmult_le_0_iff";
-val posDivAlg_type = thm "posDivAlg_type";
-val posDivAlg_correct = thm "posDivAlg_correct";
-val negDivAlg_termination = thm "negDivAlg_termination";
-val negDivAlg_unfold = thm "negDivAlg_unfold";
-val negDivAlg_eqn = thm "negDivAlg_eqn";
-val negDivAlg_induct = thm "negDivAlg_induct";
-val negDivAlg_type = thm "negDivAlg_type";
-val negDivAlg_correct = thm "negDivAlg_correct";
-val quorem_0 = thm "quorem_0";
-val posDivAlg_zero_divisor = thm "posDivAlg_zero_divisor";
-val posDivAlg_0 = thm "posDivAlg_0";
-val negDivAlg_minus1 = thm "negDivAlg_minus1";
-val negateSnd_eq = thm "negateSnd_eq";
-val negateSnd_type = thm "negateSnd_type";
-val quorem_neg = thm "quorem_neg";
-val divAlg_correct = thm "divAlg_correct";
-val divAlg_type = thm "divAlg_type";
-val zdiv_intify1 = thm "zdiv_intify1";
-val zdiv_intify2 = thm "zdiv_intify2";
-val zdiv_type = thm "zdiv_type";
-val zmod_intify1 = thm "zmod_intify1";
-val zmod_intify2 = thm "zmod_intify2";
-val zmod_type = thm "zmod_type";
-val DIVISION_BY_ZERO_ZDIV = thm "DIVISION_BY_ZERO_ZDIV";
-val DIVISION_BY_ZERO_ZMOD = thm "DIVISION_BY_ZERO_ZMOD";
-val zmod_zdiv_equality = thm "zmod_zdiv_equality";
-val pos_mod = thm "pos_mod";
-val pos_mod_sign = thm "pos_mod_sign";
-val neg_mod = thm "neg_mod";
-val neg_mod_sign = thm "neg_mod_sign";
-val quorem_div_mod = thm "quorem_div_mod";
-val quorem_div = thm "quorem_div";
-val quorem_mod = thm "quorem_mod";
-val zdiv_pos_pos_trivial = thm "zdiv_pos_pos_trivial";
-val zdiv_neg_neg_trivial = thm "zdiv_neg_neg_trivial";
-val zdiv_pos_neg_trivial = thm "zdiv_pos_neg_trivial";
-val zmod_pos_pos_trivial = thm "zmod_pos_pos_trivial";
-val zmod_neg_neg_trivial = thm "zmod_neg_neg_trivial";
-val zmod_pos_neg_trivial = thm "zmod_pos_neg_trivial";
-val zdiv_zminus_zminus = thm "zdiv_zminus_zminus";
-val zmod_zminus_zminus = thm "zmod_zminus_zminus";
-val self_quotient = thm "self_quotient";
-val self_remainder = thm "self_remainder";
-val zdiv_self = thm "zdiv_self";
-val zmod_self = thm "zmod_self";
-val zdiv_zero = thm "zdiv_zero";
-val zdiv_eq_minus1 = thm "zdiv_eq_minus1";
-val zmod_zero = thm "zmod_zero";
-val zdiv_minus1 = thm "zdiv_minus1";
-val zmod_minus1 = thm "zmod_minus1";
-val zdiv_pos_pos = thm "zdiv_pos_pos";
-val zmod_pos_pos = thm "zmod_pos_pos";
-val zdiv_neg_pos = thm "zdiv_neg_pos";
-val zmod_neg_pos = thm "zmod_neg_pos";
-val zdiv_pos_neg = thm "zdiv_pos_neg";
-val zmod_pos_neg = thm "zmod_pos_neg";
-val zdiv_neg_neg = thm "zdiv_neg_neg";
-val zmod_neg_neg = thm "zmod_neg_neg";
-val zmod_1 = thm "zmod_1";
-val zdiv_1 = thm "zdiv_1";
-val zmod_minus1_right = thm "zmod_minus1_right";
-val zdiv_minus1_right = thm "zdiv_minus1_right";
-val zdiv_mono1 = thm "zdiv_mono1";
-val zdiv_mono1_neg = thm "zdiv_mono1_neg";
-val zdiv_mono2 = thm "zdiv_mono2";
-val zdiv_mono2_neg = thm "zdiv_mono2_neg";
-val zdiv_zmult1_eq = thm "zdiv_zmult1_eq";
-val zmod_zmult1_eq = thm "zmod_zmult1_eq";
-val zmod_zmult1_eq' = thm "zmod_zmult1_eq'";
-val zmod_zmult_distrib = thm "zmod_zmult_distrib";
-val zdiv_zmult_self1 = thm "zdiv_zmult_self1";
-val zdiv_zmult_self2 = thm "zdiv_zmult_self2";
-val zmod_zmult_self1 = thm "zmod_zmult_self1";
-val zmod_zmult_self2 = thm "zmod_zmult_self2";
-val zdiv_zadd1_eq = thm "zdiv_zadd1_eq";
-val zmod_zadd1_eq = thm "zmod_zadd1_eq";
-val zmod_div_trivial = thm "zmod_div_trivial";
-val zmod_mod_trivial = thm "zmod_mod_trivial";
-val zmod_zadd_left_eq = thm "zmod_zadd_left_eq";
-val zmod_zadd_right_eq = thm "zmod_zadd_right_eq";
-val zdiv_zadd_self1 = thm "zdiv_zadd_self1";
-val zdiv_zadd_self2 = thm "zdiv_zadd_self2";
-val zmod_zadd_self1 = thm "zmod_zadd_self1";
-val zmod_zadd_self2 = thm "zmod_zadd_self2";
-val zdiv_zmult2_eq = thm "zdiv_zmult2_eq";
-val zmod_zmult2_eq = thm "zmod_zmult2_eq";
-val zdiv_zmult_zmult1 = thm "zdiv_zmult_zmult1";
-val zdiv_zmult_zmult2 = thm "zdiv_zmult_zmult2";
-val zmod_zmult_zmult1 = thm "zmod_zmult_zmult1";
-val zmod_zmult_zmult2 = thm "zmod_zmult_zmult2";
-val zdiv_neg_pos_less0 = thm "zdiv_neg_pos_less0";
-val zdiv_nonneg_neg_le0 = thm "zdiv_nonneg_neg_le0";
-val pos_imp_zdiv_nonneg_iff = thm "pos_imp_zdiv_nonneg_iff";
-val neg_imp_zdiv_nonneg_iff = thm "neg_imp_zdiv_nonneg_iff";
-val pos_imp_zdiv_neg_iff = thm "pos_imp_zdiv_neg_iff";
-val neg_imp_zdiv_neg_iff = thm "neg_imp_zdiv_neg_iff";
-*}
-
 end
 
--- a/src/ZF/List.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/List.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -92,19 +92,21 @@
 
 (*** Thanks to Sidi Ehmety for the following ***)
 
-constdefs
+definition
 (* Function `take' returns the first n elements of a list *)
-  take     :: "[i,i]=>i"
+  take     :: "[i,i]=>i"  where
   "take(n, as) == list_rec(lam n:nat. [],
 		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
 
-  nth :: "[i, i]=>i"
+definition
+  nth :: "[i, i]=>i"  where
   --{*returns the (n+1)th element of a list, or 0 if the
    list is too short.*}
   "nth(n, as) == list_rec(lam n:nat. 0,
  		          %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
 
-  list_update :: "[i, i, i]=>i"
+definition
+  list_update :: "[i, i, i]=>i"  where
   "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
       %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
 
@@ -121,11 +123,12 @@
   "upt(i, 0) = Nil"
   "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
 
-constdefs
-  min :: "[i,i] =>i"
+definition
+  min :: "[i,i] =>i"  where
     "min(x, y) == (if x le y then x else y)"
 
-  max :: "[i, i] =>i"
+definition
+  max :: "[i, i] =>i"  where
     "max(x, y) == (if x le y then y else x)"
 
 (*** Aspects of the datatype definition ***)
@@ -852,8 +855,8 @@
      (\<lambda>ys \<in> list(B).
         list_case(Nil, %y zs. Cons(<x,y>, zip_aux(B,l)`zs), ys))"
 
-constdefs
-  zip :: "[i, i]=>i"
+definition
+  zip :: "[i, i]=>i"  where
    "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys"
 
 
@@ -1157,8 +1160,8 @@
 
 (** sublist (a generalization of nth to sets) **)
 
-constdefs
-  sublist :: "[i, i] => i"
+definition
+  sublist :: "[i, i] => i"  where
     "sublist(xs, A)==
      map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
 
@@ -1245,170 +1248,4 @@
 lemma repeat_type [TC]: "[|a \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)"
 by (induct_tac n, auto)
 
-
-ML
-{*
-val ConsE = thm "ConsE";
-val Cons_iff = thm "Cons_iff";
-val Nil_Cons_iff = thm "Nil_Cons_iff";
-val list_unfold = thm "list_unfold";
-val list_mono = thm "list_mono";
-val list_univ = thm "list_univ";
-val list_subset_univ = thm "list_subset_univ";
-val list_into_univ = thm "list_into_univ";
-val list_case_type = thm "list_case_type";
-val tl_type = thm "tl_type";
-val drop_Nil = thm "drop_Nil";
-val drop_succ_Cons = thm "drop_succ_Cons";
-val drop_type = thm "drop_type";
-val list_rec_type = thm "list_rec_type";
-val map_type = thm "map_type";
-val map_type2 = thm "map_type2";
-val length_type = thm "length_type";
-val lt_length_in_nat = thm "lt_length_in_nat";
-val app_type = thm "app_type";
-val rev_type = thm "rev_type";
-val flat_type = thm "flat_type";
-val set_of_list_type = thm "set_of_list_type";
-val set_of_list_append = thm "set_of_list_append";
-val list_add_type = thm "list_add_type";
-val map_ident = thm "map_ident";
-val map_compose = thm "map_compose";
-val map_app_distrib = thm "map_app_distrib";
-val map_flat = thm "map_flat";
-val list_rec_map = thm "list_rec_map";
-val list_CollectD = thm "list_CollectD";
-val map_list_Collect = thm "map_list_Collect";
-val length_map = thm "length_map";
-val length_app = thm "length_app";
-val length_rev = thm "length_rev";
-val length_flat = thm "length_flat";
-val drop_length_Cons = thm "drop_length_Cons";
-val drop_length = thm "drop_length";
-val app_right_Nil = thm "app_right_Nil";
-val app_assoc = thm "app_assoc";
-val flat_app_distrib = thm "flat_app_distrib";
-val rev_map_distrib = thm "rev_map_distrib";
-val rev_app_distrib = thm "rev_app_distrib";
-val rev_rev_ident = thm "rev_rev_ident";
-val rev_flat = thm "rev_flat";
-val list_add_app = thm "list_add_app";
-val list_add_rev = thm "list_add_rev";
-val list_add_flat = thm "list_add_flat";
-val list_append_induct = thm "list_append_induct";
-val min_sym = thm "min_sym";
-val min_type = thm "min_type";
-val min_0 = thm "min_0";
-val min_02 = thm "min_02";
-val lt_min_iff = thm "lt_min_iff";
-val min_succ_succ = thm "min_succ_succ";
-val filter_append = thm "filter_append";
-val filter_type = thm "filter_type";
-val length_filter = thm "length_filter";
-val filter_is_subset = thm "filter_is_subset";
-val filter_False = thm "filter_False";
-val filter_True = thm "filter_True";
-val length_is_0_iff = thm "length_is_0_iff";
-val length_is_0_iff2 = thm "length_is_0_iff2";
-val length_tl = thm "length_tl";
-val length_greater_0_iff = thm "length_greater_0_iff";
-val length_succ_iff = thm "length_succ_iff";
-val append_is_Nil_iff = thm "append_is_Nil_iff";
-val append_is_Nil_iff2 = thm "append_is_Nil_iff2";
-val append_left_is_self_iff = thm "append_left_is_self_iff";
-val append_left_is_self_iff2 = thm "append_left_is_self_iff2";
-val append_left_is_Nil_iff = thm "append_left_is_Nil_iff";
-val append_left_is_Nil_iff2 = thm "append_left_is_Nil_iff2";
-val append_eq_append_iff = thm "append_eq_append_iff";
-val append_eq_append = thm "append_eq_append";
-val append_eq_append_iff2 = thm "append_eq_append_iff2";
-val append_self_iff = thm "append_self_iff";
-val append_self_iff2 = thm "append_self_iff2";
-val append1_eq_iff = thm "append1_eq_iff";
-val append_right_is_self_iff = thm "append_right_is_self_iff";
-val append_right_is_self_iff2 = thm "append_right_is_self_iff2";
-val hd_append = thm "hd_append";
-val tl_append = thm "tl_append";
-val rev_is_Nil_iff = thm "rev_is_Nil_iff";
-val Nil_is_rev_iff = thm "Nil_is_rev_iff";
-val rev_is_rev_iff = thm "rev_is_rev_iff";
-val rev_list_elim = thm "rev_list_elim";
-val length_drop = thm "length_drop";
-val drop_all = thm "drop_all";
-val drop_append = thm "drop_append";
-val drop_drop = thm "drop_drop";
-val take_0 = thm "take_0";
-val take_succ_Cons = thm "take_succ_Cons";
-val take_Nil = thm "take_Nil";
-val take_all = thm "take_all";
-val take_type = thm "take_type";
-val take_append = thm "take_append";
-val take_take = thm "take_take";
-val take_add = thm "take_add";
-val take_succ = thm "take_succ";
-val nth_0 = thm "nth_0";
-val nth_Cons = thm "nth_Cons";
-val nth_type = thm "nth_type";
-val nth_append = thm "nth_append";
-val set_of_list_conv_nth = thm "set_of_list_conv_nth";
-val nth_take_lemma = thm "nth_take_lemma";
-val nth_equalityI = thm "nth_equalityI";
-val take_equalityI = thm "take_equalityI";
-val nth_drop = thm "nth_drop";
-val list_on_set_of_list = thm "list_on_set_of_list";
-val zip_Nil = thm "zip_Nil";
-val zip_Nil2 = thm "zip_Nil2";
-val zip_Cons_Cons = thm "zip_Cons_Cons";
-val zip_type = thm "zip_type";
-val length_zip = thm "length_zip";
-val zip_append1 = thm "zip_append1";
-val zip_append2 = thm "zip_append2";
-val zip_append = thm "zip_append";
-val zip_rev = thm "zip_rev";
-val nth_zip = thm "nth_zip";
-val set_of_list_zip = thm "set_of_list_zip";
-val list_update_Nil = thm "list_update_Nil";
-val list_update_Cons_0 = thm "list_update_Cons_0";
-val list_update_Cons_succ = thm "list_update_Cons_succ";
-val list_update_type = thm "list_update_type";
-val length_list_update = thm "length_list_update";
-val nth_list_update = thm "nth_list_update";
-val nth_list_update_eq = thm "nth_list_update_eq";
-val nth_list_update_neq = thm "nth_list_update_neq";
-val list_update_overwrite = thm "list_update_overwrite";
-val list_update_same_conv = thm "list_update_same_conv";
-val update_zip = thm "update_zip";
-val set_update_subset_cons = thm "set_update_subset_cons";
-val set_of_list_update_subsetI = thm "set_of_list_update_subsetI";
-val upt_rec = thm "upt_rec";
-val upt_conv_Nil = thm "upt_conv_Nil";
-val upt_succ_append = thm "upt_succ_append";
-val upt_conv_Cons = thm "upt_conv_Cons";
-val upt_type = thm "upt_type";
-val upt_add_eq_append = thm "upt_add_eq_append";
-val length_upt = thm "length_upt";
-val nth_upt = thm "nth_upt";
-val take_upt = thm "take_upt";
-val map_succ_upt = thm "map_succ_upt";
-val nth_map = thm "nth_map";
-val nth_map_upt = thm "nth_map_upt";
-val sublist_0 = thm "sublist_0";
-val sublist_Nil = thm "sublist_Nil";
-val sublist_shift_lemma = thm "sublist_shift_lemma";
-val sublist_type = thm "sublist_type";
-val upt_add_eq_append2 = thm "upt_add_eq_append2";
-val sublist_append = thm "sublist_append";
-val sublist_Cons = thm "sublist_Cons";
-val sublist_singleton = thm "sublist_singleton";
-val sublist_upt_eq_take = thm "sublist_upt_eq_take";
-val sublist_Int_eq = thm "sublist_Int_eq";
-
-structure list =
-struct
-val induct = thm "list.induct"
-val elim   = thm "list.cases"
-val intrs  = thms "list.intros"
-end;  
-*}
-
 end
--- a/src/ZF/Main.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Main.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -15,14 +15,14 @@
     "F^0 (x) = x"
     "F^(succ(n)) (x) = F(F^n (x))"
 
-constdefs
-  iterates_omega :: "[i=>i,i] => i"
+definition
+  iterates_omega :: "[i=>i,i] => i"  where
     "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
 
-syntax (xsymbols)
-  iterates_omega :: "[i=>i,i] => i"   ("(_^\<omega> '(_'))" [60,1000] 60)
-syntax (HTML output)
-  iterates_omega :: "[i=>i,i] => i"   ("(_^\<omega> '(_'))" [60,1000] 60)
+notation (xsymbols)
+  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
+notation (HTML output)
+  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
 
 lemma iterates_triv:
      "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"  
@@ -51,8 +51,8 @@
 text{*Transfinite recursion for definitions based on the 
     three cases of ordinals*}
 
-constdefs
-  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i"
+definition
+  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
     "transrec3(k, a, b, c) ==                     
        transrec(k, \<lambda>x r.
          if x=0 then a
--- a/src/ZF/Nat.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Nat.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,36 +9,44 @@
 
 theory Nat imports OrdQuant Bool begin
 
-constdefs
-  nat :: i
+definition
+  nat :: i  where
     "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"
 
-  quasinat :: "i => o"
+definition
+  quasinat :: "i => o"  where
     "quasinat(n) == n=0 | (\<exists>m. n = succ(m))"
 
+definition
   (*Has an unconditional succ case, which is used in "recursor" below.*)
-  nat_case :: "[i, i=>i, i]=>i"
+  nat_case :: "[i, i=>i, i]=>i"  where
     "nat_case(a,b,k) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))"
 
-  nat_rec :: "[i, i, [i,i]=>i]=>i"
+definition
+  nat_rec :: "[i, i, [i,i]=>i]=>i"  where
     "nat_rec(k,a,b) ==   
           wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))"
 
   (*Internalized relations on the naturals*)
   
-  Le :: i
+definition
+  Le :: i  where
     "Le == {<x,y>:nat*nat. x le y}"
 
-  Lt :: i
+definition
+  Lt :: i  where
     "Lt == {<x, y>:nat*nat. x < y}"
   
-  Ge :: i
+definition
+  Ge :: i  where
     "Ge == {<x,y>:nat*nat. y le x}"
 
-  Gt :: i
+definition
+  Gt :: i  where
     "Gt == {<x,y>:nat*nat. y < x}"
 
-  greater_than :: "i=>i"
+definition
+  greater_than :: "i=>i"  where
     "greater_than(n) == {i:nat. n < i}"
 
 text{*No need for a less-than operator: a natural number is its list of
@@ -291,47 +299,4 @@
 lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat"
 by (force simp add: Le_def)
 
-ML
-{*
-val Le_def = thm "Le_def";
-val Lt_def = thm "Lt_def";
-val Ge_def = thm "Ge_def";
-val Gt_def = thm "Gt_def";
-val greater_than_def = thm "greater_than_def";
-
-val nat_bnd_mono = thm "nat_bnd_mono";
-val nat_unfold = thm "nat_unfold";
-val nat_0I = thm "nat_0I";
-val nat_succI = thm "nat_succI";
-val nat_1I = thm "nat_1I";
-val nat_2I = thm "nat_2I";
-val bool_subset_nat = thm "bool_subset_nat";
-val bool_into_nat = thm "bool_into_nat";
-val nat_induct = thm "nat_induct";
-val natE = thm "natE";
-val nat_into_Ord = thm "nat_into_Ord";
-val nat_0_le = thm "nat_0_le";
-val nat_le_refl = thm "nat_le_refl";
-val Ord_nat = thm "Ord_nat";
-val Limit_nat = thm "Limit_nat";
-val succ_natD = thm "succ_natD";
-val nat_succ_iff = thm "nat_succ_iff";
-val nat_le_Limit = thm "nat_le_Limit";
-val succ_in_naturalD = thm "succ_in_naturalD";
-val lt_nat_in_nat = thm "lt_nat_in_nat";
-val le_in_nat = thm "le_in_nat";
-val complete_induct = thm "complete_induct";
-val nat_induct_from = thm "nat_induct_from";
-val diff_induct = thm "diff_induct";
-val succ_lt_induct = thm "succ_lt_induct";
-val nat_case_0 = thm "nat_case_0";
-val nat_case_succ = thm "nat_case_succ";
-val nat_case_type = thm "nat_case_type";
-val nat_rec_0 = thm "nat_rec_0";
-val nat_rec_succ = thm "nat_rec_succ";
-val Un_nat_type = thm "Un_nat_type";
-val Int_nat_type = thm "Int_nat_type";
-val nat_nonempty = thm "nat_nonempty";
-*}
-
 end
--- a/src/ZF/OrdQuant.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/OrdQuant.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,17 +9,18 @@
 
 subsection {*Quantifiers and union operator for ordinals*}
 
-constdefs
-
+definition
   (* Ordinal Quantifiers *)
-  oall :: "[i, i => o] => o"
+  oall :: "[i, i => o] => o"  where
     "oall(A, P) == ALL x. x<A --> P(x)"
 
-  oex :: "[i, i => o] => o"
+definition
+  oex :: "[i, i => o] => o"  where
     "oex(A, P)  == EX x. x<A & P(x)"
 
+definition
   (* Ordinal Union *)
-  OUnion :: "[i, i => i] => i"
+  OUnion :: "[i, i => i] => i"  where
     "OUnion(i,B) == {z: \<Union>x\<in>i. B(x). Ord(i)}"
 
 syntax
@@ -28,9 +29,9 @@
   "@OUNION"   :: "[idt, i, i] => i"        ("(3UN _<_./ _)" 10)
 
 translations
-  "ALL x<a. P"  == "oall(a, %x. P)"
-  "EX x<a. P"   == "oex(a, %x. P)"
-  "UN x<a. B"   == "OUnion(a, %x. B)"
+  "ALL x<a. P"  == "CONST oall(a, %x. P)"
+  "EX x<a. P"   == "CONST oex(a, %x. P)"
+  "UN x<a. B"   == "CONST OUnion(a, %x. B)"
 
 syntax (xsymbols)
   "@oall"     :: "[idt, i, o] => o"        ("(3\<forall>_<_./ _)" 10)
@@ -194,11 +195,12 @@
 
 subsection {*Quantification over a class*}
 
-constdefs
-  "rall"     :: "[i=>o, i=>o] => o"
+definition
+  "rall"     :: "[i=>o, i=>o] => o"  where
     "rall(M, P) == ALL x. M(x) --> P(x)"
 
-  "rex"      :: "[i=>o, i=>o] => o"
+definition
+  "rex"      :: "[i=>o, i=>o] => o"  where
     "rex(M, P) == EX x. M(x) & P(x)"
 
 syntax
@@ -213,8 +215,8 @@
   "@rex"      :: "[pttrn, i=>o, o] => o"        ("(3\<exists>_[_]./ _)" 10)
 
 translations
-  "ALL x[M]. P"  == "rall(M, %x. P)"
-  "EX x[M]. P"   == "rex(M, %x. P)"
+  "ALL x[M]. P"  == "CONST rall(M, %x. P)"
+  "EX x[M]. P"   == "CONST rex(M, %x. P)"
 
 
 subsubsection{*Relativized universal quantifier*}
@@ -340,7 +342,8 @@
 
 subsubsection{*Sets as Classes*}
 
-constdefs setclass :: "[i,i] => o"       ("##_" [40] 40)
+definition
+  setclass :: "[i,i] => o"       ("##_" [40] 40)  where
    "setclass(A) == %x. x : A"
 
 lemma setclass_iff [simp]: "setclass(A,x) <-> x : A"
@@ -355,41 +358,8 @@
 
 ML
 {*
-val oall_def = thm "oall_def"
-val oex_def = thm "oex_def"
-val OUnion_def = thm "OUnion_def"
-
-val oallI = thm "oallI";
-val ospec = thm "ospec";
-val oallE = thm "oallE";
-val rev_oallE = thm "rev_oallE";
-val oall_simp = thm "oall_simp";
-val oall_cong = thm "oall_cong";
-val oexI = thm "oexI";
-val oexCI = thm "oexCI";
-val oexE = thm "oexE";
-val oex_cong = thm "oex_cong";
-val OUN_I = thm "OUN_I";
-val OUN_E = thm "OUN_E";
-val OUN_iff = thm "OUN_iff";
-val OUN_cong = thm "OUN_cong";
-val lt_induct = thm "lt_induct";
-
-val rall_def = thm "rall_def"
-val rex_def = thm "rex_def"
-
-val rallI = thm "rallI";
-val rspec = thm "rspec";
-val rallE = thm "rallE";
-val rev_oallE = thm "rev_oallE";
-val rall_cong = thm "rall_cong";
-val rexI = thm "rexI";
-val rexCI = thm "rexCI";
-val rexE = thm "rexE";
-val rex_cong = thm "rex_cong";
-
 val Ord_atomize =
-    atomize ([("OrdQuant.oall", [ospec]),("OrdQuant.rall", [rspec])]@
+    atomize ([("OrdQuant.oall", [@{thm ospec}]),("OrdQuant.rall", [@{thm rspec}])]@
                  ZF_conn_pairs,
              ZF_mem_pairs);
 change_simpset (fn ss => ss setmksimps (map mk_eq o Ord_atomize o gen_all));
@@ -400,19 +370,19 @@
 ML_setup {*
 local
 
-val unfold_rex_tac = unfold_tac [rex_def];
+val unfold_rex_tac = unfold_tac [@{thm rex_def}];
 fun prove_rex_tac ss = unfold_rex_tac ss THEN Quantifier1.prove_one_point_ex_tac;
 val rearrange_bex = Quantifier1.rearrange_bex prove_rex_tac;
 
-val unfold_rall_tac = unfold_tac [rall_def];
+val unfold_rall_tac = unfold_tac [@{thm rall_def}];
 fun prove_rall_tac ss = unfold_rall_tac ss THEN Quantifier1.prove_one_point_all_tac;
 val rearrange_ball = Quantifier1.rearrange_ball prove_rall_tac;
 
 in
 
-val defREX_regroup = Simplifier.simproc (the_context ())
+val defREX_regroup = Simplifier.simproc @{theory}
   "defined REX" ["EX x[M]. P(x) & Q(x)"] rearrange_bex;
-val defRALL_regroup = Simplifier.simproc (the_context ())
+val defRALL_regroup = Simplifier.simproc @{theory}
   "defined RALL" ["ALL x[M]. P(x) --> Q(x)"] rearrange_ball;
 
 end;
--- a/src/ZF/Order.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Order.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -11,41 +11,48 @@
 
 theory Order imports WF Perm begin
 
-constdefs
-
-  part_ord :: "[i,i]=>o"          	(*Strict partial ordering*)
+definition
+  part_ord :: "[i,i]=>o"          	(*Strict partial ordering*)  where
    "part_ord(A,r) == irrefl(A,r) & trans[A](r)"
 
-  linear   :: "[i,i]=>o"          	(*Strict total ordering*)
+definition
+  linear   :: "[i,i]=>o"          	(*Strict total ordering*)  where
    "linear(A,r) == (ALL x:A. ALL y:A. <x,y>:r | x=y | <y,x>:r)"
 
-  tot_ord  :: "[i,i]=>o"          	(*Strict total ordering*)
+definition
+  tot_ord  :: "[i,i]=>o"          	(*Strict total ordering*)  where
    "tot_ord(A,r) == part_ord(A,r) & linear(A,r)"
 
-  well_ord :: "[i,i]=>o"          	(*Well-ordering*)
+definition
+  well_ord :: "[i,i]=>o"          	(*Well-ordering*)  where
    "well_ord(A,r) == tot_ord(A,r) & wf[A](r)"
 
-  mono_map :: "[i,i,i,i]=>i"      	(*Order-preserving maps*)
+definition
+  mono_map :: "[i,i,i,i]=>i"      	(*Order-preserving maps*)  where
    "mono_map(A,r,B,s) ==
 	      {f: A->B. ALL x:A. ALL y:A. <x,y>:r --> <f`x,f`y>:s}"
 
-  ord_iso  :: "[i,i,i,i]=>i"		(*Order isomorphisms*)
+definition
+  ord_iso  :: "[i,i,i,i]=>i"		(*Order isomorphisms*)  where
    "ord_iso(A,r,B,s) ==
 	      {f: bij(A,B). ALL x:A. ALL y:A. <x,y>:r <-> <f`x,f`y>:s}"
 
-  pred     :: "[i,i,i]=>i"		(*Set of predecessors*)
+definition
+  pred     :: "[i,i,i]=>i"		(*Set of predecessors*)  where
    "pred(A,x,r) == {y:A. <y,x>:r}"
 
-  ord_iso_map :: "[i,i,i,i]=>i"      	(*Construction for linearity theorem*)
+definition
+  ord_iso_map :: "[i,i,i,i]=>i"      	(*Construction for linearity theorem*)  where
    "ord_iso_map(A,r,B,s) ==
      \<Union>x\<in>A. \<Union>y\<in>B. \<Union>f \<in> ord_iso(pred(A,x,r), r, pred(B,y,s), s). {<x,y>}"
 
-  first :: "[i, i, i] => o"
+definition
+  first :: "[i, i, i] => o"  where
     "first(u, X, R) == u:X & (ALL v:X. v~=u --> <u,v> : R)"
 
 
-syntax (xsymbols)
-    ord_iso :: "[i,i,i,i]=>i"      ("(\<langle>_, _\<rangle> \<cong>/ \<langle>_, _\<rangle>)" 51)
+notation (xsymbols)
+  ord_iso  ("(\<langle>_, _\<rangle> \<cong>/ \<langle>_, _\<rangle>)" 51)
 
 
 subsection{*Immediate Consequences of the Definitions*}
@@ -157,7 +164,7 @@
 done
 
 lemma wf_on_Int_iff: "wf[A](r Int A*A) <-> wf[A](r)"
-apply (unfold wf_on_def wf_def, fast) (*10 times faster than Blast_tac!*)
+apply (unfold wf_on_def wf_def, fast) (*10 times faster than blast!*)
 done
 
 lemma well_ord_Int_iff: "well_ord(A,r Int A*A) <-> well_ord(A,r)"
@@ -638,91 +645,4 @@
 apply (erule theI)
 done
 
-ML {*
-val pred_def = thm "pred_def"
-val linear_def = thm "linear_def"
-val part_ord_def = thm "part_ord_def"
-val tot_ord_def = thm "tot_ord_def"
-val well_ord_def = thm "well_ord_def"
-val ord_iso_def = thm "ord_iso_def"
-val mono_map_def = thm "mono_map_def";
-
-val part_ord_Imp_asym = thm "part_ord_Imp_asym";
-val linearE = thm "linearE";
-val well_ordI = thm "well_ordI";
-val well_ord_is_wf = thm "well_ord_is_wf";
-val well_ord_is_trans_on = thm "well_ord_is_trans_on";
-val well_ord_is_linear = thm "well_ord_is_linear";
-val pred_iff = thm "pred_iff";
-val predI = thm "predI";
-val predE = thm "predE";
-val pred_subset_under = thm "pred_subset_under";
-val pred_subset = thm "pred_subset";
-val pred_pred_eq = thm "pred_pred_eq";
-val trans_pred_pred_eq = thm "trans_pred_pred_eq";
-val part_ord_subset = thm "part_ord_subset";
-val linear_subset = thm "linear_subset";
-val tot_ord_subset = thm "tot_ord_subset";
-val well_ord_subset = thm "well_ord_subset";
-val irrefl_Int_iff = thm "irrefl_Int_iff";
-val trans_on_Int_iff = thm "trans_on_Int_iff";
-val part_ord_Int_iff = thm "part_ord_Int_iff";
-val linear_Int_iff = thm "linear_Int_iff";
-val tot_ord_Int_iff = thm "tot_ord_Int_iff";
-val wf_on_Int_iff = thm "wf_on_Int_iff";
-val well_ord_Int_iff = thm "well_ord_Int_iff";
-val irrefl_0 = thm "irrefl_0";
-val trans_on_0 = thm "trans_on_0";
-val part_ord_0 = thm "part_ord_0";
-val linear_0 = thm "linear_0";
-val tot_ord_0 = thm "tot_ord_0";
-val wf_on_0 = thm "wf_on_0";
-val well_ord_0 = thm "well_ord_0";
-val tot_ord_unit = thm "tot_ord_unit";
-val well_ord_unit = thm "well_ord_unit";
-val mono_map_is_fun = thm "mono_map_is_fun";
-val mono_map_is_inj = thm "mono_map_is_inj";
-val ord_isoI = thm "ord_isoI";
-val ord_iso_is_mono_map = thm "ord_iso_is_mono_map";
-val ord_iso_is_bij = thm "ord_iso_is_bij";
-val ord_iso_apply = thm "ord_iso_apply";
-val ord_iso_converse = thm "ord_iso_converse";
-val ord_iso_refl = thm "ord_iso_refl";
-val ord_iso_sym = thm "ord_iso_sym";
-val mono_map_trans = thm "mono_map_trans";
-val ord_iso_trans = thm "ord_iso_trans";
-val mono_ord_isoI = thm "mono_ord_isoI";
-val well_ord_mono_ord_isoI = thm "well_ord_mono_ord_isoI";
-val part_ord_ord_iso = thm "part_ord_ord_iso";
-val linear_ord_iso = thm "linear_ord_iso";
-val wf_on_ord_iso = thm "wf_on_ord_iso";
-val well_ord_ord_iso = thm "well_ord_ord_iso";
-val well_ord_iso_predE = thm "well_ord_iso_predE";
-val well_ord_iso_pred_eq = thm "well_ord_iso_pred_eq";
-val ord_iso_image_pred = thm "ord_iso_image_pred";
-val ord_iso_restrict_pred = thm "ord_iso_restrict_pred";
-val well_ord_iso_preserving = thm "well_ord_iso_preserving";
-val well_ord_iso_unique = thm "well_ord_iso_unique";
-val ord_iso_map_subset = thm "ord_iso_map_subset";
-val domain_ord_iso_map = thm "domain_ord_iso_map";
-val range_ord_iso_map = thm "range_ord_iso_map";
-val converse_ord_iso_map = thm "converse_ord_iso_map";
-val function_ord_iso_map = thm "function_ord_iso_map";
-val ord_iso_map_fun = thm "ord_iso_map_fun";
-val ord_iso_map_mono_map = thm "ord_iso_map_mono_map";
-val ord_iso_map_ord_iso = thm "ord_iso_map_ord_iso";
-val domain_ord_iso_map_subset = thm "domain_ord_iso_map_subset";
-val domain_ord_iso_map_cases = thm "domain_ord_iso_map_cases";
-val range_ord_iso_map_cases = thm "range_ord_iso_map_cases";
-val well_ord_trichotomy = thm "well_ord_trichotomy";
-val irrefl_converse = thm "irrefl_converse";
-val trans_on_converse = thm "trans_on_converse";
-val part_ord_converse = thm "part_ord_converse";
-val linear_converse = thm "linear_converse";
-val tot_ord_converse = thm "tot_ord_converse";
-val first_is_elem = thm "first_is_elem";
-val well_ord_imp_ex1_first = thm "well_ord_imp_ex1_first";
-val the_first_in = thm "the_first_in";
-*}
-
 end
--- a/src/ZF/OrderArith.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/OrderArith.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -8,28 +8,31 @@
 header{*Combining Orderings: Foundations of Ordinal Arithmetic*}
 
 theory OrderArith imports Order Sum Ordinal begin
-constdefs
 
+definition
   (*disjoint sum of two relations; underlies ordinal addition*)
-  radd    :: "[i,i,i,i]=>i"
+  radd    :: "[i,i,i,i]=>i"  where
     "radd(A,r,B,s) == 
                 {z: (A+B) * (A+B).  
                     (EX x y. z = <Inl(x), Inr(y)>)   |   
                     (EX x' x. z = <Inl(x'), Inl(x)> & <x',x>:r)   |      
                     (EX y' y. z = <Inr(y'), Inr(y)> & <y',y>:s)}"
 
+definition
   (*lexicographic product of two relations; underlies ordinal multiplication*)
-  rmult   :: "[i,i,i,i]=>i"
+  rmult   :: "[i,i,i,i]=>i"  where
     "rmult(A,r,B,s) == 
                 {z: (A*B) * (A*B).  
                     EX x' y' x y. z = <<x',y'>, <x,y>> &         
                        (<x',x>: r | (x'=x & <y',y>: s))}"
 
+definition
   (*inverse image of a relation*)
-  rvimage :: "[i,i,i]=>i"
+  rvimage :: "[i,i,i]=>i"  where
     "rvimage(A,f,r) == {z: A*A. EX x y. z = <x,y> & <f`x,f`y>: r}"
 
-  measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i"
+definition
+  measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i"  where
     "measure(A,f) == {<x,y>: A*A. f(x) < f(y)}"
 
 
@@ -406,12 +409,12 @@
 by (blast intro: wf_rvimage wf_Memrel)
 
 
-constdefs
-  wfrank :: "[i,i]=>i"
+definition
+  wfrank :: "[i,i]=>i"  where
     "wfrank(r,a) == wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))"
 
-constdefs
-  wftype :: "i=>i"
+definition
+  wftype :: "i=>i"  where
     "wftype(r) == \<Union>y \<in> range(r). succ(wfrank(r,y))"
 
 lemma wfrank: "wf(r) ==> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))"
@@ -564,60 +567,4 @@
 apply (rule fun_extension, auto)
 by blast
 
-
-ML {*
-val measure_def = thm "measure_def";
-val radd_Inl_Inr_iff = thm "radd_Inl_Inr_iff";
-val radd_Inl_iff = thm "radd_Inl_iff";
-val radd_Inr_iff = thm "radd_Inr_iff";
-val radd_Inr_Inl_iff = thm "radd_Inr_Inl_iff";
-val raddE = thm "raddE";
-val radd_type = thm "radd_type";
-val field_radd = thm "field_radd";
-val linear_radd = thm "linear_radd";
-val wf_on_radd = thm "wf_on_radd";
-val wf_radd = thm "wf_radd";
-val well_ord_radd = thm "well_ord_radd";
-val sum_bij = thm "sum_bij";
-val sum_ord_iso_cong = thm "sum_ord_iso_cong";
-val sum_disjoint_bij = thm "sum_disjoint_bij";
-val sum_assoc_bij = thm "sum_assoc_bij";
-val sum_assoc_ord_iso = thm "sum_assoc_ord_iso";
-val rmult_iff = thm "rmult_iff";
-val rmultE = thm "rmultE";
-val rmult_type = thm "rmult_type";
-val field_rmult = thm "field_rmult";
-val linear_rmult = thm "linear_rmult";
-val wf_on_rmult = thm "wf_on_rmult";
-val wf_rmult = thm "wf_rmult";
-val well_ord_rmult = thm "well_ord_rmult";
-val prod_bij = thm "prod_bij";
-val prod_ord_iso_cong = thm "prod_ord_iso_cong";
-val singleton_prod_bij = thm "singleton_prod_bij";
-val singleton_prod_ord_iso = thm "singleton_prod_ord_iso";
-val prod_sum_singleton_bij = thm "prod_sum_singleton_bij";
-val prod_sum_singleton_ord_iso = thm "prod_sum_singleton_ord_iso";
-val sum_prod_distrib_bij = thm "sum_prod_distrib_bij";
-val sum_prod_distrib_ord_iso = thm "sum_prod_distrib_ord_iso";
-val prod_assoc_bij = thm "prod_assoc_bij";
-val prod_assoc_ord_iso = thm "prod_assoc_ord_iso";
-val rvimage_iff = thm "rvimage_iff";
-val rvimage_type = thm "rvimage_type";
-val field_rvimage = thm "field_rvimage";
-val rvimage_converse = thm "rvimage_converse";
-val irrefl_rvimage = thm "irrefl_rvimage";
-val trans_on_rvimage = thm "trans_on_rvimage";
-val part_ord_rvimage = thm "part_ord_rvimage";
-val linear_rvimage = thm "linear_rvimage";
-val tot_ord_rvimage = thm "tot_ord_rvimage";
-val wf_rvimage = thm "wf_rvimage";
-val wf_on_rvimage = thm "wf_on_rvimage";
-val well_ord_rvimage = thm "well_ord_rvimage";
-val ord_iso_rvimage = thm "ord_iso_rvimage";
-val ord_iso_rvimage_eq = thm "ord_iso_rvimage_eq";
-val measure_eq_rvimage_Memrel = thm "measure_eq_rvimage_Memrel";
-val wf_measure = thm "wf_measure";
-val measure_iff = thm "measure_iff";
-*}
-
 end
--- a/src/ZF/OrderType.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/OrderType.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -13,35 +13,41 @@
 Ordinal arithmetic is traditionally defined in terms of order types, as it is
 here.  But a definition by transfinite recursion would be much simpler!*}
 
-constdefs
-  
-  ordermap  :: "[i,i]=>i"
+definition  
+  ordermap  :: "[i,i]=>i"  where
    "ordermap(A,r) == lam x:A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
 
-  ordertype :: "[i,i]=>i"
+definition  
+  ordertype :: "[i,i]=>i"  where
    "ordertype(A,r) == ordermap(A,r)``A"
 
+definition  
   (*alternative definition of ordinal numbers*)
-  Ord_alt   :: "i => o"
+  Ord_alt   :: "i => o"  where
    "Ord_alt(X) == well_ord(X, Memrel(X)) & (ALL u:X. u=pred(X, u, Memrel(X)))"
 
+definition  
   (*coercion to ordinal: if not, just 0*)
-  ordify    :: "i=>i"
+  ordify    :: "i=>i"  where
     "ordify(x) == if Ord(x) then x else 0"
 
+definition  
   (*ordinal multiplication*)
-  omult      :: "[i,i]=>i"           (infixl "**" 70)
+  omult      :: "[i,i]=>i"           (infixl "**" 70)  where
    "i ** j == ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))"
 
+definition  
   (*ordinal addition*)
-  raw_oadd   :: "[i,i]=>i"
+  raw_oadd   :: "[i,i]=>i"  where
     "raw_oadd(i,j) == ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))"
 
-  oadd      :: "[i,i]=>i"           (infixl "++" 65)
+definition  
+  oadd      :: "[i,i]=>i"           (infixl "++" 65)  where
     "i ++ j == raw_oadd(ordify(i),ordify(j))"
 
+definition  
   (*ordinal subtraction*)
-  odiff      :: "[i,i]=>i"           (infixl "--" 65)
+  odiff      :: "[i,i]=>i"           (infixl "--" 65)  where
     "i -- j == ordertype(i-j, Memrel(i))"
 
   
@@ -1007,121 +1013,4 @@
 lemma well_ord_Lt: "well_ord(nat,Lt)"
 by (simp add: well_ord_def wf_Lt wf_imp_wf_on tot_ord_Lt)
 
-
-
-ML {*
-val ordermap_def = thm "ordermap_def";
-val ordertype_def = thm "ordertype_def";
-val Ord_alt_def = thm "Ord_alt_def";
-val ordify_def = thm "ordify_def";
-
-val Ord_in_Ord' = thm "Ord_in_Ord'";
-val le_well_ord_Memrel = thm "le_well_ord_Memrel";
-val well_ord_Memrel = thm "well_ord_Memrel";
-val lt_pred_Memrel = thm "lt_pred_Memrel";
-val pred_Memrel = thm "pred_Memrel";
-val Ord_iso_implies_eq = thm "Ord_iso_implies_eq";
-val ordermap_type = thm "ordermap_type";
-val ordermap_eq_image = thm "ordermap_eq_image";
-val ordermap_pred_unfold = thm "ordermap_pred_unfold";
-val ordermap_unfold = thm "ordermap_unfold";
-val Ord_ordermap = thm "Ord_ordermap";
-val Ord_ordertype = thm "Ord_ordertype";
-val ordermap_mono = thm "ordermap_mono";
-val converse_ordermap_mono = thm "converse_ordermap_mono";
-val ordermap_surj = thm "ordermap_surj";
-val ordermap_bij = thm "ordermap_bij";
-val ordertype_ord_iso = thm "ordertype_ord_iso";
-val ordertype_eq = thm "ordertype_eq";
-val ordertype_eq_imp_ord_iso = thm "ordertype_eq_imp_ord_iso";
-val le_ordertype_Memrel = thm "le_ordertype_Memrel";
-val ordertype_Memrel = thm "ordertype_Memrel";
-val ordertype_0 = thm "ordertype_0";
-val bij_ordertype_vimage = thm "bij_ordertype_vimage";
-val ordermap_pred_eq_ordermap = thm "ordermap_pred_eq_ordermap";
-val ordertype_unfold = thm "ordertype_unfold";
-val ordertype_pred_subset = thm "ordertype_pred_subset";
-val ordertype_pred_lt = thm "ordertype_pred_lt";
-val ordertype_pred_unfold = thm "ordertype_pred_unfold";
-val Ord_is_Ord_alt = thm "Ord_is_Ord_alt";
-val Ord_alt_is_Ord = thm "Ord_alt_is_Ord";
-val bij_sum_0 = thm "bij_sum_0";
-val ordertype_sum_0_eq = thm "ordertype_sum_0_eq";
-val bij_0_sum = thm "bij_0_sum";
-val ordertype_0_sum_eq = thm "ordertype_0_sum_eq";
-val pred_Inl_bij = thm "pred_Inl_bij";
-val ordertype_pred_Inl_eq = thm "ordertype_pred_Inl_eq";
-val pred_Inr_bij = thm "pred_Inr_bij";
-val ordertype_pred_Inr_eq = thm "ordertype_pred_Inr_eq";
-val Ord_ordify = thm "Ord_ordify";
-val ordify_idem = thm "ordify_idem";
-val Ord_raw_oadd = thm "Ord_raw_oadd";
-val Ord_oadd = thm "Ord_oadd";
-val raw_oadd_0 = thm "raw_oadd_0";
-val oadd_0 = thm "oadd_0";
-val raw_oadd_0_left = thm "raw_oadd_0_left";
-val oadd_0_left = thm "oadd_0_left";
-val oadd_eq_if_raw_oadd = thm "oadd_eq_if_raw_oadd";
-val raw_oadd_eq_oadd = thm "raw_oadd_eq_oadd";
-val lt_oadd1 = thm "lt_oadd1";
-val oadd_le_self = thm "oadd_le_self";
-val id_ord_iso_Memrel = thm "id_ord_iso_Memrel";
-val ordertype_sum_Memrel = thm "ordertype_sum_Memrel";
-val oadd_lt_mono2 = thm "oadd_lt_mono2";
-val oadd_lt_cancel2 = thm "oadd_lt_cancel2";
-val oadd_lt_iff2 = thm "oadd_lt_iff2";
-val oadd_inject = thm "oadd_inject";
-val lt_oadd_disj = thm "lt_oadd_disj";
-val oadd_assoc = thm "oadd_assoc";
-val oadd_unfold = thm "oadd_unfold";
-val oadd_1 = thm "oadd_1";
-val oadd_succ = thm "oadd_succ";
-val oadd_UN = thm "oadd_UN";
-val oadd_Limit = thm "oadd_Limit";
-val oadd_le_self2 = thm "oadd_le_self2";
-val oadd_le_mono1 = thm "oadd_le_mono1";
-val oadd_lt_mono = thm "oadd_lt_mono";
-val oadd_le_mono = thm "oadd_le_mono";
-val oadd_le_iff2 = thm "oadd_le_iff2";
-val bij_sum_Diff = thm "bij_sum_Diff";
-val ordertype_sum_Diff = thm "ordertype_sum_Diff";
-val Ord_odiff = thm "Ord_odiff";
-val raw_oadd_ordertype_Diff = thm "raw_oadd_ordertype_Diff";
-val oadd_odiff_inverse = thm "oadd_odiff_inverse";
-val odiff_oadd_inverse = thm "odiff_oadd_inverse";
-val odiff_lt_mono2 = thm "odiff_lt_mono2";
-val Ord_omult = thm "Ord_omult";
-val pred_Pair_eq = thm "pred_Pair_eq";
-val ordertype_pred_Pair_eq = thm "ordertype_pred_Pair_eq";
-val lt_omult = thm "lt_omult";
-val omult_oadd_lt = thm "omult_oadd_lt";
-val omult_unfold = thm "omult_unfold";
-val omult_0 = thm "omult_0";
-val omult_0_left = thm "omult_0_left";
-val omult_1 = thm "omult_1";
-val omult_1_left = thm "omult_1_left";
-val oadd_omult_distrib = thm "oadd_omult_distrib";
-val omult_succ = thm "omult_succ";
-val omult_assoc = thm "omult_assoc";
-val omult_UN = thm "omult_UN";
-val omult_Limit = thm "omult_Limit";
-val lt_omult1 = thm "lt_omult1";
-val omult_le_self = thm "omult_le_self";
-val omult_le_mono1 = thm "omult_le_mono1";
-val omult_lt_mono2 = thm "omult_lt_mono2";
-val omult_le_mono2 = thm "omult_le_mono2";
-val omult_le_mono = thm "omult_le_mono";
-val omult_lt_mono = thm "omult_lt_mono";
-val omult_le_self2 = thm "omult_le_self2";
-val omult_inject = thm "omult_inject";
-
-val wf_Lt = thm "wf_Lt";
-val irrefl_Lt = thm "irrefl_Lt";
-val trans_Lt = thm "trans_Lt";
-val part_ord_Lt = thm "part_ord_Lt";
-val linear_Lt = thm "linear_Lt";
-val tot_ord_Lt = thm "tot_ord_Lt";
-val well_ord_Lt = thm "well_ord_Lt";
-*}
-
 end
--- a/src/ZF/Ordinal.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Ordinal.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,21 +9,24 @@
 
 theory Ordinal imports WF Bool equalities begin
 
-constdefs
-
-  Memrel        :: "i=>i"
+definition
+  Memrel        :: "i=>i"  where
     "Memrel(A)   == {z: A*A . EX x y. z=<x,y> & x:y }"
 
-  Transset  :: "i=>o"
+definition
+  Transset  :: "i=>o"  where
     "Transset(i) == ALL x:i. x<=i"
 
-  Ord  :: "i=>o"
+definition
+  Ord  :: "i=>o"  where
     "Ord(i)      == Transset(i) & (ALL x:i. Transset(x))"
 
-  lt        :: "[i,i] => o"  (infixl "<" 50)   (*less-than on ordinals*)
+definition
+  lt        :: "[i,i] => o"  (infixl "<" 50)   (*less-than on ordinals*)  where
     "i<j         == i:j & Ord(j)"
 
-  Limit         :: "i=>o"
+definition
+  Limit         :: "i=>o"  where
     "Limit(i)    == Ord(i) & 0<i & (ALL y. y<i --> succ(y)<i)"
 
 abbreviation
@@ -736,134 +739,4 @@
 apply (blast intro!: equalityI)
 done
 
-ML 
-{*
-val Memrel_def = thm "Memrel_def";
-val Transset_def = thm "Transset_def";
-val Ord_def = thm "Ord_def";
-val lt_def = thm "lt_def";
-val Limit_def = thm "Limit_def";
-
-val Transset_iff_Pow = thm "Transset_iff_Pow";
-val Transset_iff_Union_succ = thm "Transset_iff_Union_succ";
-val Transset_iff_Union_subset = thm "Transset_iff_Union_subset";
-val Transset_doubleton_D = thm "Transset_doubleton_D";
-val Transset_Pair_D = thm "Transset_Pair_D";
-val Transset_includes_domain = thm "Transset_includes_domain";
-val Transset_includes_range = thm "Transset_includes_range";
-val Transset_0 = thm "Transset_0";
-val Transset_Un = thm "Transset_Un";
-val Transset_Int = thm "Transset_Int";
-val Transset_succ = thm "Transset_succ";
-val Transset_Pow = thm "Transset_Pow";
-val Transset_Union = thm "Transset_Union";
-val Transset_Union_family = thm "Transset_Union_family";
-val Transset_Inter_family = thm "Transset_Inter_family";
-val OrdI = thm "OrdI";
-val Ord_is_Transset = thm "Ord_is_Transset";
-val Ord_contains_Transset = thm "Ord_contains_Transset";
-val Ord_in_Ord = thm "Ord_in_Ord";
-val Ord_succD = thm "Ord_succD";
-val Ord_subset_Ord = thm "Ord_subset_Ord";
-val OrdmemD = thm "OrdmemD";
-val Ord_trans = thm "Ord_trans";
-val Ord_succ_subsetI = thm "Ord_succ_subsetI";
-val Ord_0 = thm "Ord_0";
-val Ord_succ = thm "Ord_succ";
-val Ord_1 = thm "Ord_1";
-val Ord_succ_iff = thm "Ord_succ_iff";
-val Ord_Un = thm "Ord_Un";
-val Ord_Int = thm "Ord_Int";
-val Ord_Inter = thm "Ord_Inter";
-val Ord_INT = thm "Ord_INT";
-val ON_class = thm "ON_class";
-val ltI = thm "ltI";
-val ltE = thm "ltE";
-val ltD = thm "ltD";
-val not_lt0 = thm "not_lt0";
-val lt_Ord = thm "lt_Ord";
-val lt_Ord2 = thm "lt_Ord2";
-val le_Ord2 = thm "le_Ord2";
-val lt0E = thm "lt0E";
-val lt_trans = thm "lt_trans";
-val lt_not_sym = thm "lt_not_sym";
-val lt_asym = thm "lt_asym";
-val lt_irrefl = thm "lt_irrefl";
-val lt_not_refl = thm "lt_not_refl";
-val le_iff = thm "le_iff";
-val leI = thm "leI";
-val le_eqI = thm "le_eqI";
-val le_refl = thm "le_refl";
-val le_refl_iff = thm "le_refl_iff";
-val leCI = thm "leCI";
-val leE = thm "leE";
-val le_anti_sym = thm "le_anti_sym";
-val le0_iff = thm "le0_iff";
-val le0D = thm "le0D";
-val Memrel_iff = thm "Memrel_iff";
-val MemrelI = thm "MemrelI";
-val MemrelE = thm "MemrelE";
-val Memrel_type = thm "Memrel_type";
-val Memrel_mono = thm "Memrel_mono";
-val Memrel_0 = thm "Memrel_0";
-val Memrel_1 = thm "Memrel_1";
-val wf_Memrel = thm "wf_Memrel";
-val trans_Memrel = thm "trans_Memrel";
-val Transset_Memrel_iff = thm "Transset_Memrel_iff";
-val Transset_induct = thm "Transset_induct";
-val Ord_induct = thm "Ord_induct";
-val trans_induct = thm "trans_induct";
-val Ord_linear = thm "Ord_linear";
-val Ord_linear_lt = thm "Ord_linear_lt";
-val Ord_linear2 = thm "Ord_linear2";
-val Ord_linear_le = thm "Ord_linear_le";
-val le_imp_not_lt = thm "le_imp_not_lt";
-val not_lt_imp_le = thm "not_lt_imp_le";
-val Ord_mem_iff_lt = thm "Ord_mem_iff_lt";
-val not_lt_iff_le = thm "not_lt_iff_le";
-val not_le_iff_lt = thm "not_le_iff_lt";
-val Ord_0_le = thm "Ord_0_le";
-val Ord_0_lt = thm "Ord_0_lt";
-val Ord_0_lt_iff = thm "Ord_0_lt_iff";
-val zero_le_succ_iff = thm "zero_le_succ_iff";
-val subset_imp_le = thm "subset_imp_le";
-val le_imp_subset = thm "le_imp_subset";
-val le_subset_iff = thm "le_subset_iff";
-val le_succ_iff = thm "le_succ_iff";
-val all_lt_imp_le = thm "all_lt_imp_le";
-val lt_trans1 = thm "lt_trans1";
-val lt_trans2 = thm "lt_trans2";
-val le_trans = thm "le_trans";
-val succ_leI = thm "succ_leI";
-val succ_leE = thm "succ_leE";
-val succ_le_iff = thm "succ_le_iff";
-val succ_le_imp_le = thm "succ_le_imp_le";
-val lt_subset_trans = thm "lt_subset_trans";
-val Un_upper1_le = thm "Un_upper1_le";
-val Un_upper2_le = thm "Un_upper2_le";
-val Un_least_lt = thm "Un_least_lt";
-val Un_least_lt_iff = thm "Un_least_lt_iff";
-val Un_least_mem_iff = thm "Un_least_mem_iff";
-val Int_greatest_lt = thm "Int_greatest_lt";
-val Ord_Union = thm "Ord_Union";
-val Ord_UN = thm "Ord_UN";
-val UN_least_le = thm "UN_least_le";
-val UN_succ_least_lt = thm "UN_succ_least_lt";
-val UN_upper_le = thm "UN_upper_le";
-val le_implies_UN_le_UN = thm "le_implies_UN_le_UN";
-val Ord_equality = thm "Ord_equality";
-val Ord_Union_subset = thm "Ord_Union_subset";
-val Limit_Union_eq = thm "Limit_Union_eq";
-val Limit_is_Ord = thm "Limit_is_Ord";
-val Limit_has_0 = thm "Limit_has_0";
-val Limit_has_succ = thm "Limit_has_succ";
-val non_succ_LimitI = thm "non_succ_LimitI";
-val succ_LimitE = thm "succ_LimitE";
-val not_succ_Limit = thm "not_succ_Limit";
-val Limit_le_succD = thm "Limit_le_succD";
-val Ord_cases_disj = thm "Ord_cases_disj";
-val Ord_cases = thm "Ord_cases";
-val trans_induct3 = thm "trans_induct3";
-*}
-
 end
--- a/src/ZF/Perm.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Perm.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -13,27 +13,30 @@
 
 theory Perm imports func begin
 
-constdefs
-
+definition
   (*composition of relations and functions; NOT Suppes's relative product*)
-  comp     :: "[i,i]=>i"      (infixr "O" 60)
+  comp     :: "[i,i]=>i"      (infixr "O" 60)  where
     "r O s == {xz : domain(s)*range(r) . 
                EX x y z. xz=<x,z> & <x,y>:s & <y,z>:r}"
 
+definition
   (*the identity function for A*)
-  id    :: "i=>i"
+  id    :: "i=>i"  where
     "id(A) == (lam x:A. x)"
 
+definition
   (*one-to-one functions from A to B*)
-  inj   :: "[i,i]=>i"
+  inj   :: "[i,i]=>i"  where
     "inj(A,B) == { f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x}"
 
+definition
   (*onto functions from A to B*)
-  surj  :: "[i,i]=>i"
+  surj  :: "[i,i]=>i"  where
     "surj(A,B) == { f: A->B . ALL y:B. EX x:A. f`x=y}"
 
+definition
   (*one-to-one and onto functions*)
-  bij   :: "[i,i]=>i"
+  bij   :: "[i,i]=>i"  where
     "bij(A,B) == inj(A,B) Int surj(A,B)"
 
 
@@ -547,91 +550,4 @@
 apply (force intro: apply_type  simp add: fun_extend)
 done
 
-
-ML
-{*
-val comp_def = thm "comp_def";
-val id_def = thm "id_def";
-val inj_def = thm "inj_def";
-val surj_def = thm "surj_def";
-val bij_def = thm "bij_def";
-
-val surj_is_fun = thm "surj_is_fun";
-val fun_is_surj = thm "fun_is_surj";
-val surj_range = thm "surj_range";
-val f_imp_surjective = thm "f_imp_surjective";
-val lam_surjective = thm "lam_surjective";
-val cantor_surj = thm "cantor_surj";
-val inj_is_fun = thm "inj_is_fun";
-val inj_equality = thm "inj_equality";
-val inj_apply_equality = thm "inj_apply_equality";
-val f_imp_injective = thm "f_imp_injective";
-val lam_injective = thm "lam_injective";
-val bij_is_inj = thm "bij_is_inj";
-val bij_is_surj = thm "bij_is_surj";
-val bij_is_fun = thm "bij_is_fun";
-val lam_bijective = thm "lam_bijective";
-val RepFun_bijective = thm "RepFun_bijective";
-val idI = thm "idI";
-val idE = thm "idE";
-val id_type = thm "id_type";
-val id_conv = thm "id_conv";
-val id_mono = thm "id_mono";
-val id_subset_inj = thm "id_subset_inj";
-val id_inj = thm "id_inj";
-val id_surj = thm "id_surj";
-val id_bij = thm "id_bij";
-val subset_iff_id = thm "subset_iff_id";
-val inj_converse_fun = thm "inj_converse_fun";
-val left_inverse = thm "left_inverse";
-val left_inverse_bij = thm "left_inverse_bij";
-val right_inverse = thm "right_inverse";
-val right_inverse_bij = thm "right_inverse_bij";
-val inj_converse_inj = thm "inj_converse_inj";
-val inj_converse_surj = thm "inj_converse_surj";
-val bij_converse_bij = thm "bij_converse_bij";
-val compI = thm "compI";
-val compE = thm "compE";
-val compEpair = thm "compEpair";
-val converse_comp = thm "converse_comp";
-val range_comp = thm "range_comp";
-val range_comp_eq = thm "range_comp_eq";
-val domain_comp = thm "domain_comp";
-val domain_comp_eq = thm "domain_comp_eq";
-val image_comp = thm "image_comp";
-val comp_mono = thm "comp_mono";
-val comp_rel = thm "comp_rel";
-val comp_assoc = thm "comp_assoc";
-val left_comp_id = thm "left_comp_id";
-val right_comp_id = thm "right_comp_id";
-val comp_function = thm "comp_function";
-val comp_fun = thm "comp_fun";
-val comp_fun_apply = thm "comp_fun_apply";
-val comp_lam = thm "comp_lam";
-val comp_inj = thm "comp_inj";
-val comp_surj = thm "comp_surj";
-val comp_bij = thm "comp_bij";
-val comp_mem_injD1 = thm "comp_mem_injD1";
-val comp_mem_injD2 = thm "comp_mem_injD2";
-val comp_mem_surjD1 = thm "comp_mem_surjD1";
-val comp_mem_surjD2 = thm "comp_mem_surjD2";
-val left_comp_inverse = thm "left_comp_inverse";
-val right_comp_inverse = thm "right_comp_inverse";
-val comp_eq_id_iff = thm "comp_eq_id_iff";
-val fg_imp_bijective = thm "fg_imp_bijective";
-val nilpotent_imp_bijective = thm "nilpotent_imp_bijective";
-val invertible_imp_bijective = thm "invertible_imp_bijective";
-val inj_disjoint_Un = thm "inj_disjoint_Un";
-val surj_disjoint_Un = thm "surj_disjoint_Un";
-val bij_disjoint_Un = thm "bij_disjoint_Un";
-val surj_image = thm "surj_image";
-val restrict_image = thm "restrict_image";
-val restrict_inj = thm "restrict_inj";
-val restrict_surj = thm "restrict_surj";
-val restrict_bij = thm "restrict_bij";
-val inj_weaken_type = thm "inj_weaken_type";
-val inj_succ_restrict = thm "inj_succ_restrict";
-val inj_extend = thm "inj_extend";
-*}
-
 end
--- a/src/ZF/QPair.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/QPair.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -21,45 +21,53 @@
 1966.
 *}
 
-constdefs
-  QPair     :: "[i, i] => i"                      ("<(_;/ _)>")
+definition
+  QPair     :: "[i, i] => i"                      ("<(_;/ _)>")  where
     "<a;b> == a+b"
 
-  qfst :: "i => i"
+definition
+  qfst :: "i => i"  where
     "qfst(p) == THE a. EX b. p=<a;b>"
 
-  qsnd :: "i => i"
+definition
+  qsnd :: "i => i"  where
     "qsnd(p) == THE b. EX a. p=<a;b>"
 
-  qsplit    :: "[[i, i] => 'a, i] => 'a::{}"  (*for pattern-matching*)
+definition
+  qsplit    :: "[[i, i] => 'a, i] => 'a::{}"  (*for pattern-matching*)  where
     "qsplit(c,p) == c(qfst(p), qsnd(p))"
 
-  qconverse :: "i => i"
+definition
+  qconverse :: "i => i"  where
     "qconverse(r) == {z. w:r, EX x y. w=<x;y> & z=<y;x>}"
 
-  QSigma    :: "[i, i => i] => i"
+definition
+  QSigma    :: "[i, i => i] => i"  where
     "QSigma(A,B)  ==  \<Union>x\<in>A. \<Union>y\<in>B(x). {<x;y>}"
 
 syntax
   "_QSUM"   :: "[idt, i, i] => i"               ("(3QSUM _:_./ _)" 10)
 translations
-  "QSUM x:A. B"  => "QSigma(A, %x. B)"
+  "QSUM x:A. B"  => "CONST QSigma(A, %x. B)"
 
 abbreviation
   qprod  (infixr "<*>" 80) where
   "A <*> B == QSigma(A, %_. B)"
 
-constdefs
-  qsum    :: "[i,i]=>i"                         (infixr "<+>" 65)
+definition
+  qsum    :: "[i,i]=>i"                         (infixr "<+>" 65)  where
     "A <+> B      == ({0} <*> A) Un ({1} <*> B)"
 
-  QInl :: "i=>i"
+definition
+  QInl :: "i=>i"  where
     "QInl(a)      == <0;a>"
 
-  QInr :: "i=>i"
+definition
+  QInr :: "i=>i"  where
     "QInr(b)      == <1;b>"
 
-  qcase     :: "[i=>i, i=>i, i]=>i"
+definition
+  qcase     :: "[i=>i, i=>i, i]=>i"  where
     "qcase(c,d)   == qsplit(%y z. cond(y, d(z), c(z)))"
 
 
@@ -316,71 +324,4 @@
 lemma qsum_mono: "[| A<=C;  B<=D |] ==> A <+> B <= C <+> D"
 by blast
 
-ML
-{*
-val qsum_defs = thms "qsum_defs";
-
-val QPair_empty = thm "QPair_empty";
-val QPair_iff = thm "QPair_iff";
-val QPair_inject = thm "QPair_inject";
-val QPair_inject1 = thm "QPair_inject1";
-val QPair_inject2 = thm "QPair_inject2";
-val QSigmaI = thm "QSigmaI";
-val QSigmaE = thm "QSigmaE";
-val QSigmaE = thm "QSigmaE";
-val QSigmaE2 = thm "QSigmaE2";
-val QSigmaD1 = thm "QSigmaD1";
-val QSigmaD2 = thm "QSigmaD2";
-val QSigma_cong = thm "QSigma_cong";
-val QSigma_empty1 = thm "QSigma_empty1";
-val QSigma_empty2 = thm "QSigma_empty2";
-val qfst_conv = thm "qfst_conv";
-val qsnd_conv = thm "qsnd_conv";
-val qfst_type = thm "qfst_type";
-val qsnd_type = thm "qsnd_type";
-val QPair_qfst_qsnd_eq = thm "QPair_qfst_qsnd_eq";
-val qsplit = thm "qsplit";
-val qsplit_type = thm "qsplit_type";
-val expand_qsplit = thm "expand_qsplit";
-val qsplitI = thm "qsplitI";
-val qsplitE = thm "qsplitE";
-val qsplitD = thm "qsplitD";
-val qconverseI = thm "qconverseI";
-val qconverseD = thm "qconverseD";
-val qconverseE = thm "qconverseE";
-val qconverse_qconverse = thm "qconverse_qconverse";
-val qconverse_type = thm "qconverse_type";
-val qconverse_prod = thm "qconverse_prod";
-val qconverse_empty = thm "qconverse_empty";
-val QInlI = thm "QInlI";
-val QInrI = thm "QInrI";
-val qsumE = thm "qsumE";
-val QInl_iff = thm "QInl_iff";
-val QInr_iff = thm "QInr_iff";
-val QInl_QInr_iff = thm "QInl_QInr_iff";
-val QInr_QInl_iff = thm "QInr_QInl_iff";
-val qsum_empty = thm "qsum_empty";
-val QInl_inject = thm "QInl_inject";
-val QInr_inject = thm "QInr_inject";
-val QInl_neq_QInr = thm "QInl_neq_QInr";
-val QInr_neq_QInl = thm "QInr_neq_QInl";
-val QInlD = thm "QInlD";
-val QInrD = thm "QInrD";
-val qsum_iff = thm "qsum_iff";
-val qsum_subset_iff = thm "qsum_subset_iff";
-val qsum_equal_iff = thm "qsum_equal_iff";
-val qcase_QInl = thm "qcase_QInl";
-val qcase_QInr = thm "qcase_QInr";
-val qcase_type = thm "qcase_type";
-val Part_QInl = thm "Part_QInl";
-val Part_QInr = thm "Part_QInr";
-val Part_QInr2 = thm "Part_QInr2";
-val Part_qsum_equality = thm "Part_qsum_equality";
-val QPair_mono = thm "QPair_mono";
-val QSigma_mono = thm "QSigma_mono";
-val QInl_mono = thm "QInl_mono";
-val QInr_mono = thm "QInr_mono";
-val qsum_mono = thm "qsum_mono";
-*}
-
 end
--- a/src/ZF/QUniv.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/QUniv.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -21,8 +21,8 @@
   induction	TrueI
   case_eqns	qcase_QInl qcase_QInr
 
-constdefs
-  quniv :: "i => i"
+definition
+  quniv :: "i => i"  where
    "quniv(A) == Pow(univ(eclose(A)))"
 
 
@@ -202,46 +202,4 @@
         add: Limit_Vfrom_eq Int_UN_distrib UN_mono QPair_Int_Vset_subset_trans)
 done
 
-ML
-{*
-val Transset_includes_summands = thm "Transset_includes_summands";
-val Transset_sum_Int_subset = thm "Transset_sum_Int_subset";
-val qunivI = thm "qunivI";
-val qunivD = thm "qunivD";
-val quniv_mono = thm "quniv_mono";
-val univ_eclose_subset_quniv = thm "univ_eclose_subset_quniv";
-val univ_subset_quniv = thm "univ_subset_quniv";
-val univ_into_quniv = thm "univ_into_quniv";
-val Pow_univ_subset_quniv = thm "Pow_univ_subset_quniv";
-val univ_subset_into_quniv = thm "univ_subset_into_quniv";
-val zero_in_quniv = thm "zero_in_quniv";
-val one_in_quniv = thm "one_in_quniv";
-val two_in_quniv = thm "two_in_quniv";
-val A_subset_quniv = thm "A_subset_quniv";
-val A_into_quniv = thm "A_into_quniv";
-val QPair_subset_univ = thm "QPair_subset_univ";
-val QInl_subset_univ = thm "QInl_subset_univ";
-val naturals_subset_nat = thm "naturals_subset_nat";
-val naturals_subset_univ = thm "naturals_subset_univ";
-val QInr_subset_univ = thm "QInr_subset_univ";
-val QPair_in_quniv = thm "QPair_in_quniv";
-val QSigma_quniv = thm "QSigma_quniv";
-val QSigma_subset_quniv = thm "QSigma_subset_quniv";
-val quniv_QPair_D = thm "quniv_QPair_D";
-val quniv_QPair_E = thm "quniv_QPair_E";
-val quniv_QPair_iff = thm "quniv_QPair_iff";
-val QInl_in_quniv = thm "QInl_in_quniv";
-val QInr_in_quniv = thm "QInr_in_quniv";
-val qsum_quniv = thm "qsum_quniv";
-val qsum_subset_quniv = thm "qsum_subset_quniv";
-val nat_subset_quniv = thm "nat_subset_quniv";
-val nat_into_quniv = thm "nat_into_quniv";
-val bool_subset_quniv = thm "bool_subset_quniv";
-val bool_into_quniv = thm "bool_into_quniv";
-val QPair_Int_Vfrom_succ_subset = thm "QPair_Int_Vfrom_succ_subset";
-val QPair_Int_Vfrom_subset = thm "QPair_Int_Vfrom_subset";
-val QPair_Int_Vset_subset_trans = thm "QPair_Int_Vset_subset_trans";
-val QPair_Int_Vset_subset_UN = thm "QPair_Int_Vset_subset_UN";
-*}
-
 end
--- a/src/ZF/Resid/Confluence.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Resid/Confluence.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -7,12 +7,13 @@
 
 theory Confluence imports Reduction begin
 
-constdefs
-  confluence    :: "i=>o"
+definition
+  confluence    :: "i=>o"  where
     "confluence(R) ==   
        \<forall>x y. <x,y> \<in> R --> (\<forall>z.<x,z> \<in> R --> (\<exists>u.<y,u> \<in> R & <z,u> \<in> R))"
 
-  strip         :: "o"
+definition
+  strip         :: "o"  where
     "strip == \<forall>x y. (x ===> y) --> 
                     (\<forall>z.(x =1=> z) --> (\<exists>u.(y =1=> u) & (z===>u)))" 
 
--- a/src/ZF/Resid/Residuals.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Resid/Residuals.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -10,7 +10,6 @@
 
 consts
   Sres          :: "i"
-  res_func      :: "[i,i]=>i"     (infixl "|>" 70)
 
 abbreviation
   "residuals(u,v,w) == <u,v,w> \<in> Sres"
@@ -29,8 +28,9 @@
                  residuals(App(b,Fun(u1),u2),App(1,Fun(v1),v2),w2/w1)"
   type_intros    subst_type nat_typechecks redexes.intros bool_typechecks
 
-defs
-  res_func_def:  "u |> v == THE w. residuals(u,v,w)"
+definition
+  res_func      :: "[i,i]=>i"     (infixl "|>" 70)  where
+  "u |> v == THE w. residuals(u,v,w)"
 
 
 subsection{*Setting up rule lists*}
--- a/src/ZF/Resid/Substitution.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Resid/Substitution.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -19,8 +19,8 @@
 
   "lift_aux(App(b,f,a)) = (\<lambda>k \<in> nat. App(b, lift_aux(f)`k, lift_aux(a)`k))"
 
-constdefs
-  lift_rec      :: "[i,i]=> i"
+definition
+  lift_rec      :: "[i,i]=> i"  where
     "lift_rec(r,k) == lift_aux(r)`k"
 
 abbreviation
@@ -40,8 +40,8 @@
   "subst_aux(App(b,f,a)) =
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. App(b, subst_aux(f)`r`k, subst_aux(a)`r`k))"
 
-constdefs
-  subst_rec     :: "[i,i,i]=> i"	(**NOTE THE ARGUMENT ORDER BELOW**)
+definition
+  subst_rec     :: "[i,i,i]=> i"	(**NOTE THE ARGUMENT ORDER BELOW**)  where
     "subst_rec(u,r,k) == subst_aux(r)`u`k"
 
 abbreviation
--- a/src/ZF/Sum.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Sum.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -196,55 +196,4 @@
 lemma Part_sum_equality: "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"
 by blast
 
-ML
-{*
-val sum_def = thm "sum_def";
-val Inl_def = thm "Inl_def";
-val Inr_def = thm "Inr_def";
-val sum_defs = thms "sum_defs";
-
-val Part_iff = thm "Part_iff";
-val Part_eqI = thm "Part_eqI";
-val PartI = thm "PartI";
-val PartE = thm "PartE";
-val Part_subset = thm "Part_subset";
-val Sigma_bool = thm "Sigma_bool";
-val InlI = thm "InlI";
-val InrI = thm "InrI";
-val sumE = thm "sumE";
-val Inl_iff = thm "Inl_iff";
-val Inr_iff = thm "Inr_iff";
-val Inl_Inr_iff = thm "Inl_Inr_iff";
-val Inr_Inl_iff = thm "Inr_Inl_iff";
-val sum_empty = thm "sum_empty";
-val Inl_inject = thm "Inl_inject";
-val Inr_inject = thm "Inr_inject";
-val Inl_neq_Inr = thm "Inl_neq_Inr";
-val Inr_neq_Inl = thm "Inr_neq_Inl";
-val InlD = thm "InlD";
-val InrD = thm "InrD";
-val sum_iff = thm "sum_iff";
-val sum_subset_iff = thm "sum_subset_iff";
-val sum_equal_iff = thm "sum_equal_iff";
-val sum_eq_2_times = thm "sum_eq_2_times";
-val case_Inl = thm "case_Inl";
-val case_Inr = thm "case_Inr";
-val case_type = thm "case_type";
-val expand_case = thm "expand_case";
-val case_cong = thm "case_cong";
-val case_case = thm "case_case";
-val Part_mono = thm "Part_mono";
-val Part_Collect = thm "Part_Collect";
-val Part_CollectE = thm "Part_CollectE";
-val Part_Inl = thm "Part_Inl";
-val Part_Inr = thm "Part_Inr";
-val PartD1 = thm "PartD1";
-val Part_id = thm "Part_id";
-val Part_Inr2 = thm "Part_Inr2";
-val Part_sum_equality = thm "Part_sum_equality";
-
-*}
-
-
-
 end
--- a/src/ZF/Tools/datatype_package.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Tools/datatype_package.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -310,7 +310,7 @@
                          args)),
              subst_rec (Term.betapplys (recursor_case, args))));
 
-        val recursor_trans = recursor_def RS def_Vrecursor RS trans;
+        val recursor_trans = recursor_def RS @{thm def_Vrecursor} RS trans;
 
         fun prove_recursor_eqn arg =
           Goal.prove_global thy1 [] []
--- a/src/ZF/Tools/inductive_package.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Tools/inductive_package.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -189,8 +189,8 @@
   val bnd_mono =
     Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs))
       (fn _ => EVERY
-        [rtac (Collect_subset RS bnd_monoI) 1,
-         REPEAT (ares_tac (basic_monos @ monos) 1)]);
+        [rtac (@{thm Collect_subset} RS @{thm bnd_monoI}) 1,
+         REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
 
   val dom_subset = standard (big_rec_def RS Fp.subs);
 
@@ -204,19 +204,19 @@
   val Part_trans =
       case rec_names of
            [_] => asm_rl
-         | _   => standard (Part_subset RS subset_trans);
+         | _   => standard (@{thm Part_subset} RS @{thm subset_trans});
 
   (*To type-check recursive occurrences of the inductive sets, possibly
     enclosed in some monotonic operator M.*)
   val rec_typechecks =
      [dom_subset] RL (asm_rl :: ([Part_trans] RL monos))
-     RL [subsetD];
+     RL [@{thm subsetD}];
 
   (*Type-checking is hardest aspect of proof;
     disjIn selects the correct disjunct after unfolding*)
   fun intro_tacsf disjIn =
     [DETERM (stac unfold 1),
-     REPEAT (resolve_tac [Part_eqI,CollectI] 1),
+     REPEAT (resolve_tac [@{thm Part_eqI}, @{thm CollectI}] 1),
      (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*)
      rtac disjIn 2,
      (*Not ares_tac, since refl must be tried before equality assumptions;
@@ -229,12 +229,12 @@
      if !Ind_Syntax.trace then print_tac "The type-checking subgoal:"
      else all_tac,
      REPEAT (FIRSTGOAL (        dresolve_tac rec_typechecks
-                        ORELSE' eresolve_tac (asm_rl::PartE::SigmaE2::
+                        ORELSE' eresolve_tac (asm_rl::@{thm PartE}::@{thm SigmaE2}::
                                               type_elims)
                         ORELSE' hyp_subst_tac)),
      if !Ind_Syntax.trace then print_tac "The subgoal after monos, type_elims:"
      else all_tac,
-     DEPTH_SOLVE (swap_res_tac (SigmaI::subsetI::type_intrs) 1)];
+     DEPTH_SOLVE (swap_res_tac (@{thm SigmaI}::@{thm subsetI}::type_intrs) 1)];
 
   (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*)
   val mk_disj_rls = BalancedTree.accesses
@@ -341,12 +341,12 @@
             rtac (impI RS allI) 1,
             DETERM (etac raw_induct 1),
             (*Push Part inside Collect*)
-            full_simp_tac (min_ss addsimps [Part_Collect]) 1,
+            full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1,
             (*This CollectE and disjE separates out the introduction rules*)
-            REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE])),
+            REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, disjE])),
             (*Now break down the individual cases.  No disjE here in case
               some premise involves disjunction.*)
-            REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE]
+            REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, exE, conjE]
                                ORELSE' bound_hyp_subst_tac)),
             ind_tac (rev (map (rewrite_rule part_rec_defs) prems)) (length prems)]);
 
@@ -409,8 +409,8 @@
              else ();
 
 
-     val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp],
-                             resolve_tac [allI, impI, conjI, Part_eqI],
+     val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, @{thm PartE}, mp],
+                             resolve_tac [allI, impI, conjI, @{thm Part_eqI}],
                              dresolve_tac [spec, mp, Pr.fsplitD]];
 
      val need_mutual = length rec_names > 1;
@@ -445,8 +445,8 @@
        where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
        Instead the following rules extract the relevant conjunct.
      *)
-     val cmonos = [subset_refl RS Collect_mono] RL monos
-                   RLN (2,[rev_subsetD]);
+     val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos
+                   RLN (2,[@{thm rev_subsetD}]);
 
      (*Minimizes backtracking by delivering the correct premise to each goal*)
      fun mutual_ind_tac [] 0 = all_tac
@@ -458,7 +458,7 @@
                   using freeness of the Sum constructors; proves all but one
                   conjunct by contradiction*)
                 rewrite_goals_tac all_defs  THEN
-                simp_tac (mut_ss addsimps [Part_iff]) 1  THEN
+                simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1  THEN
                 IF_UNSOLVED (*simp_tac may have finished it off!*)
                   ((*simplify assumptions*)
                    (*some risk of excessive simplification here -- might have
--- a/src/ZF/Tools/typechk.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Tools/typechk.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -94,7 +94,7 @@
 
 (*Compiles a term-net for speed*)
 val basic_res_tac = net_resolve_tac [TrueI,refl,reflexive_thm,iff_refl,
-                                     ballI,allI,conjI,impI];
+                                     @{thm ballI},allI,conjI,impI];
 
 (*Instantiates variables in typing conditions.
   drawback: does not simplify conjunctions*)
--- a/src/ZF/Trancl.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Trancl.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,36 +9,45 @@
 
 theory Trancl imports Fixedpt Perm begin
 
-constdefs
-  refl     :: "[i,i]=>o"
+definition
+  refl     :: "[i,i]=>o"  where
     "refl(A,r) == (ALL x: A. <x,x> : r)"
 
-  irrefl   :: "[i,i]=>o"
+definition
+  irrefl   :: "[i,i]=>o"  where
     "irrefl(A,r) == ALL x: A. <x,x> ~: r"
 
-  sym      :: "i=>o"
+definition
+  sym      :: "i=>o"  where
     "sym(r) == ALL x y. <x,y>: r --> <y,x>: r"
 
-  asym     :: "i=>o"
+definition
+  asym     :: "i=>o"  where
     "asym(r) == ALL x y. <x,y>:r --> ~ <y,x>:r"
 
-  antisym  :: "i=>o"
+definition
+  antisym  :: "i=>o"  where
     "antisym(r) == ALL x y.<x,y>:r --> <y,x>:r --> x=y"
 
-  trans    :: "i=>o"
+definition
+  trans    :: "i=>o"  where
     "trans(r) == ALL x y z. <x,y>: r --> <y,z>: r --> <x,z>: r"
 
-  trans_on :: "[i,i]=>o"  ("trans[_]'(_')")
+definition
+  trans_on :: "[i,i]=>o"  ("trans[_]'(_')")  where
     "trans[A](r) == ALL x:A. ALL y:A. ALL z:A.       
                           <x,y>: r --> <y,z>: r --> <x,z>: r"
 
-  rtrancl :: "i=>i"  ("(_^*)" [100] 100)  (*refl/transitive closure*)
+definition
+  rtrancl :: "i=>i"  ("(_^*)" [100] 100)  (*refl/transitive closure*)  where
     "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))"
 
-  trancl  :: "i=>i"  ("(_^+)" [100] 100)  (*transitive closure*)
+definition
+  trancl  :: "i=>i"  ("(_^+)" [100] 100)  (*transitive closure*)  where
     "r^+ == r O r^*"
 
-  equiv    :: "[i,i]=>o"
+definition
+  equiv    :: "[i,i]=>o"  where
     "equiv(A,r) == r <= A*A & refl(A,r) & sym(r) & trans(r)"
 
 
@@ -364,59 +373,4 @@
 apply (auto simp add: trancl_converse)
 done
 
-ML
-{*
-val refl_def = thm "refl_def";
-val irrefl_def = thm "irrefl_def";
-val equiv_def = thm "equiv_def";
-val sym_def = thm "sym_def";
-val asym_def = thm "asym_def";
-val antisym_def = thm "antisym_def";
-val trans_def = thm "trans_def";
-val trans_on_def = thm "trans_on_def";
-
-val irreflI = thm "irreflI";
-val symI = thm "symI";
-val symI = thm "symI";
-val antisymI = thm "antisymI";
-val antisymE = thm "antisymE";
-val transD = thm "transD";
-val trans_onD = thm "trans_onD";
-
-val rtrancl_bnd_mono = thm "rtrancl_bnd_mono";
-val rtrancl_mono = thm "rtrancl_mono";
-val rtrancl_unfold = thm "rtrancl_unfold";
-val rtrancl_type = thm "rtrancl_type";
-val rtrancl_refl = thm "rtrancl_refl";
-val rtrancl_into_rtrancl = thm "rtrancl_into_rtrancl";
-val r_into_rtrancl = thm "r_into_rtrancl";
-val r_subset_rtrancl = thm "r_subset_rtrancl";
-val rtrancl_field = thm "rtrancl_field";
-val rtrancl_full_induct = thm "rtrancl_full_induct";
-val rtrancl_induct = thm "rtrancl_induct";
-val trans_rtrancl = thm "trans_rtrancl";
-val rtrancl_trans = thm "rtrancl_trans";
-val rtranclE = thm "rtranclE";
-val trans_trancl = thm "trans_trancl";
-val trancl_trans = thm "trancl_trans";
-val trancl_into_rtrancl = thm "trancl_into_rtrancl";
-val r_into_trancl = thm "r_into_trancl";
-val r_subset_trancl = thm "r_subset_trancl";
-val rtrancl_into_trancl1 = thm "rtrancl_into_trancl1";
-val rtrancl_into_trancl2 = thm "rtrancl_into_trancl2";
-val trancl_induct = thm "trancl_induct";
-val tranclE = thm "tranclE";
-val trancl_type = thm "trancl_type";
-val trancl_mono = thm "trancl_mono";
-val rtrancl_idemp = thm "rtrancl_idemp";
-val rtrancl_subset = thm "rtrancl_subset";
-val rtrancl_converseD = thm "rtrancl_converseD";
-val rtrancl_converseI = thm "rtrancl_converseI";
-val rtrancl_converse = thm "rtrancl_converse";
-val trancl_converseD = thm "trancl_converseD";
-val trancl_converseI = thm "trancl_converseI";
-val trancl_converse = thm "trancl_converse";
-val converse_trancl_induct = thm "converse_trancl_induct";
-*}
-
 end
--- a/src/ZF/UNITY/AllocBase.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/AllocBase.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -45,25 +45,23 @@
   "all_distinct0(Cons(a, l)) =
      (if a \<in> set_of_list(l) then 0 else all_distinct0(l))"
 
-constdefs
-  all_distinct  :: "i=>o"
+definition
+  all_distinct  :: "i=>o"  where
    "all_distinct(l) == all_distinct0(l)=1"
   
-constdefs  
-  state_of :: "i =>i" --{* coersion from anyting to state *}
+definition  
+  state_of :: "i =>i" --{* coersion from anyting to state *}  where
    "state_of(s) == if s \<in> state then s else st0"
 
-  lift :: "i =>(i=>i)" --{* simplifies the expression of programs*}
+definition
+  lift :: "i =>(i=>i)" --{* simplifies the expression of programs*}  where
    "lift(x) == %s. s`x"
 
 text{* function to show that the set of variables is infinite *}
 consts
   nat_list_inj :: "i=>i"
-  nat_var_inj  :: "i=>i"
   var_inj      :: "i=>i"
 
-defs
-  nat_var_inj_def: "nat_var_inj(n) == Var(nat_list_inj(n))"
 primrec
   "nat_list_inj(0) = Nil"
   "nat_list_inj(succ(n)) = Cons(n, nat_list_inj(n))"
@@ -71,6 +69,10 @@
 primrec
   "var_inj(Var(l)) = length(l)"
 
+definition
+  nat_var_inj  :: "i=>i"  where
+  "nat_var_inj(n) == Var(nat_list_inj(n))"
+
 
 subsection{*Various simple lemmas*}
 
--- a/src/ZF/UNITY/AllocImpl.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/AllocImpl.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -24,8 +24,7 @@
   alloc_default_val_assumes [simp]:
   "default_val(NbR)  = 0 & default_val(available_tok)=0"
 
-constdefs
-  alloc_giv_act :: i
+definition
   "alloc_giv_act ==
        {<s, t> \<in> state*state.
 	\<exists>k. k = length(s`giv) &
@@ -33,16 +32,16 @@
 		  available_tok := s`available_tok #- nth(k, s`ask)) &
 	    k < length(s`ask) & nth(k, s`ask) le s`available_tok}"
 
-  alloc_rel_act :: i
+definition
   "alloc_rel_act ==
        {<s, t> \<in> state*state.
         t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
 	      NbR := succ(s`NbR)) &
   	s`NbR < length(s`rel)}"
 
+definition
   (*The initial condition s`giv=[] is missing from the
     original definition: S. O. Ehmety *)
-  alloc_prog :: i
   "alloc_prog ==
        mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil},
 		  {alloc_giv_act, alloc_rel_act},
@@ -659,6 +658,4 @@
 apply (auto simp add: alloc_progress_def)
 done
 
-
-
 end
--- a/src/ZF/UNITY/ClientImpl.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/ClientImpl.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -24,10 +24,8 @@
 
 (*Array indexing is translated to list indexing as A[n] == nth(n-1,A). *)
 
-constdefs
+definition
  (** Release some client_tokens **)
-  
-  client_rel_act ::i
     "client_rel_act ==
      {<s,t> \<in> state*state.
       \<exists>nrel \<in> nat. nrel = length(s`rel) &
@@ -38,15 +36,14 @@
   (** Choose a new token requirement **)
   (** Including t=s suppresses fairness, allowing the non-trivial part
       of the action to be ignored **)
-
-  client_tok_act :: i
- "client_tok_act == {<s,t> \<in> state*state. t=s |
+definition
+  "client_tok_act == {<s,t> \<in> state*state. t=s |
 		     t = s(tok:=succ(s`tok mod NbT))}"
 
-  client_ask_act :: i
+definition
   "client_ask_act == {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
   
-  client_prog :: i
+definition
   "client_prog ==
    mk_program({s \<in> state. s`tok \<le> NbT & s`giv = Nil &
 	               s`ask = Nil & s`rel = Nil},
--- a/src/ZF/UNITY/Comp.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Comp.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -18,29 +18,34 @@
 
 theory Comp imports Union Increasing begin
 
-constdefs
-
-  component :: "[i,i]=>o"  (infixl "component" 65) 
+definition
+  component :: "[i,i]=>o"  (infixl "component" 65)  where
   "F component H == (EX G. F Join G = H)"
 
-  strict_component :: "[i,i]=>o" (infixl "strict'_component" 65)
+definition
+  strict_component :: "[i,i]=>o" (infixl "strict'_component" 65)  where
   "F strict_component H == F component H & F~=H"
 
+definition
   (* A stronger form of the component relation *)
-  component_of :: "[i,i]=>o"   (infixl "component'_of" 65)
+  component_of :: "[i,i]=>o"   (infixl "component'_of" 65)  where
   "F component_of H  == (EX G. F ok G & F Join G = H)"
   
-  strict_component_of :: "[i,i]=>o" (infixl "strict'_component'_of" 65)
+definition
+  strict_component_of :: "[i,i]=>o" (infixl "strict'_component'_of" 65)  where
   "F strict_component_of H  == F component_of H  & F~=H"
 
+definition
   (* Preserves a state function f, in particular a variable *)
-  preserves :: "(i=>i)=>i"
+  preserves :: "(i=>i)=>i"  where
   "preserves(f) == {F:program. ALL z. F: stable({s:state. f(s) = z})}"
 
-  fun_pair :: "[i=>i, i =>i] =>(i=>i)"
+definition
+  fun_pair :: "[i=>i, i =>i] =>(i=>i)"  where
   "fun_pair(f, g) == %x. <f(x), g(x)>"
 
-  localize  :: "[i=>i, i] => i"
+definition
+  localize  :: "[i=>i, i] => i"  where
   "localize(f,F) == mk_program(Init(F), Acts(F),
 		       AllowedActs(F) Int (UN G:preserves(f). Acts(G)))"
 
--- a/src/ZF/UNITY/Constrains.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Constrains.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -7,8 +7,8 @@
 
 theory Constrains
 imports UNITY
+begin
 
-begin
 consts traces :: "[i, i] => i"
   (* Initial states and program => (final state, reversed trace to it)... 
       the domain may also be state*list(state) *)
@@ -39,23 +39,22 @@
   type_intros UnI1 UnI2 fieldI2 UN_I
 
   
-consts
-  Constrains :: "[i,i] => i"  (infixl "Co"     60)
-  op_Unless  :: "[i, i] => i"  (infixl "Unless" 60)
+definition
+  Constrains :: "[i,i] => i"  (infixl "Co" 60)  where
+  "A Co B == {F:program. F:(reachable(F) Int A) co B}"
 
-defs
-  Constrains_def:
-    "A Co B == {F:program. F:(reachable(F) Int A) co B}"
+definition
+  op_Unless  :: "[i, i] => i"  (infixl "Unless" 60)  where
+  "A Unless B == (A-B) Co (A Un B)"
 
-  Unless_def:
-    "A Unless B == (A-B) Co (A Un B)"
+definition
+  Stable     :: "i => i"  where
+  "Stable(A) == A Co A"
 
-constdefs
-  Stable     :: "i => i"
-    "Stable(A) == A Co A"
+definition
   (*Always is the weak form of "invariant"*)
-  Always :: "i => i"
-    "Always(A) == initially(A) Int Stable(A)"
+  Always :: "i => i"  where
+  "Always(A) == initially(A) Int Stable(A)"
 
 
 (*** traces and reachable ***)
@@ -330,8 +329,7 @@
 (** Unless **)
 
 lemma Unless_type: "A Unless B <=program"
-
-apply (unfold Unless_def)
+apply (unfold op_Unless_def)
 apply (rule Constrains_type)
 done
 
@@ -457,80 +455,11 @@
 
 ML
 {*
-val reachable_type = thm "reachable_type";
-val st_set_reachable = thm "st_set_reachable";
-val reachable_Int_state = thm "reachable_Int_state";
-val state_Int_reachable = thm "state_Int_reachable";
-val reachable_equiv_traces = thm "reachable_equiv_traces";
-val Init_into_reachable = thm "Init_into_reachable";
-val stable_reachable = thm "stable_reachable";
-val invariant_reachable = thm "invariant_reachable";
-val invariant_includes_reachable = thm "invariant_includes_reachable";
-val constrains_reachable_Int = thm "constrains_reachable_Int";
-val Constrains_eq_constrains = thm "Constrains_eq_constrains";
-val Constrains_def2 = thm "Constrains_def2";
-val constrains_imp_Constrains = thm "constrains_imp_Constrains";
-val ConstrainsI = thm "ConstrainsI";
-val Constrains_type = thm "Constrains_type";
-val Constrains_empty = thm "Constrains_empty";
-val Constrains_state = thm "Constrains_state";
-val Constrains_weaken_R = thm "Constrains_weaken_R";
-val Constrains_weaken_L = thm "Constrains_weaken_L";
-val Constrains_weaken = thm "Constrains_weaken";
-val Constrains_Un = thm "Constrains_Un";
-val Constrains_UN = thm "Constrains_UN";
-val Constrains_Int = thm "Constrains_Int";
-val Constrains_INT = thm "Constrains_INT";
-val Constrains_imp_subset = thm "Constrains_imp_subset";
-val Constrains_trans = thm "Constrains_trans";
-val Constrains_cancel = thm "Constrains_cancel";
-val stable_imp_Stable = thm "stable_imp_Stable";
-val Stable_eq = thm "Stable_eq";
-val Stable_eq_stable = thm "Stable_eq_stable";
-val StableI = thm "StableI";
-val StableD = thm "StableD";
-val Stable_Un = thm "Stable_Un";
-val Stable_Int = thm "Stable_Int";
-val Stable_Constrains_Un = thm "Stable_Constrains_Un";
-val Stable_Constrains_Int = thm "Stable_Constrains_Int";
-val Stable_UN = thm "Stable_UN";
-val Stable_INT = thm "Stable_INT";
-val Stable_reachable = thm "Stable_reachable";
-val Stable_type = thm "Stable_type";
-val Elimination = thm "Elimination";
-val Elimination2 = thm "Elimination2";
-val Unless_type = thm "Unless_type";
-val AlwaysI = thm "AlwaysI";
-val AlwaysD = thm "AlwaysD";
-val AlwaysE = thm "AlwaysE";
-val Always_imp_Stable = thm "Always_imp_Stable";
-val Always_includes_reachable = thm "Always_includes_reachable";
-val invariant_imp_Always = thm "invariant_imp_Always";
-val Always_reachable = thm "Always_reachable";
-val Always_eq_invariant_reachable = thm "Always_eq_invariant_reachable";
-val Always_eq_includes_reachable = thm "Always_eq_includes_reachable";
-val Always_type = thm "Always_type";
-val Always_state_eq = thm "Always_state_eq";
-val state_AlwaysI = thm "state_AlwaysI";
-val Always_eq_UN_invariant = thm "Always_eq_UN_invariant";
-val Always_weaken = thm "Always_weaken";
-val Int_absorb2 = thm "Int_absorb2";
-val Always_Constrains_pre = thm "Always_Constrains_pre";
-val Always_Constrains_post = thm "Always_Constrains_post";
-val Always_ConstrainsI = thm "Always_ConstrainsI";
-val Always_ConstrainsD = thm "Always_ConstrainsD";
-val Always_Constrains_weaken = thm "Always_Constrains_weaken";
-val Always_Int_distrib = thm "Always_Int_distrib";
-val Always_INT_distrib = thm "Always_INT_distrib";
-val Always_Int_I = thm "Always_Int_I";
-val Always_Diff_Un_eq = thm "Always_Diff_Un_eq";
-val Always_thin = thm "Always_thin";
-
 (*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
-val Always_Int_tac = dtac Always_Int_I THEN' assume_tac THEN' etac Always_thin;
+val Always_Int_tac = dtac @{thm Always_Int_I} THEN' assume_tac THEN' etac @{thm Always_thin};
 
 (*Combines a list of invariance THEOREMS into one.*)
-val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS Always_Int_I);
+val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I});
 
 (*To allow expansion of the program's definition when appropriate*)
 structure ProgramDefs = NamedThmsFun(val name = "program" val description = "program definitions");
@@ -541,13 +470,13 @@
   let val css as (cs, ss) = local_clasimpset_of ctxt in
    SELECT_GOAL
       (EVERY [REPEAT (Always_Int_tac 1),
-              REPEAT (etac Always_ConstrainsI 1
+              REPEAT (etac @{thm Always_ConstrainsI} 1
                       ORELSE
-                      resolve_tac [StableI, stableI,
-                                   constrains_imp_Constrains] 1),
-              rtac constrainsI 1,
+                      resolve_tac [@{thm StableI}, @{thm stableI},
+                                   @{thm constrains_imp_Constrains}] 1),
+              rtac @{thm constrainsI} 1,
               (* Three subgoals *)
-              rewrite_goal_tac [st_set_def] 3,
+              rewrite_goal_tac [@{thm st_set_def}] 3,
               REPEAT (force_tac css 2),
               full_simp_tac (ss addsimps (ProgramDefs.get ctxt)) 1,
               ALLGOALS (clarify_tac cs),
@@ -561,7 +490,7 @@
 
 (*For proving invariants*)
 fun always_tac ctxt i = 
-    rtac AlwaysI i THEN force_tac (local_clasimpset_of ctxt) i THEN constrains_tac ctxt i;
+    rtac @{thm AlwaysI} i THEN force_tac (local_clasimpset_of ctxt) i THEN constrains_tac ctxt i;
 *}
 
 setup ProgramDefs.setup
--- a/src/ZF/UNITY/Distributor.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Distributor.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -12,8 +12,9 @@
 text{*Distributor specification (the number of outputs is Nclients)*}
 
 text{*spec (14)*}
-constdefs
-  distr_follows :: "[i, i, i, i =>i] =>i"
+
+definition
+  distr_follows :: "[i, i, i, i =>i] =>i"  where
     "distr_follows(A, In, iIn, Out) ==
      (lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
      (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
@@ -25,12 +26,14 @@
           (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) & nth(k, s`iIn) = n}))
 	  Wrt prefix(A)/list(A))"
 
-  distr_allowed_acts :: "[i=>i]=>i"
+definition
+  distr_allowed_acts :: "[i=>i]=>i"  where
     "distr_allowed_acts(Out) ==
      {D \<in> program. AllowedActs(D) =
      cons(id(state), \<Union>G \<in> (\<Inter>n\<in>nat. preserves(lift(Out(n)))). Acts(G))}"
 
-  distr_spec :: "[i, i, i, i =>i]=>i"
+definition
+  distr_spec :: "[i, i, i, i =>i]=>i"  where
     "distr_spec(A, In, iIn, Out) ==
      distr_follows(A, In, iIn, Out) \<inter> distr_allowed_acts(Out)"
 
--- a/src/ZF/UNITY/FP.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/FP.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -12,11 +12,12 @@
 
 theory FP imports UNITY begin
 
-constdefs   
-  FP_Orig :: "i=>i"
+definition   
+  FP_Orig :: "i=>i"  where
     "FP_Orig(F) == Union({A \<in> Pow(state). \<forall>B. F \<in> stable(A Int B)})"
 
-  FP :: "i=>i"
+definition
+  FP :: "i=>i"  where
     "FP(F) == {s\<in>state. F \<in> stable({s})}"
 
 
@@ -80,20 +81,4 @@
       ==> A-FP(F) = (\<Union>act \<in> Acts(F). A - {s \<in> state. act``{s} \<subseteq> {s}})"
 by (unfold FP_def stable_def constrains_def st_set_def, blast)
 
-ML
-{*
-val FP_Orig_type = thm "FP_Orig_type";
-val st_set_FP_Orig = thm "st_set_FP_Orig";
-val FP_type = thm "FP_type";
-val st_set_FP = thm "st_set_FP";
-val stable_FP_Orig_Int = thm "stable_FP_Orig_Int";
-val FP_Orig_weakest2 = thm "FP_Orig_weakest2";
-val stable_FP_Int = thm "stable_FP_Int";
-val FP_subset_FP_Orig = thm "FP_subset_FP_Orig";
-val FP_Orig_subset_FP = thm "FP_Orig_subset_FP";
-val FP_equivalence = thm "FP_equivalence";
-val FP_weakest = thm "FP_weakest";
-val Diff_FP = thm "Diff_FP";
-*}
-
 end
--- a/src/ZF/UNITY/Follows.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Follows.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -10,8 +10,8 @@
 
 theory Follows imports SubstAx Increasing begin
 
-constdefs
-  Follows :: "[i, i, i=>i, i=>i] => i"
+definition
+  Follows :: "[i, i, i=>i, i=>i] => i"  where
   "Follows(A, r, f, g) ==
             Increasing(A, r, g) Int
             Increasing(A, r,f) Int
@@ -473,58 +473,4 @@
 apply (rule Follows_munion, auto)
 done
 
-ML
-{*
-val Follows_cong = thm "Follows_cong";
-val subset_Always_comp = thm "subset_Always_comp";
-val imp_Always_comp = thm "imp_Always_comp";
-val imp_Always_comp2 = thm "imp_Always_comp2";
-val subset_LeadsTo_comp = thm "subset_LeadsTo_comp";
-val imp_LeadsTo_comp = thm "imp_LeadsTo_comp";
-val imp_LeadsTo_comp_right = thm "imp_LeadsTo_comp_right";
-val imp_LeadsTo_comp_left = thm "imp_LeadsTo_comp_left";
-val imp_LeadsTo_comp2 = thm "imp_LeadsTo_comp2";
-val Follows_type = thm "Follows_type";
-val Follows_into_program = thm "Follows_into_program";
-val FollowsD = thm "FollowsD";
-val Follows_constantI = thm "Follows_constantI";
-val subset_Follows_comp = thm "subset_Follows_comp";
-val imp_Follows_comp = thm "imp_Follows_comp";
-val imp_Follows_comp2 = thm "imp_Follows_comp2";
-val Follows_trans = thm "Follows_trans";
-val Follows_imp_Increasing_left = thm "Follows_imp_Increasing_left";
-val Follows_imp_Increasing_right = thm "Follows_imp_Increasing_right";
-val Follows_imp_Always = thm "Follows_imp_Always";
-val Follows_imp_LeadsTo = thm "Follows_imp_LeadsTo";
-val Follows_LeadsTo_pfixLe = thm "Follows_LeadsTo_pfixLe";
-val Follows_LeadsTo_pfixGe = thm "Follows_LeadsTo_pfixGe";
-val Always_Follows1 = thm "Always_Follows1";
-val Always_Follows2 = thm "Always_Follows2";
-val refl_SetLe = thm "refl_SetLe";
-val trans_on_SetLe = thm "trans_on_SetLe";
-val antisym_SetLe = thm "antisym_SetLe";
-val part_order_SetLe = thm "part_order_SetLe";
-val increasing_Un = thm "increasing_Un";
-val Increasing_Un = thm "Increasing_Un";
-val Always_Un = thm "Always_Un";
-val Follows_Un = thm "Follows_Un";
-val refl_MultLe = thm "refl_MultLe";
-val MultLe_refl1 = thm "MultLe_refl1";
-val MultLe_refl2 = thm "MultLe_refl2";
-val trans_on_MultLe = thm "trans_on_MultLe";
-val MultLe_type = thm "MultLe_type";
-val MultLe_trans = thm "MultLe_trans";
-val part_order_imp_part_ord = thm "part_order_imp_part_ord";
-val antisym_MultLe = thm "antisym_MultLe";
-val part_order_MultLe = thm "part_order_MultLe";
-val empty_le_MultLe = thm "empty_le_MultLe";
-val empty_le_MultLe2 = thm "empty_le_MultLe2";
-val munion_mono = thm "munion_mono";
-val increasing_munion = thm "increasing_munion";
-val Increasing_munion = thm "Increasing_munion";
-val Always_munion = thm "Always_munion";
-val Follows_munion = thm "Follows_munion";
-val Follows_msetsum_UN = thm "Follows_msetsum_UN";
-*}
-
 end
--- a/src/ZF/UNITY/GenPrefix.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/GenPrefix.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -15,8 +15,8 @@
 imports Main
 begin
 
-constdefs (*really belongs in ZF/Trancl*)
-  part_order :: "[i, i] => o"
+definition (*really belongs in ZF/Trancl*)
+  part_order :: "[i, i] => o"  where
   "part_order(A, r) == refl(A,r) & trans[A](r) & antisym(r)"
 
 consts
@@ -37,11 +37,12 @@
 	      ==> <xs, ys@zs>:gen_prefix(A, r)"
     type_intros app_type list.Nil list.Cons
 
-constdefs
-   prefix :: "i=>i"
+definition
+  prefix :: "i=>i"  where
   "prefix(A) == gen_prefix(A, id(A))"
 
-   strict_prefix :: "i=>i"
+definition
+  strict_prefix :: "i=>i"  where
   "strict_prefix(A) == prefix(A) - id(list(A))"
 
 
--- a/src/ZF/UNITY/Guar.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Guar.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -34,49 +34,57 @@
 done
 
 
-constdefs
-
   (*Existential and Universal properties.  We formalize the two-program
     case, proving equivalence with Chandy and Sanders's n-ary definitions*)
 
-   ex_prop :: "i => o"
+definition
+   ex_prop :: "i => o"  where
    "ex_prop(X) == X<=program &
     (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X | G \<in> X --> (F Join G) \<in> X)"
 
-  strict_ex_prop  :: "i => o"
+definition
+  strict_ex_prop  :: "i => o"  where
   "strict_ex_prop(X) == X<=program &
    (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> (F \<in> X | G \<in> X) <-> (F Join G \<in> X))"
 
-  uv_prop  :: "i => o"
+definition
+  uv_prop  :: "i => o"  where
    "uv_prop(X) == X<=program &
     (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G --> F \<in> X & G \<in> X --> (F Join G) \<in> X))"
   
- strict_uv_prop  :: "i => o"
+definition
+ strict_uv_prop  :: "i => o"  where
    "strict_uv_prop(X) == X<=program &
     (SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G -->(F \<in> X & G \<in> X) <-> (F Join G \<in> X)))"
 
-  guar :: "[i, i] => i" (infixl "guarantees" 55)
+definition
+  guar :: "[i, i] => i" (infixl "guarantees" 55)  where
               (*higher than membership, lower than Co*)
   "X guarantees Y == {F \<in> program. \<forall>G \<in> program. F ok G --> F Join G \<in> X --> F Join G \<in> Y}"
   
+definition
   (* Weakest guarantees *)
-   wg :: "[i,i] => i"
+   wg :: "[i,i] => i"  where
   "wg(F,Y) == Union({X \<in> Pow(program). F:(X guarantees Y)})"
 
+definition
   (* Weakest existential property stronger than X *)
-   wx :: "i =>i"
+   wx :: "i =>i"  where
    "wx(X) == Union({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
 
+definition
   (*Ill-defined programs can arise through "Join"*)
-  welldef :: i
+  welldef :: i  where
   "welldef == {F \<in> program. Init(F) \<noteq> 0}"
   
-  refines :: "[i, i, i] => o" ("(3_ refines _ wrt _)" [10,10,10] 10)
+definition
+  refines :: "[i, i, i] => o" ("(3_ refines _ wrt _)" [10,10,10] 10)  where
   "G refines F wrt X ==
    \<forall>H \<in> program. (F ok H  & G ok H & F Join H \<in> welldef Int X)
     --> (G Join H \<in> welldef Int X)"
 
-  iso_refines :: "[i,i, i] => o"  ("(3_ iso'_refines _ wrt _)" [10,10,10] 10)
+definition
+  iso_refines :: "[i,i, i] => o"  ("(3_ iso'_refines _ wrt _)" [10,10,10] 10)  where
   "G iso_refines F wrt X ==  F \<in> welldef Int X --> G \<in> welldef Int X"
 
 
--- a/src/ZF/UNITY/Increasing.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Increasing.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -11,14 +11,14 @@
 
 theory Increasing imports Constrains Monotonicity begin
 
-constdefs
-
-  increasing :: "[i, i, i=>i] => i" ("increasing[_]'(_, _')")
+definition
+  increasing :: "[i, i, i=>i] => i" ("increasing[_]'(_, _')")  where
   "increasing[A](r, f) ==
     {F \<in> program. (\<forall>k \<in> A. F \<in> stable({s \<in> state. <k, f(s)> \<in> r})) &
                 (\<forall>x \<in> state. f(x):A)}"
   
-  Increasing :: "[i, i, i=>i] => i" ("Increasing[_]'(_, _')")
+definition
+  Increasing :: "[i, i, i=>i] => i" ("Increasing[_]'(_, _')")  where
   "Increasing[A](r, f) ==
     {F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) &
                 (\<forall>x \<in> state. f(x):A)}"
@@ -225,5 +225,4 @@
 apply simp_all
 done
 
-  
 end
--- a/src/ZF/UNITY/Merge.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Merge.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -12,26 +12,29 @@
 (** Merge specification (the number of inputs is Nclients) ***)
 (** Parameter A represents the type of items to Merge **)
 
-constdefs
+definition
   (*spec (10)*)
-  merge_increasing :: "[i, i, i] =>i"
+  merge_increasing :: "[i, i, i] =>i"  where
     "merge_increasing(A, Out, iOut) == program guarantees
          (lift(Out) IncreasingWrt  prefix(A)/list(A)) Int
          (lift(iOut) IncreasingWrt prefix(nat)/list(nat))"
 
+definition
   (*spec (11)*)
-  merge_eq_Out :: "[i, i] =>i"
+  merge_eq_Out :: "[i, i] =>i"  where
   "merge_eq_Out(Out, iOut) == program guarantees
          Always({s \<in> state. length(s`Out) = length(s`iOut)})"
 
+definition
   (*spec (12)*)
-  merge_bounded :: "i=>i"
+  merge_bounded :: "i=>i"  where
   "merge_bounded(iOut) == program guarantees
          Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})"
   
+definition
   (*spec (13)*)
   (* Parameter A represents the type of tokens *)
-  merge_follows :: "[i, i=>i, i, i] =>i"
+  merge_follows :: "[i, i=>i, i, i] =>i"  where
     "merge_follows(A, In, Out, iOut) ==
      (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))
 		   guarantees
@@ -40,18 +43,21 @@
                       nth(k, s`iOut) = n})) Fols lift(In(n))
          Wrt prefix(A)/list(A))"
 
+definition
   (*spec: preserves part*)
-  merge_preserves :: "[i=>i] =>i"
+  merge_preserves :: "[i=>i] =>i"  where
     "merge_preserves(In) == \<Inter>n \<in> nat. preserves(lift(In(n)))"
 
+definition
 (* environmental constraints*)
-  merge_allowed_acts :: "[i, i] =>i"
+  merge_allowed_acts :: "[i, i] =>i"  where
   "merge_allowed_acts(Out, iOut) ==
          {F \<in> program. AllowedActs(F) =
             cons(id(state), (\<Union>G \<in> preserves(lift(Out)) \<inter>
                                    preserves(lift(iOut)). Acts(G)))}"
   
-  merge_spec :: "[i, i =>i, i, i]=>i"
+definition
+  merge_spec :: "[i, i =>i, i, i]=>i"  where
   "merge_spec(A, In, Out, iOut) ==
    merge_increasing(A, Out, iOut) \<inter> merge_eq_Out(Out, iOut) \<inter>
    merge_bounded(iOut) \<inter>  merge_follows(A, In, Out, iOut)
@@ -189,5 +195,3 @@
 done
 
 end
-
-  
\ No newline at end of file
--- a/src/ZF/UNITY/Monotonicity.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Monotonicity.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -8,17 +8,18 @@
 
 header{*Monotonicity of an Operator WRT a Relation*}
 
-theory Monotonicity imports GenPrefix MultisetSum begin
+theory Monotonicity imports GenPrefix MultisetSum
+begin
 
-constdefs
-
-  mono1 :: "[i, i, i, i, i=>i] => o"
+definition
+  mono1 :: "[i, i, i, i, i=>i] => o"  where
   "mono1(A, r, B, s, f) ==
     (\<forall>x \<in> A. \<forall>y \<in> A. <x,y> \<in> r --> <f(x), f(y)> \<in> s) & (\<forall>x \<in> A. f(x) \<in> B)"
 
   (* monotonicity of a 2-place meta-function f *)
 
-  mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o"
+definition
+  mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o"  where
   "mono2(A, r, B, s, C, t, f) == 
     (\<forall>x \<in> A. \<forall>y \<in> A. \<forall>u \<in> B. \<forall>v \<in> B.
               <x,y> \<in> r & <u,v> \<in> s --> <f(x,u), f(y,v)> \<in> t) &
@@ -26,14 +27,15 @@
 
  (* Internalized relations on sets and multisets *)
 
-  SetLe :: "i =>i"
+definition
+  SetLe :: "i =>i"  where
   "SetLe(A) == {<x,y> \<in> Pow(A)*Pow(A). x <= y}"
 
-  MultLe :: "[i,i] =>i"
+definition
+  MultLe :: "[i,i] =>i"  where
   "MultLe(A, r) == multirel(A, r - id(A)) Un id(Mult(A))"
 
 
-
 lemma mono1D: 
   "[| mono1(A, r, B, s, f); <x, y> \<in> r; x \<in> A; y \<in> A |] ==> <f(x), f(y)> \<in> s"
 by (unfold mono1_def, auto)
@@ -115,19 +117,4 @@
 lemma mono_succ [iff]: "mono1(nat, Le, nat, Le, succ)"
 by (unfold mono1_def Le_def, auto)
 
-ML{*
-val mono1D = thm "mono1D";
-val mono2D = thm "mono2D";
-val take_mono_left_lemma = thm "take_mono_left_lemma";
-val take_mono_left = thm "take_mono_left";
-val take_mono_right = thm "take_mono_right";
-val take_mono = thm "take_mono";
-val mono_take = thm "mono_take";
-val length_mono = thm "length_mono";
-val mono_length = thm "mono_length";
-val mono_Un = thm "mono_Un";
-val mono_munion = thm "mono_munion";
-val mono_succ = thm "mono_succ";
-*}
-
 end
\ No newline at end of file
--- a/src/ZF/UNITY/MultisetSum.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/MultisetSum.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,22 +9,24 @@
 imports "../Induct/Multiset"
 begin
 
-constdefs
-
-  lcomm :: "[i, i, [i,i]=>i]=>o"
+definition
+  lcomm :: "[i, i, [i,i]=>i]=>o"  where
   "lcomm(A, B, f) ==
    (\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> B. f(x, f(y, z))= f(y, f(x, z)))  &
    (\<forall>x \<in> A. \<forall>y \<in> B. f(x, y) \<in> B)"
 
-  general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i"
+definition
+  general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i"  where
    "general_setsum(C, B, e, f, g) ==
     if Finite(C) then fold[cons(e, B)](%x y. f(g(x), y), e, C) else e"
 
-  msetsum :: "[i=>i, i, i]=>i"
+definition
+  msetsum :: "[i=>i, i, i]=>i"  where
   "msetsum(g, C, B) == normalize(general_setsum(C, Mult(B), 0, op +#, g))"
 
-  (* sum for naturals *)
-  nsetsum :: "[i=>i, i] =>i"
+
+definition  (* sum for naturals *)
+  nsetsum :: "[i=>i, i] =>i"  where
   "nsetsum(g, C) == general_setsum(C, nat, 0, op #+, g)"
 
 
@@ -181,29 +183,4 @@
 apply (erule Finite_induct, auto)
 done
 
-ML
-{*
-val mset_of_normalize = thm "mset_of_normalize";
-val general_setsum_0 = thm "general_setsum_0";
-val general_setsum_cons = thm "general_setsum_cons";
-val lcomm_mono = thm "lcomm_mono";
-val munion_lcomm = thm "munion_lcomm";
-val multiset_general_setsum = thm "multiset_general_setsum";
-val msetsum_0 = thm "msetsum_0";
-val msetsum_cons = thm "msetsum_cons";
-val msetsum_multiset = thm "msetsum_multiset";
-val mset_of_msetsum = thm "mset_of_msetsum";
-val msetsum_Un_Int = thm "msetsum_Un_Int";
-val msetsum_Un_disjoint = thm "msetsum_Un_disjoint";
-val UN_Fin_lemma = thm "UN_Fin_lemma";
-val msetsum_UN_disjoint = thm "msetsum_UN_disjoint";
-val msetsum_addf = thm "msetsum_addf";
-val msetsum_cong = thm "msetsum_cong";
-val multiset_union_diff = thm "multiset_union_diff";
-val msetsum_Un = thm "msetsum_Un";
-val nsetsum_0 = thm "nsetsum_0";
-val nsetsum_cons = thm "nsetsum_cons";
-val nsetsum_type = thm "nsetsum_type";
-*}
-
 end
\ No newline at end of file
--- a/src/ZF/UNITY/Mutex.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Mutex.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -1,6 +1,12 @@
 (*  ID:         $Id$
     Author:     Sidi O Ehmety, Computer Laboratory
     Copyright   2001  University of Cambridge
+
+Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
+
+Variables' types are introduced globally so that type verification
+reduces to the usual ZF typechecking \<in> an ill-tyed expression will
+reduce to the empty set.
 *)
 
 header{*Mutual Exclusion*}
@@ -28,75 +34,61 @@
   u_type:  "type_of(u)=bool & default_val(u)=0"
   v_type:  "type_of(v)=bool & default_val(v)=0"
 
-constdefs
+definition
   (** The program for process U **)
-   U0 :: i
-    "U0 == {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
+  "U0 == {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
 
-  U1 :: i
+definition
   "U1 == {<s,t>:state*state. t = s(p:= s`v, m:=#2) &  s`m = #1}"
 
-  U2 :: i
-    "U2 == {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
+definition
+  "U2 == {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
 
-  U3 :: i
-    "U3 == {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
+definition
+  "U3 == {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
 
-  U4 :: i
-    "U4 == {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
+definition
+  "U4 == {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
 
 
    (** The program for process V **)
 
-  V0 :: i
-    "V0 == {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
+definition
+  "V0 == {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
 
-  V1 :: i
-    "V1 == {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
+definition
+  "V1 == {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
 
-  V2 :: i
-    "V2 == {<s,t>:state*state. t  = s(n:=#3) & s`p=1 & s`n = #2}"
+definition
+  "V2 == {<s,t>:state*state. t  = s(n:=#3) & s`p=1 & s`n = #2}"
 
-  V3 :: i
+definition
   "V3 == {<s,t>:state*state. t = s (v:=0, n:=#4) & s`n = #3}"
 
-  V4 :: i
-    "V4 == {<s,t>:state*state. t  = s (p:=0, n:=#0) & s`n = #4}"
+definition
+  "V4 == {<s,t>:state*state. t  = s (p:=0, n:=#0) & s`n = #4}"
 
- Mutex :: i
- "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
+definition
+  "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
               {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))"
 
   (** The correct invariants **)
 
-  IU :: i
-    "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
+definition
+  "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
 	             & (s`m = #3 --> s`p=0)}"
 
-  IV :: i
-    "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
+definition
+  "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
 	              & (s`n = #3 --> s`p=1)}"
 
   (** The faulty invariant (for U alone) **)
 
-  bad_IU :: i
-    "bad_IU == {s:state. (s`u = 1 <-> (#1 $<= s`m & s`m  $<= #3))&
+definition
+  "bad_IU == {s:state. (s`u = 1 <-> (#1 $<= s`m & s`m  $<= #3))&
                    (#3 $<= s`m & s`m $<= #4 --> s`p=0)}"
 
 
-(*  Title:      ZF/UNITY/Mutex.ML
-    ID:         $Id \<in> Mutex.ML,v 1.4 2003/05/27 09:39:05 paulson Exp $
-    Author:     Sidi O Ehmety, Computer Laboratory
-    Copyright   2001  University of Cambridge
-
-Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
-
-Variables' types are introduced globally so that type verification
-reduces to the usual ZF typechecking \<in> an ill-tyed expression will
-reduce to the empty set.
-
-*)
-
 (** Variables' types **)
 
 declare p_type [simp] u_type [simp] v_type [simp] m_type [simp] n_type [simp]
@@ -211,7 +203,7 @@
 (*** Progress for U ***)
 
 lemma U_F0: "Mutex \<in> {s \<in> state. s`m=#2} Unless {s \<in> state. s`m=#3}"
-by (unfold Unless_def Mutex_def, safety)
+by (unfold op_Unless_def Mutex_def, safety)
 
 lemma U_F1:
      "Mutex \<in> {s \<in> state. s`m=#1} LeadsTo {s \<in> state. s`p = s`v & s`m = #2}"
@@ -263,7 +255,7 @@
 (*** Progress for V ***)
 
 lemma V_F0: "Mutex \<in> {s \<in> state. s`n=#2} Unless {s \<in> state. s`n=#3}"
-by (unfold Unless_def Mutex_def, safety)
+by (unfold op_Unless_def Mutex_def, safety)
 
 lemma V_F1: "Mutex \<in> {s \<in> state. s`n=#1} LeadsTo {s \<in> state. s`p = not(s`u) & s`n = #2}"
 by (unfold Mutex_def, ensures_tac "V1")
--- a/src/ZF/UNITY/State.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/State.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -21,19 +21,20 @@
   type_of :: "i=>i"
   default_val :: "i=>i"
 
-constdefs
-  state   :: i
-   "state == \<Pi> x \<in> var. cons(default_val(x), type_of(x))"
+definition
+  "state == \<Pi> x \<in> var. cons(default_val(x), type_of(x))"
 
-  st0     :: i
-   "st0 == \<lambda>x \<in> var. default_val(x)"
+definition
+  "st0 == \<lambda>x \<in> var. default_val(x)"
   
-  st_set  :: "i=>o"
+definition
+  st_set  :: "i=>o"  where
 (* To prevent typing conditions like `A<=state' from
    being used in combination with the rules `constrains_weaken', etc. *)
   "st_set(A) == A<=state"
 
-  st_compl :: "i=>i"
+definition
+  st_compl :: "i=>i"  where
   "st_compl(A) == state-A"
 
 
@@ -111,30 +112,4 @@
      "\<forall>d\<in>D. f(d) \<in> A ==> (\<Union>k\<in>A. {d\<in>D. P(d) & f(d) = k}) = {d\<in>D. P(d)}"
 by blast
 
-
-ML
-{*
-val st_set_def = thm "st_set_def";
-val state_def = thm "state_def";
-
-val st0_in_state = thm "st0_in_state";
-val st_set_Collect = thm "st_set_Collect";
-val st_set_0 = thm "st_set_0";
-val st_set_state = thm "st_set_state";
-val st_set_Un_iff = thm "st_set_Un_iff";
-val st_set_Union_iff = thm "st_set_Union_iff";
-val st_set_Int = thm "st_set_Int";
-val st_set_Inter = thm "st_set_Inter";
-val st_set_DiffI = thm "st_set_DiffI";
-val Collect_Int_state = thm "Collect_Int_state";
-val state_Int_Collect = thm "state_Int_Collect";
-val st_setI = thm "st_setI";
-val st_setD = thm "st_setD";
-val st_set_subset = thm "st_set_subset";
-val state_update_type = thm "state_update_type";
-val st_set_compl = thm "st_set_compl";
-val st_compl_iff = thm "st_compl_iff";
-val st_compl_Collect = thm "st_compl_Collect";
-*}
-
 end
--- a/src/ZF/UNITY/SubstAx.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/SubstAx.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -9,19 +9,19 @@
 
 theory SubstAx
 imports WFair Constrains 
-
 begin
 
-constdefs
-  (* The definitions below are not `conventional', but yields simpler rules *)
-   Ensures :: "[i,i] => i"            (infixl "Ensures" 60)
-    "A Ensures B == {F:program. F : (reachable(F) Int A) ensures (reachable(F) Int B) }"
+definition
+  (* The definitions below are not `conventional', but yield simpler rules *)
+  Ensures :: "[i,i] => i"            (infixl "Ensures" 60)  where
+  "A Ensures B == {F:program. F : (reachable(F) Int A) ensures (reachable(F) Int B) }"
 
-  LeadsTo :: "[i, i] => i"            (infixl "LeadsTo" 60)
-    "A LeadsTo B == {F:program. F:(reachable(F) Int A) leadsTo (reachable(F) Int B)}"
+definition
+  LeadsTo :: "[i, i] => i"            (infixl "LeadsTo" 60)  where
+  "A LeadsTo B == {F:program. F:(reachable(F) Int A) leadsTo (reachable(F) Int B)}"
 
-syntax (xsymbols)
-  "LeadsTo" :: "[i, i] => i" (infixl " \<longmapsto>w " 60)
+notation (xsymbols)
+  LeadsTo  (infixl " \<longmapsto>w " 60)
 
 
 
@@ -260,7 +260,7 @@
 
 lemma PSP_Unless: 
 "[| F \<in> A LeadsTo A'; F \<in> B Unless B'|]==> F:(A Int B) LeadsTo ((A' Int B) Un B')"
-apply (unfold Unless_def)
+apply (unfold op_Unless_def)
 apply (drule PSP, assumption)
 apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo)
 done
@@ -348,82 +348,29 @@
 apply (rule_tac [3] subset_refl, auto) 
 done
 
-ML
-{*
-val LeadsTo_eq = thm "LeadsTo_eq";
-val LeadsTo_type = thm "LeadsTo_type";
-val Always_LeadsTo_pre = thm "Always_LeadsTo_pre";
-val Always_LeadsTo_post = thm "Always_LeadsTo_post";
-val Always_LeadsToI = thm "Always_LeadsToI";
-val Always_LeadsToD = thm "Always_LeadsToD";
-val LeadsTo_Basis = thm "LeadsTo_Basis";
-val LeadsTo_Trans = thm "LeadsTo_Trans";
-val LeadsTo_Union = thm "LeadsTo_Union";
-val leadsTo_imp_LeadsTo = thm "leadsTo_imp_LeadsTo";
-val LeadsTo_Un_duplicate = thm "LeadsTo_Un_duplicate";
-val LeadsTo_Un_duplicate2 = thm "LeadsTo_Un_duplicate2";
-val LeadsTo_UN = thm "LeadsTo_UN";
-val LeadsTo_Un = thm "LeadsTo_Un";
-val single_LeadsTo_I = thm "single_LeadsTo_I";
-val subset_imp_LeadsTo = thm "subset_imp_LeadsTo";
-val empty_LeadsTo = thm "empty_LeadsTo";
-val LeadsTo_state = thm "LeadsTo_state";
-val LeadsTo_weaken_R = thm "LeadsTo_weaken_R";
-val LeadsTo_weaken_L = thm "LeadsTo_weaken_L";
-val LeadsTo_weaken = thm "LeadsTo_weaken";
-val Always_LeadsTo_weaken = thm "Always_LeadsTo_weaken";
-val LeadsTo_Un_post = thm "LeadsTo_Un_post";
-val LeadsTo_Trans_Un = thm "LeadsTo_Trans_Un";
-val LeadsTo_Un_distrib = thm "LeadsTo_Un_distrib";
-val LeadsTo_UN_distrib = thm "LeadsTo_UN_distrib";
-val LeadsTo_Union_distrib = thm "LeadsTo_Union_distrib";
-val EnsuresI = thm "EnsuresI";
-val Always_LeadsTo_Basis = thm "Always_LeadsTo_Basis";
-val LeadsTo_Diff = thm "LeadsTo_Diff";
-val LeadsTo_UN_UN = thm "LeadsTo_UN_UN";
-val LeadsTo_Un_Un = thm "LeadsTo_Un_Un";
-val LeadsTo_cancel2 = thm "LeadsTo_cancel2";
-val Un_Diff = thm "Un_Diff";
-val LeadsTo_cancel_Diff2 = thm "LeadsTo_cancel_Diff2";
-val LeadsTo_cancel1 = thm "LeadsTo_cancel1";
-val Diff_Un2 = thm "Diff_Un2";
-val LeadsTo_cancel_Diff1 = thm "LeadsTo_cancel_Diff1";
-val LeadsTo_empty = thm "LeadsTo_empty";
-val PSP_Stable = thm "PSP_Stable";
-val PSP_Stable2 = thm "PSP_Stable2";
-val PSP = thm "PSP";
-val PSP2 = thm "PSP2";
-val PSP_Unless = thm "PSP_Unless";
-val LeadsTo_wf_induct = thm "LeadsTo_wf_induct";
-val LessThan_induct = thm "LessThan_induct";
-val Completion = thm "Completion";
-val Finite_completion = thm "Finite_completion";
-val Stable_completion = thm "Stable_completion";
-val Finite_stable_completion = thm "Finite_stable_completion";
-
-
+ML {*
 (*proves "ensures/leadsTo" properties when the program is specified*)
 fun ensures_tac ctxt sact =
   let val css as (cs, ss) = local_clasimpset_of ctxt in
     SELECT_GOAL
       (EVERY [REPEAT (Always_Int_tac 1),
-              etac Always_LeadsTo_Basis 1 
+              etac @{thm Always_LeadsTo_Basis} 1 
                   ORELSE   (*subgoal may involve LeadsTo, leadsTo or ensures*)
-                  REPEAT (ares_tac [LeadsTo_Basis, leadsTo_Basis,
-                                    EnsuresI, ensuresI] 1),
+                  REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
+                                    @{thm EnsuresI}, @{thm ensuresI}] 1),
               (*now there are two subgoals: co & transient*)
               simp_tac (ss addsimps (ProgramDefs.get ctxt)) 2,
-              res_inst_tac [("act", sact)] transientI 2,
+              res_inst_tac [("act", sact)] @{thm transientI} 2,
                  (*simplify the command's domain*)
-              simp_tac (ss addsimps [domain_def]) 3, 
+              simp_tac (ss addsimps [@{thm domain_def}]) 3, 
               (* proving the domain part *)
              clarify_tac cs 3, dtac Cla.swap 3, force_tac css 4,
-             rtac ReplaceI 3, force_tac css 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,
-             REPEAT (rtac state_update_type 3),
+             REPEAT (rtac @{thm state_update_type} 3),
              constrains_tac ctxt 1,
              ALLGOALS (clarify_tac cs),
-             ALLGOALS (asm_full_simp_tac (ss addsimps [st_set_def])),
+             ALLGOALS (asm_full_simp_tac (ss addsimps [@{thm st_set_def}])),
                         ALLGOALS (clarify_tac cs),
             ALLGOALS (asm_lr_simp_tac ss)])
   end;
--- a/src/ZF/UNITY/UNITY.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/UNITY.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -17,65 +17,81 @@
   "constrains" :: "[i, i] => i"  (infixl "co"     60)
   op_unless    :: "[i, i] => i"  (infixl "unless" 60)
 
-constdefs
-   program  :: i
+definition
+  program  :: i  where
   "program == {<init, acts, allowed>:
 	       Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
 	       id(state) \<in> acts & id(state) \<in> allowed}"
 
-  mk_program :: "[i,i,i]=>i"
+definition
+  mk_program :: "[i,i,i]=>i"  where
   --{* The definition yields a program thanks to the coercions
        init \<inter> state, acts \<inter> Pow(state*state), etc. *}
   "mk_program(init, acts, allowed) ==
     <init \<inter> state, cons(id(state), acts \<inter> Pow(state*state)),
               cons(id(state), allowed \<inter> Pow(state*state))>"
 
-  SKIP :: i
+definition
+  SKIP :: i  where
   "SKIP == mk_program(state, 0, Pow(state*state))"
 
   (* Coercion from anything to program *)
-  programify :: "i=>i"
+definition
+  programify :: "i=>i"  where
   "programify(F) == if F \<in> program then F else SKIP"
 
-  RawInit :: "i=>i"
+definition
+  RawInit :: "i=>i"  where
   "RawInit(F) == fst(F)"
 
-  Init :: "i=>i"
+definition
+  Init :: "i=>i"  where
   "Init(F) == RawInit(programify(F))"
 
-  RawActs :: "i=>i"
+definition
+  RawActs :: "i=>i"  where
   "RawActs(F) == cons(id(state), fst(snd(F)))"
 
-  Acts :: "i=>i"
+definition
+  Acts :: "i=>i"  where
   "Acts(F) == RawActs(programify(F))"
 
-  RawAllowedActs :: "i=>i"
+definition
+  RawAllowedActs :: "i=>i"  where
   "RawAllowedActs(F) == cons(id(state), snd(snd(F)))"
 
-  AllowedActs :: "i=>i"
+definition
+  AllowedActs :: "i=>i"  where
   "AllowedActs(F) == RawAllowedActs(programify(F))"
 
 
-  Allowed :: "i =>i"
+definition
+  Allowed :: "i =>i"  where
   "Allowed(F) == {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}"
 
-  initially :: "i=>i"
+definition
+  initially :: "i=>i"  where
   "initially(A) == {F \<in> program. Init(F)\<subseteq>A}"
 
-  stable     :: "i=>i"
+definition
+  stable     :: "i=>i"  where
    "stable(A) == A co A"
 
-  strongest_rhs :: "[i, i] => i"
+definition
+  strongest_rhs :: "[i, i] => i"  where
   "strongest_rhs(F, A) == Inter({B \<in> Pow(state). F \<in> A co B})"
 
-  invariant :: "i => i"
+definition
+  invariant :: "i => i"  where
   "invariant(A) == initially(A) \<inter> stable(A)"
 
   (* meta-function composition *)
-  metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl "comp" 65)
+definition
+  metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl "comp" 65)  where
   "f comp g == %x. f(g(x))"
 
-  pg_compl :: "i=>i"
+definition
+  pg_compl :: "i=>i"  where
   "pg_compl(X)== program - X"
 
 defs
@@ -607,101 +623,9 @@
      "[| F \<in> A co B; st_set(B) |] ==> strongest_rhs(F,A) \<subseteq> B"
 by (auto simp add: constrains_def strongest_rhs_def st_set_def)
 
-ML
-{*
-val constrains_def = thm "constrains_def";
-val stable_def = thm "stable_def";
-val invariant_def = thm "invariant_def";
-val unless_def = thm "unless_def";
-val initially_def = thm "initially_def";
-val SKIP_def = thm "SKIP_def";
-val Allowed_def = thm "Allowed_def";
-val programify_def = thm "programify_def";
-val metacomp_def = thm "metacomp_def";
-
-val id_in_Acts = thm "id_in_Acts";
-val id_in_RawAllowedActs = thm "id_in_RawAllowedActs";
-val id_in_AllowedActs = thm "id_in_AllowedActs";
-val cons_id_Acts = thm "cons_id_Acts";
-val cons_id_AllowedActs = thm "cons_id_AllowedActs";
-val Init_type = thm "Init_type";
-val st_set_Init = thm "st_set_Init";
-val Acts_type = thm "Acts_type";
-val AllowedActs_type = thm "AllowedActs_type";
-val ActsD = thm "ActsD";
-val AllowedActsD = thm "AllowedActsD";
-val mk_program_in_program = thm "mk_program_in_program";
-val Init_eq = thm "Init_eq";
-val Acts_eq = thm "Acts_eq";
-val AllowedActs_eq = thm "AllowedActs_eq";
-val Init_SKIP = thm "Init_SKIP";
-val Acts_SKIP = thm "Acts_SKIP";
-val AllowedActs_SKIP = thm "AllowedActs_SKIP";
-val raw_surjective_mk_program = thm "raw_surjective_mk_program";
-val surjective_mk_program = thm "surjective_mk_program";
-val program_equalityI = thm "program_equalityI";
-val program_equalityE = thm "program_equalityE";
-val program_equality_iff = thm "program_equality_iff";
-val def_prg_Init = thm "def_prg_Init";
-val def_prg_Acts = thm "def_prg_Acts";
-val def_prg_AllowedActs = thm "def_prg_AllowedActs";
-val def_prg_simps = thm "def_prg_simps";
-val def_act_simp = thm "def_act_simp";
-val def_set_simp = thm "def_set_simp";
-val constrains_type = thm "constrains_type";
-val constrainsI = thm "constrainsI";
-val constrainsD = thm "constrainsD";
-val constrainsD2 = thm "constrainsD2";
-val constrains_empty = thm "constrains_empty";
-val constrains_empty2 = thm "constrains_empty2";
-val constrains_state = thm "constrains_state";
-val constrains_state2 = thm "constrains_state2";
-val constrains_weaken_R = thm "constrains_weaken_R";
-val constrains_weaken_L = thm "constrains_weaken_L";
-val constrains_weaken = thm "constrains_weaken";
-val constrains_Un = thm "constrains_Un";
-val constrains_UN = thm "constrains_UN";
-val constrains_Un_distrib = thm "constrains_Un_distrib";
-val constrains_UN_distrib = thm "constrains_UN_distrib";
-val constrains_Int_distrib = thm "constrains_Int_distrib";
-val constrains_INT_distrib = thm "constrains_INT_distrib";
-val constrains_Int = thm "constrains_Int";
-val constrains_INT = thm "constrains_INT";
-val constrains_All = thm "constrains_All";
-val constrains_imp_subset = thm "constrains_imp_subset";
-val constrains_trans = thm "constrains_trans";
-val constrains_cancel = thm "constrains_cancel";
-val unless_type = thm "unless_type";
-val unlessI = thm "unlessI";
-val unlessD = thm "unlessD";
-val initially_type = thm "initially_type";
-val initiallyI = thm "initiallyI";
-val initiallyD = thm "initiallyD";
-val stable_type = thm "stable_type";
-val stableI = thm "stableI";
-val stableD = thm "stableD";
-val stableD2 = thm "stableD2";
-val stable_state = thm "stable_state";
-val stable_unless = thm "stable_unless";
-val stable_Un = thm "stable_Un";
-val stable_UN = thm "stable_UN";
-val stable_Int = thm "stable_Int";
-val stable_INT = thm "stable_INT";
-val stable_All = thm "stable_All";
-val stable_constrains_Un = thm "stable_constrains_Un";
-val stable_constrains_Int = thm "stable_constrains_Int";
-val invariant_type = thm "invariant_type";
-val invariantI = thm "invariantI";
-val invariantD = thm "invariantD";
-val invariantD2 = thm "invariantD2";
-val invariant_Int = thm "invariant_Int";
-val elimination = thm "elimination";
-val elimination2 = thm "elimination2";
-val constrains_strongest_rhs = thm "constrains_strongest_rhs";
-val strongest_rhs_is_strongest = thm "strongest_rhs_is_strongest";
-
-fun simp_of_act def = def RS def_act_simp;
-fun simp_of_set def = def RS def_set_simp;
+ML {*
+fun simp_of_act def = def RS @{thm def_act_simp};
+fun simp_of_set def = def RS @{thm def_set_simp};
 *}
 
 end
--- a/src/ZF/UNITY/Union.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/Union.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -11,29 +11,33 @@
 
 *)
 
-theory Union imports SubstAx FP begin
+theory Union imports SubstAx FP
+begin
 
-constdefs
-
+definition
   (*FIXME: conjoin Init(F) Int Init(G) \<noteq> 0 *) 
-  ok :: "[i, i] => o"     (infixl "ok" 65)
+  ok :: "[i, i] => o"     (infixl "ok" 65)  where
     "F ok G == Acts(F) \<subseteq> AllowedActs(G) &
                Acts(G) \<subseteq> AllowedActs(F)"
 
+definition
   (*FIXME: conjoin (\<Inter>i \<in> I. Init(F(i))) \<noteq> 0 *) 
-  OK  :: "[i, i=>i] => o"
+  OK  :: "[i, i=>i] => o"  where
     "OK(I,F) == (\<forall>i \<in> I. \<forall>j \<in> I-{i}. Acts(F(i)) \<subseteq> AllowedActs(F(j)))"
 
-   JOIN  :: "[i, i=>i] => i"
+definition
+  JOIN  :: "[i, i=>i] => i"  where
   "JOIN(I,F) == if I = 0 then SKIP else
                  mk_program(\<Inter>i \<in> I. Init(F(i)), \<Union>i \<in> I. Acts(F(i)),
                  \<Inter>i \<in> I. AllowedActs(F(i)))"
 
-  Join :: "[i, i] => i"    (infixl "Join" 65)
+definition
+  Join :: "[i, i] => i"    (infixl "Join" 65)  where
   "F Join G == mk_program (Init(F) Int Init(G), Acts(F) Un Acts(G),
 				AllowedActs(F) Int AllowedActs(G))"
+definition
   (*Characterizes safety properties.  Used with specifying AllowedActs*)
-  safety_prop :: "i => o"
+  safety_prop :: "i => o"  where
   "safety_prop(X) == X\<subseteq>program &
       SKIP \<in> X & (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) --> G \<in> X)"
   
@@ -42,13 +46,15 @@
   "@JOIN"      :: "[pttrn, i, i] => i"       ("(3JN _:_./ _)" 10)
 
 translations
-  "JN x:A. B"   == "JOIN(A, (%x. B))"
+  "JN x:A. B"   == "CONST JOIN(A, (%x. B))"
   "JN x y. B"   == "JN x. JN y. B"
-  "JN x. B"     == "JOIN(state,(%x. B))"
+  "JN x. B"     == "CONST JOIN(CONST state,(%x. B))"
 
-syntax (symbols)
-   SKIP     :: i                      ("\<bottom>")
-  Join      :: "[i, i] => i"          (infixl "\<squnion>" 65)
+notation (xsymbols)
+  SKIP  ("\<bottom>") and
+  Join  (infixl "\<squnion>" 65)
+
+syntax (xsymbols)
   "@JOIN1"  :: "[pttrns, i] => i"     ("(3\<Squnion> _./ _)" 10)
   "@JOIN"   :: "[pttrn, i, i] => i"   ("(3\<Squnion> _ \<in> _./ _)" 10)
 
@@ -576,102 +582,4 @@
 apply (cut_tac F = G in Acts_type, auto)
 done
 
-
-ML
-{*
-val safety_prop_def = thm "safety_prop_def";
-
-val reachable_SKIP = thm "reachable_SKIP";
-val ok_programify_left = thm "ok_programify_left";
-val ok_programify_right = thm "ok_programify_right";
-val Join_programify_left = thm "Join_programify_left";
-val Join_programify_right = thm "Join_programify_right";
-val SKIP_in_constrains_iff = thm "SKIP_in_constrains_iff";
-val SKIP_in_Constrains_iff = thm "SKIP_in_Constrains_iff";
-val SKIP_in_stable = thm "SKIP_in_stable";
-val SKIP_in_Stable = thm "SKIP_in_Stable";
-val Join_in_program = thm "Join_in_program";
-val JOIN_in_program = thm "JOIN_in_program";
-val Init_Join = thm "Init_Join";
-val Acts_Join = thm "Acts_Join";
-val AllowedActs_Join = thm "AllowedActs_Join";
-val Join_commute = thm "Join_commute";
-val Join_left_commute = thm "Join_left_commute";
-val Join_assoc = thm "Join_assoc";
-val cons_id = thm "cons_id";
-val Join_SKIP_left = thm "Join_SKIP_left";
-val Join_SKIP_right = thm "Join_SKIP_right";
-val Join_absorb = thm "Join_absorb";
-val Join_left_absorb = thm "Join_left_absorb";
-val OK_programify = thm "OK_programify";
-val JN_programify = thm "JN_programify";
-val JN_empty = thm "JN_empty";
-val Init_JN = thm "Init_JN";
-val Acts_JN = thm "Acts_JN";
-val AllowedActs_JN = thm "AllowedActs_JN";
-val JN_cons = thm "JN_cons";
-val JN_cong = thm "JN_cong";
-val JN_absorb = thm "JN_absorb";
-val JN_Un = thm "JN_Un";
-val JN_constant = thm "JN_constant";
-val JN_Join_distrib = thm "JN_Join_distrib";
-val JN_Join_miniscope = thm "JN_Join_miniscope";
-val JN_Join_diff = thm "JN_Join_diff";
-val JN_constrains = thm "JN_constrains";
-val Join_constrains = thm "Join_constrains";
-val Join_unless = thm "Join_unless";
-val Join_constrains_weaken = thm "Join_constrains_weaken";
-val JN_constrains_weaken = thm "JN_constrains_weaken";
-val JN_stable = thm "JN_stable";
-val initially_JN_I = thm "initially_JN_I";
-val invariant_JN_I = thm "invariant_JN_I";
-val Join_stable = thm "Join_stable";
-val initially_JoinI = thm "initially_JoinI";
-val invariant_JoinI = thm "invariant_JoinI";
-val FP_JN = thm "FP_JN";
-val JN_transient = thm "JN_transient";
-val Join_transient = thm "Join_transient";
-val Join_transient_I1 = thm "Join_transient_I1";
-val Join_transient_I2 = thm "Join_transient_I2";
-val JN_ensures = thm "JN_ensures";
-val Join_ensures = thm "Join_ensures";
-val stable_Join_constrains = thm "stable_Join_constrains";
-val stable_Join_Always1 = thm "stable_Join_Always1";
-val stable_Join_Always2 = thm "stable_Join_Always2";
-val stable_Join_ensures1 = thm "stable_Join_ensures1";
-val stable_Join_ensures2 = thm "stable_Join_ensures2";
-val ok_SKIP1 = thm "ok_SKIP1";
-val ok_SKIP2 = thm "ok_SKIP2";
-val ok_Join_commute = thm "ok_Join_commute";
-val ok_commute = thm "ok_commute";
-val ok_sym = thm "ok_sym";
-val ok_iff_OK = thm "ok_iff_OK";
-val ok_Join_iff1 = thm "ok_Join_iff1";
-val ok_Join_iff2 = thm "ok_Join_iff2";
-val ok_Join_commute_I = thm "ok_Join_commute_I";
-val ok_JN_iff1 = thm "ok_JN_iff1";
-val ok_JN_iff2 = thm "ok_JN_iff2";
-val OK_iff_ok = thm "OK_iff_ok";
-val OK_imp_ok = thm "OK_imp_ok";
-val Allowed_SKIP = thm "Allowed_SKIP";
-val Allowed_Join = thm "Allowed_Join";
-val Allowed_JN = thm "Allowed_JN";
-val ok_iff_Allowed = thm "ok_iff_Allowed";
-val OK_iff_Allowed = thm "OK_iff_Allowed";
-val safety_prop_Acts_iff = thm "safety_prop_Acts_iff";
-val safety_prop_AllowedActs_iff_Allowed = thm "safety_prop_AllowedActs_iff_Allowed";
-val Allowed_eq = thm "Allowed_eq";
-val def_prg_Allowed = thm "def_prg_Allowed";
-val safety_prop_constrains = thm "safety_prop_constrains";
-val safety_prop_constrainsI = thm "safety_prop_constrainsI";
-val safety_prop_stable = thm "safety_prop_stable";
-val safety_prop_stableI = thm "safety_prop_stableI";
-val safety_prop_Int = thm "safety_prop_Int";
-val safety_prop_Inter = thm "safety_prop_Inter";
-val def_UNION_ok_iff = thm "def_UNION_ok_iff";
-
-val Join_ac = thms "Join_ac";
-*}
-
-
 end
--- a/src/ZF/UNITY/WFair.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/UNITY/WFair.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -14,15 +14,15 @@
 assuming weak fairness. From Misra, "A Logic for Concurrent Programming",
 1994.*}
 
-constdefs
-  
+definition
   (* This definition specifies weak fairness.  The rest of the theory
     is generic to all forms of fairness.*) 
- transient :: "i=>i"
+  transient :: "i=>i"  where
   "transient(A) =={F:program. (EX act: Acts(F). A<=domain(act) &
 			       act``A <= state-A) & st_set(A)}"
 
-  ensures :: "[i,i] => i"       (infixl "ensures" 60)
+definition
+  ensures :: "[i,i] => i"       (infixl "ensures" 60)  where
   "A ensures B == ((A-B) co (A Un B)) Int transient(A-B)"
   
 consts
@@ -42,18 +42,18 @@
   monos        Pow_mono
   type_intros  Union_Pow_iff [THEN iffD2] UnionI PowI
  
-constdefs
-
+definition
   (* The Visible version of the LEADS-TO relation*)
-  leadsTo :: "[i, i] => i"       (infixl "leadsTo" 60)
+  leadsTo :: "[i, i] => i"       (infixl "leadsTo" 60)  where
   "A leadsTo B == {F:program. <A,B>:leads(state, F)}"
   
+definition
   (* wlt(F, B) is the largest set that leads to B*)
-  wlt :: "[i, i] => i"
+  wlt :: "[i, i] => i"  where
     "wlt(F, B) == Union({A:Pow(state). F: A leadsTo B})"
 
-syntax (xsymbols)
-  "leadsTo" :: "[i, i] => i" (infixl "\<longmapsto>" 60)
+notation (xsymbols)
+  leadsTo  (infixl "\<longmapsto>" 60)
 
 (** Ad-hoc set-theory rules **)
 
@@ -708,93 +708,4 @@
 apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) 
 done
 
-ML
-{*
-val Int_Union_Union = thm "Int_Union_Union";
-val Int_Union_Union2 = thm "Int_Union_Union2";
-val transient_type = thm "transient_type";
-val transientD2 = thm "transientD2";
-val stable_transient_empty = thm "stable_transient_empty";
-val transient_strengthen = thm "transient_strengthen";
-val transientI = thm "transientI";
-val transientE = thm "transientE";
-val transient_state = thm "transient_state";
-val transient_state2 = thm "transient_state2";
-val transient_empty = thm "transient_empty";
-val ensures_type = thm "ensures_type";
-val ensuresI = thm "ensuresI";
-val ensuresI2 = thm "ensuresI2";
-val ensuresD = thm "ensuresD";
-val ensures_weaken_R = thm "ensures_weaken_R";
-val stable_ensures_Int = thm "stable_ensures_Int";
-val stable_transient_ensures = thm "stable_transient_ensures";
-val ensures_eq = thm "ensures_eq";
-val subset_imp_ensures = thm "subset_imp_ensures";
-val leads_left = thm "leads_left";
-val leads_right = thm "leads_right";
-val leadsTo_type = thm "leadsTo_type";
-val leadsToD2 = thm "leadsToD2";
-val leadsTo_Basis = thm "leadsTo_Basis";
-val subset_imp_leadsTo = thm "subset_imp_leadsTo";
-val leadsTo_Trans = thm "leadsTo_Trans";
-val transient_imp_leadsTo = thm "transient_imp_leadsTo";
-val leadsTo_Un_duplicate = thm "leadsTo_Un_duplicate";
-val leadsTo_Un_duplicate2 = thm "leadsTo_Un_duplicate2";
-val leadsTo_Union = thm "leadsTo_Union";
-val leadsTo_Union_Int = thm "leadsTo_Union_Int";
-val leadsTo_UN = thm "leadsTo_UN";
-val leadsTo_Un = thm "leadsTo_Un";
-val single_leadsTo_I = thm "single_leadsTo_I";
-val leadsTo_refl = thm "leadsTo_refl";
-val leadsTo_refl_iff = thm "leadsTo_refl_iff";
-val empty_leadsTo = thm "empty_leadsTo";
-val leadsTo_state = thm "leadsTo_state";
-val leadsTo_weaken_R = thm "leadsTo_weaken_R";
-val leadsTo_weaken_L = thm "leadsTo_weaken_L";
-val leadsTo_weaken = thm "leadsTo_weaken";
-val transient_imp_leadsTo2 = thm "transient_imp_leadsTo2";
-val leadsTo_Un_distrib = thm "leadsTo_Un_distrib";
-val leadsTo_UN_distrib = thm "leadsTo_UN_distrib";
-val leadsTo_Union_distrib = thm "leadsTo_Union_distrib";
-val leadsTo_Diff = thm "leadsTo_Diff";
-val leadsTo_UN_UN = thm "leadsTo_UN_UN";
-val leadsTo_Un_Un = thm "leadsTo_Un_Un";
-val leadsTo_cancel2 = thm "leadsTo_cancel2";
-val leadsTo_cancel_Diff2 = thm "leadsTo_cancel_Diff2";
-val leadsTo_cancel1 = thm "leadsTo_cancel1";
-val leadsTo_cancel_Diff1 = thm "leadsTo_cancel_Diff1";
-val leadsTo_induct = thm "leadsTo_induct";
-val leadsTo_induct2 = thm "leadsTo_induct2";
-val leadsTo_induct_pre_aux = thm "leadsTo_induct_pre_aux";
-val leadsTo_induct_pre = thm "leadsTo_induct_pre";
-val leadsTo_empty = thm "leadsTo_empty";
-val psp_stable = thm "psp_stable";
-val psp_stable2 = thm "psp_stable2";
-val psp_ensures = thm "psp_ensures";
-val psp = thm "psp";
-val psp2 = thm "psp2";
-val psp_unless = thm "psp_unless";
-val leadsTo_wf_induct_aux = thm "leadsTo_wf_induct_aux";
-val leadsTo_wf_induct = thm "leadsTo_wf_induct";
-val nat_measure_field = thm "nat_measure_field";
-val Image_inverse_lessThan = thm "Image_inverse_lessThan";
-val lessThan_induct = thm "lessThan_induct";
-val wlt_type = thm "wlt_type";
-val wlt_st_set = thm "wlt_st_set";
-val wlt_leadsTo_iff = thm "wlt_leadsTo_iff";
-val wlt_leadsTo = thm "wlt_leadsTo";
-val leadsTo_subset = thm "leadsTo_subset";
-val leadsTo_eq_subset_wlt = thm "leadsTo_eq_subset_wlt";
-val wlt_increasing = thm "wlt_increasing";
-val leadsTo_123_aux = thm "leadsTo_123_aux";
-val leadsTo_123 = thm "leadsTo_123";
-val wlt_constrains_wlt = thm "wlt_constrains_wlt";
-val completion_aux = thm "completion_aux";
-val completion = thm "completion";
-val finite_completion_aux = thm "finite_completion_aux";
-val finite_completion = thm "finite_completion";
-val stable_completion = thm "stable_completion";
-val finite_stable_completion = thm "finite_stable_completion";
-*}
-
 end
--- a/src/ZF/Univ.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Univ.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -13,8 +13,8 @@
 
 theory Univ imports Epsilon Cardinal begin
 
-constdefs
-  Vfrom       :: "[i,i]=>i"
+definition
+  Vfrom       :: "[i,i]=>i"  where
     "Vfrom(A,i) == transrec(i, %x f. A Un (\<Union>y\<in>x. Pow(f`y)))"
 
 abbreviation
@@ -22,16 +22,18 @@
   "Vset(x) == Vfrom(0,x)"
 
 
-constdefs
-  Vrec        :: "[i, [i,i]=>i] =>i"
+definition
+  Vrec        :: "[i, [i,i]=>i] =>i"  where
     "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
  		 	   H(z, lam w:Vset(x). g`rank(w)`w)) ` a"
 
-  Vrecursor   :: "[[i,i]=>i, i] =>i"
+definition
+  Vrecursor   :: "[[i,i]=>i, i] =>i"  where
     "Vrecursor(H,a) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
 				H(lam w:Vset(x). g`rank(w)`w, z)) ` a"
 
-  univ        :: "i=>i"
+definition
+  univ        :: "i=>i"  where
     "univ(A) == Vfrom(A,nat)"
 
 
@@ -787,115 +789,8 @@
 
 ML
 {*
-
-val Vfrom = thm "Vfrom";
-val Vfrom_mono = thm "Vfrom_mono";
-val Vfrom_rank_subset1 = thm "Vfrom_rank_subset1";
-val Vfrom_rank_subset2 = thm "Vfrom_rank_subset2";
-val Vfrom_rank_eq = thm "Vfrom_rank_eq";
-val zero_in_Vfrom = thm "zero_in_Vfrom";
-val i_subset_Vfrom = thm "i_subset_Vfrom";
-val A_subset_Vfrom = thm "A_subset_Vfrom";
-val subset_mem_Vfrom = thm "subset_mem_Vfrom";
-val singleton_in_Vfrom = thm "singleton_in_Vfrom";
-val doubleton_in_Vfrom = thm "doubleton_in_Vfrom";
-val Pair_in_Vfrom = thm "Pair_in_Vfrom";
-val succ_in_Vfrom = thm "succ_in_Vfrom";
-val Vfrom_0 = thm "Vfrom_0";
-val Vfrom_succ = thm "Vfrom_succ";
-val Vfrom_Union = thm "Vfrom_Union";
-val Limit_Vfrom_eq = thm "Limit_Vfrom_eq";
-val zero_in_VLimit = thm "zero_in_VLimit";
-val singleton_in_VLimit = thm "singleton_in_VLimit";
-val Vfrom_UnI1 = thm "Vfrom_UnI1";
-val Vfrom_UnI2 = thm "Vfrom_UnI2";
-val doubleton_in_VLimit = thm "doubleton_in_VLimit";
-val Pair_in_VLimit = thm "Pair_in_VLimit";
-val product_VLimit = thm "product_VLimit";
-val Sigma_subset_VLimit = thm "Sigma_subset_VLimit";
-val nat_subset_VLimit = thm "nat_subset_VLimit";
-val nat_into_VLimit = thm "nat_into_VLimit";
-val zero_in_VLimit = thm "zero_in_VLimit";
-val one_in_VLimit = thm "one_in_VLimit";
-val Inl_in_VLimit = thm "Inl_in_VLimit";
-val Inr_in_VLimit = thm "Inr_in_VLimit";
-val sum_VLimit = thm "sum_VLimit";
-val sum_subset_VLimit = thm "sum_subset_VLimit";
-val Transset_Vfrom = thm "Transset_Vfrom";
-val Transset_Vfrom_succ = thm "Transset_Vfrom_succ";
-val Transset_Pair_subset = thm "Transset_Pair_subset";
-val Union_in_Vfrom = thm "Union_in_Vfrom";
-val Union_in_VLimit = thm "Union_in_VLimit";
-val in_VLimit = thm "in_VLimit";
-val prod_in_Vfrom = thm "prod_in_Vfrom";
-val prod_in_VLimit = thm "prod_in_VLimit";
-val sum_in_Vfrom = thm "sum_in_Vfrom";
-val sum_in_VLimit = thm "sum_in_VLimit";
-val fun_in_Vfrom = thm "fun_in_Vfrom";
-val fun_in_VLimit = thm "fun_in_VLimit";
-val Pow_in_Vfrom = thm "Pow_in_Vfrom";
-val Pow_in_VLimit = thm "Pow_in_VLimit";
-val Vset = thm "Vset";
-val Vset_succ = thm "Vset_succ";
-val Transset_Vset = thm "Transset_Vset";
-val VsetD = thm "VsetD";
-val VsetI = thm "VsetI";
-val Vset_Ord_rank_iff = thm "Vset_Ord_rank_iff";
-val Vset_rank_iff = thm "Vset_rank_iff";
-val rank_Vset = thm "rank_Vset";
-val arg_subset_Vset_rank = thm "arg_subset_Vset_rank";
-val Int_Vset_subset = thm "Int_Vset_subset";
-val rank_Inl = thm "rank_Inl";
-val rank_Inr = thm "rank_Inr";
-val Vrec = thm "Vrec";
-val def_Vrec = thm "def_Vrec";
-val Vrecursor = thm "Vrecursor";
-val def_Vrecursor = thm "def_Vrecursor";
-val univ_mono = thm "univ_mono";
-val Transset_univ = thm "Transset_univ";
-val univ_eq_UN = thm "univ_eq_UN";
-val subset_univ_eq_Int = thm "subset_univ_eq_Int";
-val univ_Int_Vfrom_subset = thm "univ_Int_Vfrom_subset";
-val univ_Int_Vfrom_eq = thm "univ_Int_Vfrom_eq";
-val zero_in_univ = thm "zero_in_univ";
-val A_subset_univ = thm "A_subset_univ";
-val A_into_univ = thm "A_into_univ";
-val singleton_in_univ = thm "singleton_in_univ";
-val doubleton_in_univ = thm "doubleton_in_univ";
-val Pair_in_univ = thm "Pair_in_univ";
-val Union_in_univ = thm "Union_in_univ";
-val product_univ = thm "product_univ";
-val nat_subset_univ = thm "nat_subset_univ";
-val nat_into_univ = thm "nat_into_univ";
-val one_in_univ = thm "one_in_univ";
-val two_in_univ = thm "two_in_univ";
-val bool_subset_univ = thm "bool_subset_univ";
-val bool_into_univ = thm "bool_into_univ";
-val Inl_in_univ = thm "Inl_in_univ";
-val Inr_in_univ = thm "Inr_in_univ";
-val sum_univ = thm "sum_univ";
-val sum_subset_univ = thm "sum_subset_univ";
-val Fin_VLimit = thm "Fin_VLimit";
-val Fin_subset_VLimit = thm "Fin_subset_VLimit";
-val Fin_univ = thm "Fin_univ";
-val nat_fun_VLimit = thm "nat_fun_VLimit";
-val nat_fun_subset_VLimit = thm "nat_fun_subset_VLimit";
-val nat_fun_univ = thm "nat_fun_univ";
-val FiniteFun_VLimit1 = thm "FiniteFun_VLimit1";
-val FiniteFun_univ1 = thm "FiniteFun_univ1";
-val FiniteFun_VLimit = thm "FiniteFun_VLimit";
-val FiniteFun_univ = thm "FiniteFun_univ";
-val FiniteFun_in_univ = thm "FiniteFun_in_univ";
-val FiniteFun_in_univ' = thm "FiniteFun_in_univ'";
-val doubleton_in_Vfrom_D = thm "doubleton_in_Vfrom_D";
-val Vfrom_doubleton_D = thm "Vfrom_doubleton_D";
-val Pair_in_Vfrom_D = thm "Pair_in_Vfrom_D";
-val product_Int_Vfrom_subset = thm "product_Int_Vfrom_subset";
-
-val rank_rls = thms "rank_rls";
-val rank_ss = simpset() addsimps [VsetI] 
-              addsimps rank_rls @ (rank_rls RLN (2, [lt_trans]));
-
+val rank_ss = @{simpset} addsimps [@{thm VsetI}] 
+              addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}]));
 *}
 
 end
--- a/src/ZF/WF.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/WF.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -19,29 +19,35 @@
 
 theory WF imports Trancl begin
 
-constdefs
-  wf           :: "i=>o"
+definition
+  wf           :: "i=>o"  where
     (*r is a well-founded relation*)
     "wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. <y,x>:r --> ~ y:Z)"
 
-  wf_on        :: "[i,i]=>o"                      ("wf[_]'(_')")
+definition
+  wf_on        :: "[i,i]=>o"                      ("wf[_]'(_')")  where
     (*r is well-founded on A*)
     "wf_on(A,r) == wf(r Int A*A)"
 
-  is_recfun    :: "[i, i, [i,i]=>i, i] =>o"
+definition
+  is_recfun    :: "[i, i, [i,i]=>i, i] =>o"  where
     "is_recfun(r,a,H,f) == (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))"
 
-  the_recfun   :: "[i, i, [i,i]=>i] =>i"
+definition
+  the_recfun   :: "[i, i, [i,i]=>i] =>i"  where
     "the_recfun(r,a,H) == (THE f. is_recfun(r,a,H,f))"
 
-  wftrec :: "[i, i, [i,i]=>i] =>i"
+definition
+  wftrec :: "[i, i, [i,i]=>i] =>i"  where
     "wftrec(r,a,H) == H(a, the_recfun(r,a,H))"
 
-  wfrec :: "[i, i, [i,i]=>i] =>i"
+definition
+  wfrec :: "[i, i, [i,i]=>i] =>i"  where
     (*public version.  Does not require r to be transitive*)
     "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))"
 
-  wfrec_on     :: "[i, i, i, [i,i]=>i] =>i"       ("wfrec[_]'(_,_,_')")
+definition
+  wfrec_on     :: "[i, i, i, [i,i]=>i] =>i"       ("wfrec[_]'(_,_,_')")  where
     "wfrec[A](r,a,H) == wfrec(r Int A*A, a, H)"
 
 
@@ -362,49 +368,4 @@
      "wf(r) <-> (ALL Q x. x:Q --> (EX z:Q. ALL y. <y,z>:r --> y~:Q))"
 by (unfold wf_def, blast)
 
-
-ML
-{*
-val wf_def = thm "wf_def";
-val wf_on_def = thm "wf_on_def";
-
-val wf_imp_wf_on = thm "wf_imp_wf_on";
-val wf_on_field_imp_wf = thm "wf_on_field_imp_wf";
-val wf_iff_wf_on_field = thm "wf_iff_wf_on_field";
-val wf_on_subset_A = thm "wf_on_subset_A";
-val wf_on_subset_r = thm "wf_on_subset_r";
-val wf_onI = thm "wf_onI";
-val wf_onI2 = thm "wf_onI2";
-val wf_induct = thm "wf_induct";
-val wf_induct2 = thm "wf_induct2";
-val field_Int_square = thm "field_Int_square";
-val wf_on_induct = thm "wf_on_induct";
-val wfI = thm "wfI";
-val wf_not_refl = thm "wf_not_refl";
-val wf_not_sym = thm "wf_not_sym";
-val wf_asym = thm "wf_asym";
-val wf_on_not_refl = thm "wf_on_not_refl";
-val wf_on_not_sym = thm "wf_on_not_sym";
-val wf_on_asym = thm "wf_on_asym";
-val wf_on_chain3 = thm "wf_on_chain3";
-val wf_on_trancl = thm "wf_on_trancl";
-val wf_trancl = thm "wf_trancl";
-val underI = thm "underI";
-val underD = thm "underD";
-val is_recfun_type = thm "is_recfun_type";
-val apply_recfun = thm "apply_recfun";
-val is_recfun_equal = thm "is_recfun_equal";
-val is_recfun_cut = thm "is_recfun_cut";
-val is_recfun_functional = thm "is_recfun_functional";
-val is_the_recfun = thm "is_the_recfun";
-val unfold_the_recfun = thm "unfold_the_recfun";
-val the_recfun_cut = thm "the_recfun_cut";
-val wftrec = thm "wftrec";
-val wfrec = thm "wfrec";
-val def_wfrec = thm "def_wfrec";
-val wfrec_type = thm "wfrec_type";
-val wfrec_on = thm "wfrec_on";
-val wf_eq_minimal = thm "wf_eq_minimal";
-*}
-
 end
--- a/src/ZF/ZF.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ZF.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -655,100 +655,14 @@
 lemma cantor: "\<exists>S \<in> Pow(A). \<forall>x\<in>A. b(x) ~= S"
 by (best elim!: equalityCE del: ReplaceI RepFun_eqI)
 
-ML
-{*
-val lam_def = thm "lam_def";
-val domain_def = thm "domain_def";
-val range_def = thm "range_def";
-val image_def = thm "image_def";
-val vimage_def = thm "vimage_def";
-val field_def = thm "field_def";
-val Inter_def = thm "Inter_def";
-val Ball_def = thm "Ball_def";
-val Bex_def = thm "Bex_def";
-
-val ballI = thm "ballI";
-val bspec = thm "bspec";
-val rev_ballE = thm "rev_ballE";
-val ballE = thm "ballE";
-val rev_bspec = thm "rev_bspec";
-val ball_triv = thm "ball_triv";
-val ball_cong = thm "ball_cong";
-val bexI = thm "bexI";
-val rev_bexI = thm "rev_bexI";
-val bexCI = thm "bexCI";
-val bexE = thm "bexE";
-val bex_triv = thm "bex_triv";
-val bex_cong = thm "bex_cong";
-val subst_elem = thm "subst_elem";
-val subsetI = thm "subsetI";
-val subsetD = thm "subsetD";
-val subsetCE = thm "subsetCE";
-val rev_subsetD = thm "rev_subsetD";
-val contra_subsetD = thm "contra_subsetD";
-val rev_contra_subsetD = thm "rev_contra_subsetD";
-val subset_refl = thm "subset_refl";
-val subset_trans = thm "subset_trans";
-val subset_iff = thm "subset_iff";
-val equalityI = thm "equalityI";
-val equality_iffI = thm "equality_iffI";
-val equalityD1 = thm "equalityD1";
-val equalityD2 = thm "equalityD2";
-val equalityE = thm "equalityE";
-val equalityCE = thm "equalityCE";
-val Replace_iff = thm "Replace_iff";
-val ReplaceI = thm "ReplaceI";
-val ReplaceE = thm "ReplaceE";
-val ReplaceE2 = thm "ReplaceE2";
-val Replace_cong = thm "Replace_cong";
-val RepFunI = thm "RepFunI";
-val RepFun_eqI = thm "RepFun_eqI";
-val RepFunE = thm "RepFunE";
-val RepFun_cong = thm "RepFun_cong";
-val RepFun_iff = thm "RepFun_iff";
-val triv_RepFun = thm "triv_RepFun";
-val separation = thm "separation";
-val CollectI = thm "CollectI";
-val CollectE = thm "CollectE";
-val CollectD1 = thm "CollectD1";
-val CollectD2 = thm "CollectD2";
-val Collect_cong = thm "Collect_cong";
-val UnionI = thm "UnionI";
-val UnionE = thm "UnionE";
-val UN_iff = thm "UN_iff";
-val UN_I = thm "UN_I";
-val UN_E = thm "UN_E";
-val UN_cong = thm "UN_cong";
-val Inter_iff = thm "Inter_iff";
-val InterI = thm "InterI";
-val InterD = thm "InterD";
-val InterE = thm "InterE";
-val INT_iff = thm "INT_iff";
-val INT_I = thm "INT_I";
-val INT_E = thm "INT_E";
-val INT_cong = thm "INT_cong";
-val PowI = thm "PowI";
-val PowD = thm "PowD";
-val Pow_bottom = thm "Pow_bottom";
-val Pow_top = thm "Pow_top";
-val not_mem_empty = thm "not_mem_empty";
-val emptyE = thm "emptyE";
-val empty_subsetI = thm "empty_subsetI";
-val equals0I = thm "equals0I";
-val equals0D = thm "equals0D";
-val not_emptyI = thm "not_emptyI";
-val not_emptyE = thm "not_emptyE";
-val cantor = thm "cantor";
-*}
-
 (*Functions for ML scripts*)
 ML
 {*
 (*Converts A<=B to x\<in>A ==> x\<in>B*)
-fun impOfSubs th = th RSN (2, rev_subsetD);
+fun impOfSubs th = th RSN (2, @{thm rev_subsetD});
 
 (*Takes assumptions \<forall>x\<in>A.P(x) and a\<in>A; creates assumption P(a)*)
-val ball_tac = dtac bspec THEN' assume_tac
+val ball_tac = dtac @{thm bspec} THEN' assume_tac
 *}
 
 end
--- a/src/ZF/Zorn.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/Zorn.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -12,22 +12,24 @@
 text{*Based upon the unpublished article ``Towards the Mechanization of the
 Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.*}
 
-constdefs
-  Subset_rel :: "i=>i"
+definition
+  Subset_rel :: "i=>i"  where
    "Subset_rel(A) == {z \<in> A*A . \<exists>x y. z=<x,y> & x<=y & x\<noteq>y}"
 
-  chain      :: "i=>i"
+definition
+  chain      :: "i=>i"  where
    "chain(A)      == {F \<in> Pow(A). \<forall>X\<in>F. \<forall>Y\<in>F. X<=Y | Y<=X}"
 
-  super      :: "[i,i]=>i"
+definition
+  super      :: "[i,i]=>i"  where
    "super(A,c)    == {d \<in> chain(A). c<=d & c\<noteq>d}"
 
-  maxchain   :: "i=>i"
+definition
+  maxchain   :: "i=>i"  where
    "maxchain(A)   == {c \<in> chain(A). super(A,c)=0}"
 
-
-constdefs
-  increasing :: "i=>i"
+definition
+  increasing :: "i=>i"  where
     "increasing(A) == {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A --> x<=f`x}"
 
 
--- a/src/ZF/arith_data.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/arith_data.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -119,22 +119,22 @@
 
 
 (*Simplify #1*n and n*#1 to n*)
-val add_0s = [add_0_natify, add_0_right_natify];
-val add_succs = [add_succ, add_succ_right];
-val mult_1s = [mult_1_natify, mult_1_right_natify];
-val tc_rules = [natify_in_nat, add_type, diff_type, mult_type];
-val natifys = [natify_0, natify_ident, add_natify1, add_natify2,
-               diff_natify1, diff_natify2];
+val add_0s = [@{thm add_0_natify}, @{thm add_0_right_natify}];
+val add_succs = [@{thm add_succ}, @{thm add_succ_right}];
+val mult_1s = [@{thm mult_1_natify}, @{thm mult_1_right_natify}];
+val tc_rules = [@{thm natify_in_nat}, @{thm add_type}, @{thm diff_type}, @{thm mult_type}];
+val natifys = [@{thm natify_0}, @{thm natify_ident}, @{thm add_natify1}, @{thm add_natify2},
+               @{thm diff_natify1}, @{thm diff_natify2}];
 
 (*Final simplification: cancel + and **)
 fun simplify_meta_eq rules =
   let val ss0 =
-    FOL_ss addeqcongs [eq_cong2, iff_cong2]
+    FOL_ss addeqcongs [@{thm eq_cong2}, @{thm iff_cong2}]
       delsimps iff_simps (*these could erase the whole rule!*)
       addsimps rules
   in fn ss => mk_meta_eq o simplify (Simplifier.inherit_context ss ss0) end;
 
-val final_rules = add_0s @ mult_1s @ [mult_0, mult_0_right];
+val final_rules = add_0s @ mult_1s @ [@{thm mult_0}, @{thm mult_0_right}];
 
 structure CancelNumeralsCommon =
   struct
@@ -144,8 +144,9 @@
   val dest_coeff        = dest_coeff
   val find_first_coeff  = find_first_coeff []
 
-  val norm_ss1 = ZF_ss addsimps add_0s @ add_succs @ mult_1s @ add_ac
-  val norm_ss2 = ZF_ss addsimps add_0s @ mult_1s @ add_ac @ mult_ac @ tc_rules @ natifys
+  val norm_ss1 = ZF_ss addsimps add_0s @ add_succs @ mult_1s @ @{thms add_ac}
+  val norm_ss2 = ZF_ss addsimps add_0s @ mult_1s @ @{thms add_ac} @
+    @{thms mult_ac} @ tc_rules @ natifys
   fun norm_tac ss =
     ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1))
     THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2))
@@ -164,8 +165,8 @@
   val prove_conv = prove_conv "nateq_cancel_numerals"
   val mk_bal   = FOLogic.mk_eq
   val dest_bal = FOLogic.dest_eq
-  val bal_add1 = eq_add_iff RS iff_trans
-  val bal_add2 = eq_add_iff RS iff_trans
+  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
   end;
 
@@ -177,8 +178,8 @@
   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 = less_add_iff RS iff_trans
-  val bal_add2 = less_add_iff RS iff_trans
+  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
   end;
 
@@ -190,8 +191,8 @@
   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 = diff_add_eq RS trans
-  val bal_add2 = diff_add_eq RS trans
+  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
   end;
 
--- a/src/ZF/equalities.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/equalities.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -971,279 +971,18 @@
                     Inter_greatest Int_greatest RepFun_subset
                     Un_upper1 Un_upper2 Int_lower1 Int_lower2
 
-(*First, ML bindings from the old file subset.ML*)
-ML
-{*
-val cons_subsetI = thm "cons_subsetI";
-val subset_consI = thm "subset_consI";
-val cons_subset_iff = thm "cons_subset_iff";
-val cons_subsetE = thm "cons_subsetE";
-val subset_empty_iff = thm "subset_empty_iff";
-val subset_cons_iff = thm "subset_cons_iff";
-val subset_succI = thm "subset_succI";
-val succ_subsetI = thm "succ_subsetI";
-val succ_subsetE = thm "succ_subsetE";
-val succ_subset_iff = thm "succ_subset_iff";
-val singleton_subsetI = thm "singleton_subsetI";
-val singleton_subsetD = thm "singleton_subsetD";
-val Union_subset_iff = thm "Union_subset_iff";
-val Union_upper = thm "Union_upper";
-val Union_least = thm "Union_least";
-val subset_UN_iff_eq = thm "subset_UN_iff_eq";
-val UN_subset_iff = thm "UN_subset_iff";
-val UN_upper = thm "UN_upper";
-val UN_least = thm "UN_least";
-val Inter_subset_iff = thm "Inter_subset_iff";
-val Inter_lower = thm "Inter_lower";
-val Inter_greatest = thm "Inter_greatest";
-val INT_lower = thm "INT_lower";
-val INT_greatest = thm "INT_greatest";
-val Un_subset_iff = thm "Un_subset_iff";
-val Un_upper1 = thm "Un_upper1";
-val Un_upper2 = thm "Un_upper2";
-val Un_least = thm "Un_least";
-val Int_subset_iff = thm "Int_subset_iff";
-val Int_lower1 = thm "Int_lower1";
-val Int_lower2 = thm "Int_lower2";
-val Int_greatest = thm "Int_greatest";
-val Diff_subset = thm "Diff_subset";
-val Diff_contains = thm "Diff_contains";
-val subset_Diff_cons_iff = thm "subset_Diff_cons_iff";
-val Collect_subset = thm "Collect_subset";
-val RepFun_subset = thm "RepFun_subset";
-
-val subset_SIs = thms "subset_SIs";
-
-val subset_cs = claset() 
-    delrules [subsetI, subsetCE]
-    addSIs subset_SIs
-    addIs  [Union_upper, Inter_lower]
-    addSEs [cons_subsetE];
+ML {*
+val subset_cs = @{claset} 
+    delrules [@{thm subsetI}, @{thm subsetCE}]
+    addSIs @{thms subset_SIs}
+    addIs  [@{thm Union_upper}, @{thm Inter_lower}]
+    addSEs [@{thm cons_subsetE}];
 *}
 (*subset_cs is a claset for subset reasoning*)
 
 ML
 {*
-val ZF_cs = claset() delrules [equalityI];
-
-val in_mono = thm "in_mono";
-val conj_mono = thm "conj_mono";
-val disj_mono = thm "disj_mono";
-val imp_mono = thm "imp_mono";
-val imp_refl = thm "imp_refl";
-val ex_mono = thm "ex_mono";
-val all_mono = thm "all_mono";
-
-val converse_iff = thm "converse_iff";
-val converseI = thm "converseI";
-val converseD = thm "converseD";
-val converseE = thm "converseE";
-val converse_converse = thm "converse_converse";
-val converse_type = thm "converse_type";
-val converse_prod = thm "converse_prod";
-val converse_empty = thm "converse_empty";
-val converse_subset_iff = thm "converse_subset_iff";
-val domain_iff = thm "domain_iff";
-val domainI = thm "domainI";
-val domainE = thm "domainE";
-val domain_subset = thm "domain_subset";
-val rangeI = thm "rangeI";
-val rangeE = thm "rangeE";
-val range_subset = thm "range_subset";
-val fieldI1 = thm "fieldI1";
-val fieldI2 = thm "fieldI2";
-val fieldCI = thm "fieldCI";
-val fieldE = thm "fieldE";
-val field_subset = thm "field_subset";
-val domain_subset_field = thm "domain_subset_field";
-val range_subset_field = thm "range_subset_field";
-val domain_times_range = thm "domain_times_range";
-val field_times_field = thm "field_times_field";
-val image_iff = thm "image_iff";
-val image_singleton_iff = thm "image_singleton_iff";
-val imageI = thm "imageI";
-val imageE = thm "imageE";
-val image_subset = thm "image_subset";
-val vimage_iff = thm "vimage_iff";
-val vimage_singleton_iff = thm "vimage_singleton_iff";
-val vimageI = thm "vimageI";
-val vimageE = thm "vimageE";
-val vimage_subset = thm "vimage_subset";
-val rel_Union = thm "rel_Union";
-val rel_Un = thm "rel_Un";
-val domain_Diff_eq = thm "domain_Diff_eq";
-val range_Diff_eq = thm "range_Diff_eq";
-val cons_eq = thm "cons_eq";
-val cons_commute = thm "cons_commute";
-val cons_absorb = thm "cons_absorb";
-val cons_Diff = thm "cons_Diff";
-val equal_singleton = thm "equal_singleton";
-val Int_cons = thm "Int_cons";
-val Int_absorb = thm "Int_absorb";
-val Int_left_absorb = thm "Int_left_absorb";
-val Int_commute = thm "Int_commute";
-val Int_left_commute = thm "Int_left_commute";
-val Int_assoc = thm "Int_assoc";
-val Int_Un_distrib = thm "Int_Un_distrib";
-val Int_Un_distrib2 = thm "Int_Un_distrib2";
-val subset_Int_iff = thm "subset_Int_iff";
-val subset_Int_iff2 = thm "subset_Int_iff2";
-val Int_Diff_eq = thm "Int_Diff_eq";
-val Int_cons_left = thm "Int_cons_left";
-val Int_cons_right = thm "Int_cons_right";
-val Un_cons = thm "Un_cons";
-val Un_absorb = thm "Un_absorb";
-val Un_left_absorb = thm "Un_left_absorb";
-val Un_commute = thm "Un_commute";
-val Un_left_commute = thm "Un_left_commute";
-val Un_assoc = thm "Un_assoc";
-val Un_Int_distrib = thm "Un_Int_distrib";
-val subset_Un_iff = thm "subset_Un_iff";
-val subset_Un_iff2 = thm "subset_Un_iff2";
-val Un_empty = thm "Un_empty";
-val Un_eq_Union = thm "Un_eq_Union";
-val Diff_cancel = thm "Diff_cancel";
-val Diff_triv = thm "Diff_triv";
-val empty_Diff = thm "empty_Diff";
-val Diff_0 = thm "Diff_0";
-val Diff_eq_0_iff = thm "Diff_eq_0_iff";
-val Diff_cons = thm "Diff_cons";
-val Diff_cons2 = thm "Diff_cons2";
-val Diff_disjoint = thm "Diff_disjoint";
-val Diff_partition = thm "Diff_partition";
-val subset_Un_Diff = thm "subset_Un_Diff";
-val double_complement = thm "double_complement";
-val double_complement_Un = thm "double_complement_Un";
-val Un_Int_crazy = thm "Un_Int_crazy";
-val Diff_Un = thm "Diff_Un";
-val Diff_Int = thm "Diff_Int";
-val Un_Diff = thm "Un_Diff";
-val Int_Diff = thm "Int_Diff";
-val Diff_Int_distrib = thm "Diff_Int_distrib";
-val Diff_Int_distrib2 = thm "Diff_Int_distrib2";
-val Un_Int_assoc_iff = thm "Un_Int_assoc_iff";
-val Union_cons = thm "Union_cons";
-val Union_Un_distrib = thm "Union_Un_distrib";
-val Union_Int_subset = thm "Union_Int_subset";
-val Union_disjoint = thm "Union_disjoint";
-val Union_empty_iff = thm "Union_empty_iff";
-val Int_Union2 = thm "Int_Union2";
-val Inter_0 = thm "Inter_0";
-val Inter_Un_subset = thm "Inter_Un_subset";
-val Inter_Un_distrib = thm "Inter_Un_distrib";
-val Union_singleton = thm "Union_singleton";
-val Inter_singleton = thm "Inter_singleton";
-val Inter_cons = thm "Inter_cons";
-val Union_eq_UN = thm "Union_eq_UN";
-val Inter_eq_INT = thm "Inter_eq_INT";
-val UN_0 = thm "UN_0";
-val UN_singleton = thm "UN_singleton";
-val UN_Un = thm "UN_Un";
-val INT_Un = thm "INT_Un";
-val UN_UN_flatten = thm "UN_UN_flatten";
-val Int_UN_distrib = thm "Int_UN_distrib";
-val Un_INT_distrib = thm "Un_INT_distrib";
-val Int_UN_distrib2 = thm "Int_UN_distrib2";
-val Un_INT_distrib2 = thm "Un_INT_distrib2";
-val UN_constant = thm "UN_constant";
-val INT_constant = thm "INT_constant";
-val UN_RepFun = thm "UN_RepFun";
-val INT_RepFun = thm "INT_RepFun";
-val INT_Union_eq = thm "INT_Union_eq";
-val INT_UN_eq = thm "INT_UN_eq";
-val UN_Un_distrib = thm "UN_Un_distrib";
-val INT_Int_distrib = thm "INT_Int_distrib";
-val UN_Int_subset = thm "UN_Int_subset";
-val Diff_UN = thm "Diff_UN";
-val Diff_INT = thm "Diff_INT";
-val Sigma_cons1 = thm "Sigma_cons1";
-val Sigma_cons2 = thm "Sigma_cons2";
-val Sigma_succ1 = thm "Sigma_succ1";
-val Sigma_succ2 = thm "Sigma_succ2";
-val SUM_UN_distrib1 = thm "SUM_UN_distrib1";
-val SUM_UN_distrib2 = thm "SUM_UN_distrib2";
-val SUM_Un_distrib1 = thm "SUM_Un_distrib1";
-val SUM_Un_distrib2 = thm "SUM_Un_distrib2";
-val prod_Un_distrib2 = thm "prod_Un_distrib2";
-val SUM_Int_distrib1 = thm "SUM_Int_distrib1";
-val SUM_Int_distrib2 = thm "SUM_Int_distrib2";
-val prod_Int_distrib2 = thm "prod_Int_distrib2";
-val SUM_eq_UN = thm "SUM_eq_UN";
-val domain_of_prod = thm "domain_of_prod";
-val domain_0 = thm "domain_0";
-val domain_cons = thm "domain_cons";
-val domain_Un_eq = thm "domain_Un_eq";
-val domain_Int_subset = thm "domain_Int_subset";
-val domain_Diff_subset = thm "domain_Diff_subset";
-val domain_converse = thm "domain_converse";
-val domain_UN = thm "domain_UN";
-val domain_Union = thm "domain_Union";
-val range_of_prod = thm "range_of_prod";
-val range_0 = thm "range_0";
-val range_cons = thm "range_cons";
-val range_Un_eq = thm "range_Un_eq";
-val range_Int_subset = thm "range_Int_subset";
-val range_Diff_subset = thm "range_Diff_subset";
-val range_converse = thm "range_converse";
-val field_of_prod = thm "field_of_prod";
-val field_0 = thm "field_0";
-val field_cons = thm "field_cons";
-val field_Un_eq = thm "field_Un_eq";
-val field_Int_subset = thm "field_Int_subset";
-val field_Diff_subset = thm "field_Diff_subset";
-val field_converse = thm "field_converse";
-val image_0 = thm "image_0";
-val image_Un = thm "image_Un";
-val image_Int_subset = thm "image_Int_subset";
-val image_Int_square_subset = thm "image_Int_square_subset";
-val image_Int_square = thm "image_Int_square";
-val image_0_left = thm "image_0_left";
-val image_Un_left = thm "image_Un_left";
-val image_Int_subset_left = thm "image_Int_subset_left";
-val vimage_0 = thm "vimage_0";
-val vimage_Un = thm "vimage_Un";
-val vimage_Int_subset = thm "vimage_Int_subset";
-val vimage_eq_UN = thm "vimage_eq_UN";
-val function_vimage_Int = thm "function_vimage_Int";
-val function_vimage_Diff = thm "function_vimage_Diff";
-val function_image_vimage = thm "function_image_vimage";
-val vimage_Int_square_subset = thm "vimage_Int_square_subset";
-val vimage_Int_square = thm "vimage_Int_square";
-val vimage_0_left = thm "vimage_0_left";
-val vimage_Un_left = thm "vimage_Un_left";
-val vimage_Int_subset_left = thm "vimage_Int_subset_left";
-val converse_Un = thm "converse_Un";
-val converse_Int = thm "converse_Int";
-val converse_Diff = thm "converse_Diff";
-val converse_UN = thm "converse_UN";
-val converse_INT = thm "converse_INT";
-val Pow_0 = thm "Pow_0";
-val Pow_insert = thm "Pow_insert";
-val Un_Pow_subset = thm "Un_Pow_subset";
-val UN_Pow_subset = thm "UN_Pow_subset";
-val subset_Pow_Union = thm "subset_Pow_Union";
-val Union_Pow_eq = thm "Union_Pow_eq";
-val Union_Pow_iff = thm "Union_Pow_iff";
-val Pow_Int_eq = thm "Pow_Int_eq";
-val Pow_INT_eq = thm "Pow_INT_eq";
-val RepFun_eq_0_iff = thm "RepFun_eq_0_iff";
-val RepFun_constant = thm "RepFun_constant";
-val Collect_Un = thm "Collect_Un";
-val Collect_Int = thm "Collect_Int";
-val Collect_Diff = thm "Collect_Diff";
-val Collect_cons = thm "Collect_cons";
-val Int_Collect_self_eq = thm "Int_Collect_self_eq";
-val Collect_Collect_eq = thm "Collect_Collect_eq";
-val Collect_Int_Collect_eq = thm "Collect_Int_Collect_eq";
-val Collect_disj_eq = thm "Collect_disj_eq";
-val Collect_conj_eq = thm "Collect_conj_eq";
-
-val Int_ac = thms "Int_ac";
-val Un_ac = thms "Un_ac";
-val Int_absorb1 = thm "Int_absorb1";
-val Int_absorb2 = thm "Int_absorb2";
-val Un_absorb1 = thm "Un_absorb1";
-val Un_absorb2 = thm "Un_absorb2";
+val ZF_cs = @{claset} delrules [@{thm equalityI}];
 *}
  
 end
--- a/src/ZF/ex/CoUnit.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ex/CoUnit.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -82,8 +82,8 @@
   apply (erule counit2.cases)
   apply (unfold counit2.con_defs)
   apply (tactic {* fast_tac (subset_cs
-    addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]
-    addSEs [Ord_in_Ord, Pair_inject]) 1 *})
+    addSIs [@{thm QPair_Int_Vset_subset_UN} RS @{thm subset_trans}, @{thm QPair_mono}]
+    addSEs [@{thm Ord_in_Ord}, @{thm Pair_inject}]) 1 *})
   done
 
 lemma counit2_implies_equal: "[| x \<in> counit2;  y \<in> counit2 |] ==> x = y"
--- a/src/ZF/ex/LList.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ex/LList.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -40,13 +40,12 @@
 
 
 (*Lazy list functions; flip is not definitional!*)
-consts
-  lconst   :: "i => i"
-  flip     :: "i => i"
+definition
+  lconst   :: "i => i"  where
+  "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
 
-defs
-  lconst_def:  "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
-
+consts
+  flip     :: "i => i"
 axioms
   flip_LNil:   "flip(LNil) = LNil"
 
@@ -117,7 +116,8 @@
 apply (erule llist.cases)
 apply (simp_all add: QInl_def QInr_def llist.con_defs)
 (*LCons case: I simply can't get rid of the deepen_tac*)
-apply (tactic "deepen_tac (claset() addIs [Ord_trans] addIs [Int_lower1 RS subset_trans]) 2 1")
+apply (tactic "deepen_tac (@{claset} addIs [@{thm Ord_trans}]
+  addIs [@{thm Int_lower1} RS @{thm subset_trans}]) 2 1")
 done
 
 lemma llist_quniv: "llist(quniv(A)) \<subseteq> quniv(A)"
@@ -219,7 +219,8 @@
 apply (erule llist.cases, simp_all)
 apply (simp_all add: QInl_def QInr_def llist.con_defs)
 (*LCons case: I simply can't get rid of the deepen_tac*)
-apply (tactic "deepen_tac (claset() addIs [Ord_trans] addIs [Int_lower1 RS subset_trans]) 2 1")
+apply (tactic "deepen_tac (@{claset} addIs [@{thm Ord_trans}]
+  addIs [@{thm Int_lower1} RS @{thm subset_trans}]) 2 1")
 done
 
 lemma flip_in_quniv: "l \<in> llist(bool) ==> flip(l) \<in> quniv(bool)"
--- a/src/ZF/ex/Limit.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ex/Limit.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -21,146 +21,173 @@
 
 theory Limit  imports  Main begin
 
-constdefs
-
-  rel :: "[i,i,i]=>o"
+definition
+  rel :: "[i,i,i]=>o"  where
     "rel(D,x,y) == <x,y>:snd(D)"
 
-  set :: "i=>i"
+definition
+  set :: "i=>i"  where
     "set(D) == fst(D)"
 
-  po  :: "i=>o"
+definition
+  po  :: "i=>o"  where
     "po(D) ==
      (\<forall>x \<in> set(D). rel(D,x,x)) &
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). \<forall>z \<in> set(D).
        rel(D,x,y) --> rel(D,y,z) --> rel(D,x,z)) &
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) --> rel(D,y,x) --> x = y)"
 
-  chain :: "[i,i]=>o"
+definition
+  chain :: "[i,i]=>o"  where
     (* Chains are object level functions nat->set(D) *)
     "chain(D,X) == X \<in> nat->set(D) & (\<forall>n \<in> nat. rel(D,X`n,X`(succ(n))))"
 
-  isub :: "[i,i,i]=>o"
+definition
+  isub :: "[i,i,i]=>o"  where
     "isub(D,X,x) == x \<in> set(D) & (\<forall>n \<in> nat. rel(D,X`n,x))"
 
-  islub :: "[i,i,i]=>o"
+definition
+  islub :: "[i,i,i]=>o"  where
     "islub(D,X,x) == isub(D,X,x) & (\<forall>y. isub(D,X,y) --> rel(D,x,y))"
 
-  lub :: "[i,i]=>i"
+definition
+  lub :: "[i,i]=>i"  where
     "lub(D,X) == THE x. islub(D,X,x)"
 
-  cpo :: "i=>o"
+definition
+  cpo :: "i=>o"  where
     "cpo(D) == po(D) & (\<forall>X. chain(D,X) --> (\<exists>x. islub(D,X,x)))"
 
-  pcpo :: "i=>o"
+definition
+  pcpo :: "i=>o"  where
     "pcpo(D) == cpo(D) & (\<exists>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y))"
 
-  bot :: "i=>i"
+definition
+  bot :: "i=>i"  where
     "bot(D) == THE x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
 
-  mono :: "[i,i]=>i"
+definition
+  mono :: "[i,i]=>i"  where
     "mono(D,E) ==
      {f \<in> set(D)->set(E).
       \<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) --> rel(E,f`x,f`y)}"
 
-  cont :: "[i,i]=>i"
+definition
+  cont :: "[i,i]=>i"  where
     "cont(D,E) ==
      {f \<in> mono(D,E).
       \<forall>X. chain(D,X) --> f`(lub(D,X)) = lub(E,\<lambda>n \<in> nat. f`(X`n))}"
 
-  cf :: "[i,i]=>i"
+definition
+  cf :: "[i,i]=>i"  where
     "cf(D,E) ==
      <cont(D,E),
       {y \<in> cont(D,E)*cont(D,E). \<forall>x \<in> set(D). rel(E,(fst(y))`x,(snd(y))`x)}>"
 
-  suffix :: "[i,i]=>i"
+definition
+  suffix :: "[i,i]=>i"  where
     "suffix(X,n) == \<lambda>m \<in> nat. X`(n #+ m)"
 
-  subchain :: "[i,i]=>o"
+definition
+  subchain :: "[i,i]=>o"  where
     "subchain(X,Y) == \<forall>m \<in> nat. \<exists>n \<in> nat. X`m = Y`(m #+ n)"
 
-  dominate :: "[i,i,i]=>o"
+definition
+  dominate :: "[i,i,i]=>o"  where
     "dominate(D,X,Y) == \<forall>m \<in> nat. \<exists>n \<in> nat. rel(D,X`m,Y`n)"
 
-  matrix :: "[i,i]=>o"
+definition
+  matrix :: "[i,i]=>o"  where
     "matrix(D,M) ==
      M \<in> nat -> (nat -> set(D)) &
      (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`m)) &
      (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`n`succ(m))) &
      (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`succ(m)))"
 
-  projpair  :: "[i,i,i,i]=>o"
+definition
+  projpair  :: "[i,i,i,i]=>o"  where
     "projpair(D,E,e,p) ==
      e \<in> cont(D,E) & p \<in> cont(E,D) &
      p O e = id(set(D)) & rel(cf(E,E),e O p,id(set(E)))"
 
-  emb       :: "[i,i,i]=>o"
+definition
+  emb       :: "[i,i,i]=>o"  where
     "emb(D,E,e) == \<exists>p. projpair(D,E,e,p)"
 
-  Rp        :: "[i,i,i]=>i"
+definition
+  Rp        :: "[i,i,i]=>i"  where
     "Rp(D,E,e) == THE p. projpair(D,E,e,p)"
 
+definition
   (* Twice, constructions on cpos are more difficult. *)
-  iprod     :: "i=>i"
+  iprod     :: "i=>i"  where
     "iprod(DD) ==
      <(\<Pi> n \<in> nat. set(DD`n)),
       {x:(\<Pi> n \<in> nat. set(DD`n))*(\<Pi> n \<in> nat. set(DD`n)).
        \<forall>n \<in> nat. rel(DD`n,fst(x)`n,snd(x)`n)}>"
 
-  mkcpo     :: "[i,i=>o]=>i"
+definition
+  mkcpo     :: "[i,i=>o]=>i"  where
     (* Cannot use rel(D), is meta fun, need two more args *)
     "mkcpo(D,P) ==
      <{x \<in> set(D). P(x)},{x \<in> set(D)*set(D). rel(D,fst(x),snd(x))}>"
 
-  subcpo    :: "[i,i]=>o"
+definition
+  subcpo    :: "[i,i]=>o"  where
     "subcpo(D,E) ==
      set(D) \<subseteq> set(E) &
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) <-> rel(E,x,y)) &
      (\<forall>X. chain(D,X) --> lub(E,X):set(D))"
 
-  subpcpo   :: "[i,i]=>o"
+definition
+  subpcpo   :: "[i,i]=>o"  where
     "subpcpo(D,E) == subcpo(D,E) & bot(E):set(D)"
 
-  emb_chain :: "[i,i]=>o"
+definition
+  emb_chain :: "[i,i]=>o"  where
     "emb_chain(DD,ee) ==
      (\<forall>n \<in> nat. cpo(DD`n)) & (\<forall>n \<in> nat. emb(DD`n,DD`succ(n),ee`n))"
 
-  Dinf      :: "[i,i]=>i"
+definition
+  Dinf      :: "[i,i]=>i"  where
     "Dinf(DD,ee) ==
      mkcpo(iprod(DD))
      (%x. \<forall>n \<in> nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)"
 
-consts
-  e_less    :: "[i,i,i,i]=>i"
-  e_gr      :: "[i,i,i,i]=>i"
-
-defs  (*???NEEDS PRIMREC*)
-
-  e_less_def: (* Valid for m le n only. *)
+definition
+  e_less    :: "[i,i,i,i]=>i"  where
+  (* Valid for m le n only. *)
     "e_less(DD,ee,m,n) == rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)"
 
-  e_gr_def: (* Valid for n le m only. *)
+
+definition
+  e_gr      :: "[i,i,i,i]=>i"  where
+  (* Valid for n le m only. *)
     "e_gr(DD,ee,m,n) ==
      rec(m#-n,id(set(DD`n)),
          %x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))"
 
 
-constdefs
-  eps       :: "[i,i,i,i]=>i"
+definition
+  eps       :: "[i,i,i,i]=>i"  where
     "eps(DD,ee,m,n) == if(m le n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))"
 
-  rho_emb   :: "[i,i,i]=>i"
+definition
+  rho_emb   :: "[i,i,i]=>i"  where
     "rho_emb(DD,ee,n) == \<lambda>x \<in> set(DD`n). \<lambda>m \<in> nat. eps(DD,ee,n,m)`x"
 
-  rho_proj  :: "[i,i,i]=>i"
+definition
+  rho_proj  :: "[i,i,i]=>i"  where
     "rho_proj(DD,ee,n) == \<lambda>x \<in> set(Dinf(DD,ee)). x`n"
 
-  commute   :: "[i,i,i,i=>i]=>o"
+definition
+  commute   :: "[i,i,i,i=>i]=>o"  where
     "commute(DD,ee,E,r) ==
      (\<forall>n \<in> nat. emb(DD`n,E,r(n))) &
      (\<forall>m \<in> nat. \<forall>n \<in> nat. m le n --> r(n) O eps(DD,ee,m,n) = r(m))"
 
-  mediating :: "[i,i,i=>i,i=>i,i]=>o"
+definition
+  mediating :: "[i,i,i=>i,i=>i,i]=>o"  where
     "mediating(E,G,r,f,t) == emb(E,G,t) & (\<forall>n \<in> nat. f(n) = t O r(n))"
 
 
--- a/src/ZF/ex/Primes.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ex/Primes.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -7,23 +7,28 @@
 header{*The Divides Relation and Euclid's algorithm for the GCD*}
 
 theory Primes imports Main begin
-constdefs
-  divides :: "[i,i]=>o"              (infixl "dvd" 50) 
+
+definition
+  divides :: "[i,i]=>o"              (infixl "dvd" 50)  where
     "m dvd n == m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
 
-  is_gcd  :: "[i,i,i]=>o"     --{*definition of great common divisor*}
+definition
+  is_gcd  :: "[i,i,i]=>o"     --{*definition of great common divisor*}  where
     "is_gcd(p,m,n) == ((p dvd m) & (p dvd n))   &
                        (\<forall>d\<in>nat. (d dvd m) & (d dvd n) --> d dvd p)"
 
-  gcd     :: "[i,i]=>i"       --{*Euclid's algorithm for the gcd*}
+definition
+  gcd     :: "[i,i]=>i"       --{*Euclid's algorithm for the gcd*}  where
     "gcd(m,n) == transrec(natify(n),
 			%n f. \<lambda>m \<in> nat.
 			        if n=0 then m else f`(m mod n)`n) ` natify(m)"
 
-  coprime :: "[i,i]=>o"       --{*the coprime relation*}
+definition
+  coprime :: "[i,i]=>o"       --{*the coprime relation*}  where
     "coprime(m,n) == gcd(m,n) = 1"
   
-  prime   :: i                --{*the set of prime numbers*}
+definition
+  prime   :: i                --{*the set of prime numbers*}  where
    "prime == {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p --> m=1 | m=p)}"
 
 
--- a/src/ZF/func.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/func.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -443,8 +443,8 @@
 
 subsection{*Function Updates*}
 
-constdefs
-  update  :: "[i,i,i] => i"
+definition
+  update  :: "[i,i,i] => i"  where
    "update(f,a,b) == lam x: cons(a, domain(f)). if(x=a, b, f`x)"
 
 nonterminals
@@ -461,7 +461,7 @@
 
 translations
   "_Update (f, _updbinds(b,bs))"  == "_Update (_Update(f,b), bs)"
-  "f(x:=y)"                       == "update(f,x,y)"
+  "f(x:=y)"                       == "CONST update(f,x,y)"
 
 
 lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)"
@@ -592,114 +592,4 @@
 lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono 
                      Collect_mono Part_mono in_mono
 
-ML
-{*
-val Pi_iff = thm "Pi_iff";
-val Pi_iff_old = thm "Pi_iff_old";
-val fun_is_function = thm "fun_is_function";
-val fun_is_rel = thm "fun_is_rel";
-val Pi_cong = thm "Pi_cong";
-val fun_weaken_type = thm "fun_weaken_type";
-val apply_equality2 = thm "apply_equality2";
-val function_apply_equality = thm "function_apply_equality";
-val apply_equality = thm "apply_equality";
-val apply_0 = thm "apply_0";
-val Pi_memberD = thm "Pi_memberD";
-val function_apply_Pair = thm "function_apply_Pair";
-val apply_Pair = thm "apply_Pair";
-val apply_type = thm "apply_type";
-val apply_funtype = thm "apply_funtype";
-val apply_iff = thm "apply_iff";
-val Pi_type = thm "Pi_type";
-val Pi_Collect_iff = thm "Pi_Collect_iff";
-val Pi_weaken_type = thm "Pi_weaken_type";
-val domain_type = thm "domain_type";
-val range_type = thm "range_type";
-val Pair_mem_PiD = thm "Pair_mem_PiD";
-val lamI = thm "lamI";
-val lamE = thm "lamE";
-val lamD = thm "lamD";
-val lam_type = thm "lam_type";
-val lam_funtype = thm "lam_funtype";
-val beta = thm "beta";
-val lam_empty = thm "lam_empty";
-val domain_lam = thm "domain_lam";
-val lam_cong = thm "lam_cong";
-val lam_theI = thm "lam_theI";
-val lam_eqE = thm "lam_eqE";
-val Pi_empty1 = thm "Pi_empty1";
-val singleton_fun = thm "singleton_fun";
-val Pi_empty2 = thm "Pi_empty2";
-val fun_space_empty_iff = thm "fun_space_empty_iff";
-val fun_subset = thm "fun_subset";
-val fun_extension = thm "fun_extension";
-val eta = thm "eta";
-val fun_extension_iff = thm "fun_extension_iff";
-val fun_subset_eq = thm "fun_subset_eq";
-val Pi_lamE = thm "Pi_lamE";
-val image_lam = thm "image_lam";
-val image_fun = thm "image_fun";
-val Pi_image_cons = thm "Pi_image_cons";
-val restrict_subset = thm "restrict_subset";
-val function_restrictI = thm "function_restrictI";
-val restrict_type2 = thm "restrict_type2";
-val restrict = thm "restrict";
-val restrict_empty = thm "restrict_empty";
-val domain_restrict_lam = thm "domain_restrict_lam";
-val restrict_restrict = thm "restrict_restrict";
-val domain_restrict = thm "domain_restrict";
-val restrict_idem = thm "restrict_idem";
-val restrict_if = thm "restrict_if";
-val restrict_lam_eq = thm "restrict_lam_eq";
-val fun_cons_restrict_eq = thm "fun_cons_restrict_eq";
-val function_Union = thm "function_Union";
-val fun_Union = thm "fun_Union";
-val fun_disjoint_Un = thm "fun_disjoint_Un";
-val fun_disjoint_apply1 = thm "fun_disjoint_apply1";
-val fun_disjoint_apply2 = thm "fun_disjoint_apply2";
-val domain_of_fun = thm "domain_of_fun";
-val apply_rangeI = thm "apply_rangeI";
-val range_of_fun = thm "range_of_fun";
-val fun_extend = thm "fun_extend";
-val fun_extend3 = thm "fun_extend3";
-val fun_extend_apply = thm "fun_extend_apply";
-val singleton_apply = thm "singleton_apply";
-val cons_fun_eq = thm "cons_fun_eq";
-
-val update_def = thm "update_def";
-val update_apply = thm "update_apply";
-val update_idem = thm "update_idem";
-val domain_update = thm "domain_update";
-val update_type = thm "update_type";
-
-val Replace_mono = thm "Replace_mono";
-val RepFun_mono = thm "RepFun_mono";
-val Pow_mono = thm "Pow_mono";
-val Union_mono = thm "Union_mono";
-val UN_mono = thm "UN_mono";
-val Inter_anti_mono = thm "Inter_anti_mono";
-val cons_mono = thm "cons_mono";
-val Un_mono = thm "Un_mono";
-val Int_mono = thm "Int_mono";
-val Diff_mono = thm "Diff_mono";
-val Sigma_mono = thm "Sigma_mono";
-val sum_mono = thm "sum_mono";
-val Pi_mono = thm "Pi_mono";
-val lam_mono = thm "lam_mono";
-val converse_mono = thm "converse_mono";
-val domain_mono = thm "domain_mono";
-val domain_rel_subset = thm "domain_rel_subset";
-val range_mono = thm "range_mono";
-val range_rel_subset = thm "range_rel_subset";
-val field_mono = thm "field_mono";
-val field_rel_subset = thm "field_rel_subset";
-val image_pair_mono = thm "image_pair_mono";
-val vimage_pair_mono = thm "vimage_pair_mono";
-val image_mono = thm "image_mono";
-val vimage_mono = thm "vimage_mono";
-val Collect_mono = thm "Collect_mono";
-
-val basic_monos = thms "basic_monos";
-*}
-
 end
--- a/src/ZF/ind_syntax.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/ind_syntax.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -64,7 +64,7 @@
   read_instantiate replaces a propositional variable by a formula variable*)
 val equals_CollectD = 
     read_instantiate [("W","?Q")]
-        (make_elim (equalityD1 RS subsetD RS CollectD2));
+        (make_elim (@{thm equalityD1} RS @{thm subsetD} RS @{thm CollectD2}));
 
 
 (** For datatype definitions **)
--- a/src/ZF/int_arith.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/int_arith.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -11,38 +11,38 @@
     such as -x = #3
 **)
 
-Addsimps [inst "y" "integ_of(?w)" zminus_equation,
-          inst "x" "integ_of(?w)" equation_zminus];
+Addsimps [inst "y" "integ_of(?w)" @{thm zminus_equation},
+          inst "x" "integ_of(?w)" @{thm equation_zminus}];
 
-AddIffs [inst "y" "integ_of(?w)" zminus_zless,
-         inst "x" "integ_of(?w)" zless_zminus];
+AddIffs [inst "y" "integ_of(?w)" @{thm zminus_zless},
+         inst "x" "integ_of(?w)" @{thm zless_zminus}];
 
-AddIffs [inst "y" "integ_of(?w)" zminus_zle,
-         inst "x" "integ_of(?w)" zle_zminus];
+AddIffs [inst "y" "integ_of(?w)" @{thm zminus_zle},
+         inst "x" "integ_of(?w)" @{thm zle_zminus}];
 
-Addsimps [inst "s" "integ_of(?w)" (thm "Let_def")];
+Addsimps [inst "s" "integ_of(?w)" @{thm Let_def}];
 
 (*** Simprocs for numeric literals ***)
 
 (** Combining of literal coefficients in sums of products **)
 
 Goal "(x $< y) <-> (x$-y $< #0)";
-by (simp_tac (simpset() addsimps zcompare_rls) 1);
+by (simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
 qed "zless_iff_zdiff_zless_0";
 
 Goal "[| x: int; y: int |] ==> (x = y) <-> (x$-y = #0)";
-by (asm_simp_tac (simpset() addsimps zcompare_rls) 1);
+by (asm_simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
 qed "eq_iff_zdiff_eq_0";
 
 Goal "(x $<= y) <-> (x$-y $<= #0)";
-by (asm_simp_tac (simpset() addsimps zcompare_rls) 1);
+by (asm_simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
 qed "zle_iff_zdiff_zle_0";
 
 
 (** For combine_numerals **)
 
 Goal "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k";
-by (simp_tac (simpset() addsimps [zadd_zmult_distrib]@zadd_ac) 1);
+by (simp_tac (simpset() addsimps [@{thm zadd_zmult_distrib}]@ @{thms zadd_ac}) 1);
 qed "left_zadd_zmult_distrib";
 
 
@@ -56,37 +56,37 @@
                            zle_iff_zdiff_zle_0];
 
 Goal "(i$*u $+ m = j$*u $+ n) <-> ((i$-j)$*u $+ m = intify(n))";
-by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
-by (simp_tac (simpset() addsimps zcompare_rls) 1);
-by (simp_tac (simpset() addsimps zadd_ac) 1);
+by (simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]) 1);
+by (simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
+by (simp_tac (simpset() addsimps @{thms zadd_ac}) 1);
 qed "eq_add_iff1";
 
 Goal "(i$*u $+ m = j$*u $+ n) <-> (intify(m) = (j$-i)$*u $+ n)";
-by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
-by (simp_tac (simpset() addsimps zcompare_rls) 1);
-by (simp_tac (simpset() addsimps zadd_ac) 1);
+by (simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]) 1);
+by (simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
+by (simp_tac (simpset() addsimps @{thms zadd_ac}) 1);
 qed "eq_add_iff2";
 
 Goal "(i$*u $+ m $< j$*u $+ n) <-> ((i$-j)$*u $+ m $< n)";
-by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
-                                     zadd_ac@rel_iff_rel_0_rls) 1);
+by (asm_simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]@
+                                     @{thms zadd_ac} @ rel_iff_rel_0_rls) 1);
 qed "less_add_iff1";
 
 Goal "(i$*u $+ m $< j$*u $+ n) <-> (m $< (j$-i)$*u $+ n)";
-by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
-                                     zadd_ac@rel_iff_rel_0_rls) 1);
+by (asm_simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]@
+                                     @{thms zadd_ac} @ rel_iff_rel_0_rls) 1);
 qed "less_add_iff2";
 
 Goal "(i$*u $+ m $<= j$*u $+ n) <-> ((i$-j)$*u $+ m $<= n)";
-by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
-by (simp_tac (simpset() addsimps zcompare_rls) 1);
-by (simp_tac (simpset() addsimps zadd_ac) 1);
+by (simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]) 1);
+by (simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
+by (simp_tac (simpset() addsimps @{thms zadd_ac}) 1);
 qed "le_add_iff1";
 
 Goal "(i$*u $+ m $<= j$*u $+ n) <-> (m $<= (j$-i)$*u $+ n)";
-by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
-by (simp_tac (simpset() addsimps zcompare_rls) 1);
-by (simp_tac (simpset() addsimps zadd_ac) 1);
+by (simp_tac (simpset() addsimps [@{thm zdiff_def}, @{thm zadd_zmult_distrib}]) 1);
+by (simp_tac (simpset() addsimps @{thms zcompare_rls}) 1);
+by (simp_tac (simpset() addsimps @{thms zadd_ac}) 1);
 qed "le_add_iff2";
 
 
@@ -178,41 +178,41 @@
 
 
 (*Simplify #1*n and n*#1 to n*)
-val add_0s = [zadd_0_intify, zadd_0_right_intify];
+val add_0s = [@{thm zadd_0_intify}, @{thm zadd_0_right_intify}];
 
-val mult_1s = [zmult_1_intify, zmult_1_right_intify,
-               zmult_minus1, zmult_minus1_right];
+val mult_1s = [@{thm zmult_1_intify}, @{thm zmult_1_right_intify},
+               @{thm zmult_minus1}, @{thm zmult_minus1_right}];
 
-val tc_rules = [integ_of_type, intify_in_int,
-                int_of_type, zadd_type, zdiff_type, zmult_type] @ 
-               thms "bin.intros";
-val intifys = [intify_ident, zadd_intify1, zadd_intify2,
-               zdiff_intify1, zdiff_intify2, zmult_intify1, zmult_intify2,
-               zless_intify1, zless_intify2, zle_intify1, zle_intify2];
+val tc_rules = [@{thm integ_of_type}, @{thm intify_in_int},
+                @{thm int_of_type}, @{thm zadd_type}, @{thm zdiff_type}, @{thm zmult_type}] @ 
+               @{thms bin.intros};
+val intifys = [@{thm intify_ident}, @{thm zadd_intify1}, @{thm zadd_intify2},
+               @{thm zdiff_intify1}, @{thm zdiff_intify2}, @{thm zmult_intify1}, @{thm zmult_intify2},
+               @{thm zless_intify1}, @{thm zless_intify2}, @{thm zle_intify1}, @{thm zle_intify2}];
 
 (*To perform binary arithmetic*)
-val bin_simps = [add_integ_of_left] @ bin_arith_simps @ bin_rel_simps;
+val bin_simps = [@{thm add_integ_of_left}] @ @{thms bin_arith_simps} @ @{thms bin_rel_simps};
 
 (*To evaluate binary negations of coefficients*)
-val zminus_simps = NCons_simps @
-                   [integ_of_minus RS sym,
-                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
-                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
+val zminus_simps = @{thms NCons_simps} @
+                   [@{thm integ_of_minus} RS 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}];
 
 (*To let us treat subtraction as addition*)
-val diff_simps = [zdiff_def, zminus_zadd_distrib, zminus_zminus];
+val diff_simps = [@{thm zdiff_def}, @{thm zminus_zadd_distrib}, @{thm zminus_zminus}];
 
 (*push the unary minus down: - x * y = x * - y *)
 val int_minus_mult_eq_1_to_2 =
-    [zmult_zminus, zmult_zminus_right RS sym] MRS trans |> standard;
+    [@{thm zmult_zminus}, @{thm zmult_zminus_right} RS sym] MRS trans |> standard;
 
 (*to extract again any uncancelled minuses*)
 val int_minus_from_mult_simps =
-    [zminus_zminus, zmult_zminus, zmult_zminus_right];
+    [@{thm zminus_zminus}, @{thm zmult_zminus}, @{thm zmult_zminus_right}];
 
 (*combine unary minus with numeric literals, however nested within a product*)
 val int_mult_minus_simps =
-    [zmult_assoc, zmult_zminus RS sym, int_minus_mult_eq_1_to_2];
+    [@{thm zmult_assoc}, @{thm zmult_zminus} RS sym, int_minus_mult_eq_1_to_2];
 
 fun prep_simproc (name, pats, proc) =
   Simplifier.simproc (the_context ()) name pats proc;
@@ -226,9 +226,9 @@
   val find_first_coeff  = find_first_coeff []
   fun trans_tac _       = ArithData.gen_trans_tac iff_trans
 
-  val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac
+  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
-  val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys
+  val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys
   fun norm_tac ss =
     ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1))
     THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2))
@@ -304,9 +304,9 @@
   val prove_conv        = prove_conv_nohyps "int_combine_numerals"
   fun trans_tac _       = ArithData.gen_trans_tac trans
 
-  val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac @ intifys
+  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
-  val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys
+  val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys
   fun norm_tac ss =
     ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1))
     THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2))
@@ -341,15 +341,15 @@
   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      = zmult_assoc RS sym RS trans
+  val left_distrib      = @{thm zmult_assoc} RS sym RS trans
   val prove_conv        = prove_conv_nohyps "int_combine_numerals_prod"
   fun trans_tac _       = ArithData.gen_trans_tac trans
 
 
 
 val norm_ss1 = ZF_ss addsimps mult_1s @ diff_simps @ zminus_simps
-  val norm_ss2 = ZF_ss addsimps [zmult_zminus_right RS sym] @
-    bin_simps @ zmult_ac @ tc_rules @ intifys
+  val norm_ss2 = ZF_ss addsimps [@{thm zmult_zminus_right} RS sym] @
+    bin_simps @ @{thms zmult_ac} @ tc_rules @ intifys
   fun norm_tac ss =
     ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1))
     THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2))
--- a/src/ZF/pair.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/pair.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -162,41 +162,6 @@
      "(\<forall>z \<in> Sigma(A,B). P(z)) <-> (\<forall>x \<in> A. \<forall>y \<in> B(x). P(<x,y>))"
 by blast
 
-ML
-{*
-val singleton_eq_iff = thm "singleton_eq_iff";
-val doubleton_eq_iff = thm "doubleton_eq_iff";
-val Pair_iff = thm "Pair_iff";
-val Pair_inject = thm "Pair_inject";
-val Pair_inject1 = thm "Pair_inject1";
-val Pair_inject2 = thm "Pair_inject2";
-val Pair_not_0 = thm "Pair_not_0";
-val Pair_neq_0 = thm "Pair_neq_0";
-val Pair_neq_fst = thm "Pair_neq_fst";
-val Pair_neq_snd = thm "Pair_neq_snd";
-val Sigma_iff = thm "Sigma_iff";
-val SigmaI = thm "SigmaI";
-val SigmaD1 = thm "SigmaD1";
-val SigmaD2 = thm "SigmaD2";
-val SigmaE = thm "SigmaE";
-val SigmaE2 = thm "SigmaE2";
-val Sigma_cong = thm "Sigma_cong";
-val Sigma_empty1 = thm "Sigma_empty1";
-val Sigma_empty2 = thm "Sigma_empty2";
-val Sigma_empty_iff = thm "Sigma_empty_iff";
-val fst_conv = thm "fst_conv";
-val snd_conv = thm "snd_conv";
-val fst_type = thm "fst_type";
-val snd_type = thm "snd_type";
-val Pair_fst_snd_eq = thm "Pair_fst_snd_eq";
-val split = thm "split";
-val split_type = thm "split_type";
-val expand_split = thm "expand_split";
-val splitI = thm "splitI";
-val splitE = thm "splitE";
-val splitD = thm "splitD";
-*}
-
 end
 
 
--- a/src/ZF/simpdata.ML	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/simpdata.ML	Sun Oct 07 21:19:31 2007 +0200
@@ -32,39 +32,39 @@
 
 (*Analyse a rigid formula*)
 val ZF_conn_pairs =
-  [("Ball",     [bspec]),
+  [("Ball",     [@{thm bspec}]),
    ("All",      [spec]),
    ("op -->",   [mp]),
    ("op &",     [conjunct1,conjunct2])];
 
 (*Analyse a:b, where b is rigid*)
 val ZF_mem_pairs =
-  [("Collect",  [CollectD1,CollectD2]),
-   (@{const_name Diff},     [DiffD1,DiffD2]),
-   (@{const_name Int},   [IntD1,IntD2])];
+  [("Collect",  [@{thm CollectD1}, @{thm CollectD2}]),
+   (@{const_name Diff},     [@{thm DiffD1}, @{thm DiffD2}]),
+   (@{const_name Int},   [@{thm IntD1}, @{thm IntD2}])];
 
 val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);
 
 change_simpset (fn ss =>
   ss setmksimps (map mk_eq o ZF_atomize o gen_all)
-  addcongs [if_weak_cong]);
+  addcongs [@{thm if_weak_cong}]);
 
 local
 
-val unfold_bex_tac = unfold_tac [Bex_def];
+val unfold_bex_tac = unfold_tac [@{thm Bex_def}];
 fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac;
 val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac;
 
-val unfold_ball_tac = unfold_tac [Ball_def];
+val unfold_ball_tac = unfold_tac [@{thm Ball_def}];
 fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac;
 val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac;
 
 in
 
-val defBEX_regroup = Simplifier.simproc (the_context ())
+val defBEX_regroup = Simplifier.simproc @{theory}
   "defined BEX" ["EX x:A. P(x) & Q(x)"] rearrange_bex;
 
-val defBALL_regroup = Simplifier.simproc (the_context ())
+val defBALL_regroup = Simplifier.simproc @{theory}
   "defined BALL" ["ALL x:A. P(x) --> Q(x)"] rearrange_ball;
 
 end;
@@ -72,4 +72,5 @@
 Addsimprocs [defBALL_regroup, defBEX_regroup];
 
 
-val ZF_ss = simpset();
+val ZF_ss = @{simpset};
+
--- a/src/ZF/upair.thy	Sun Oct 07 15:49:25 2007 +0200
+++ b/src/ZF/upair.thy	Sun Oct 07 21:19:31 2007 +0200
@@ -525,95 +525,4 @@
      "Inter({b}) = b"
 by blast+
 
-
-ML
-{*
-val Upair_iff = thm "Upair_iff";
-val UpairI1 = thm "UpairI1";
-val UpairI2 = thm "UpairI2";
-val UpairE = thm "UpairE";
-val Un_iff = thm "Un_iff";
-val UnI1 = thm "UnI1";
-val UnI2 = thm "UnI2";
-val UnE = thm "UnE";
-val UnE' = thm "UnE'";
-val UnCI = thm "UnCI";
-val Int_iff = thm "Int_iff";
-val IntI = thm "IntI";
-val IntD1 = thm "IntD1";
-val IntD2 = thm "IntD2";
-val IntE = thm "IntE";
-val Diff_iff = thm "Diff_iff";
-val DiffI = thm "DiffI";
-val DiffD1 = thm "DiffD1";
-val DiffD2 = thm "DiffD2";
-val DiffE = thm "DiffE";
-val cons_iff = thm "cons_iff";
-val consI1 = thm "consI1";
-val consI2 = thm "consI2";
-val consE = thm "consE";
-val consE' = thm "consE'";
-val consCI = thm "consCI";
-val cons_not_0 = thm "cons_not_0";
-val cons_neq_0 = thm "cons_neq_0";
-val singleton_iff = thm "singleton_iff";
-val singletonI = thm "singletonI";
-val singletonE = thm "singletonE";
-val the_equality = thm "the_equality";
-val the_equality2 = thm "the_equality2";
-val theI = thm "theI";
-val the_0 = thm "the_0";
-val theI2 = thm "theI2";
-val if_true = thm "if_true";
-val if_false = thm "if_false";
-val if_cong = thm "if_cong";
-val if_weak_cong = thm "if_weak_cong";
-val if_P = thm "if_P";
-val if_not_P = thm "if_not_P";
-val split_if = thm "split_if";
-val split_if_eq1 = thm "split_if_eq1";
-val split_if_eq2 = thm "split_if_eq2";
-val split_if_mem1 = thm "split_if_mem1";
-val split_if_mem2 = thm "split_if_mem2";
-val if_iff = thm "if_iff";
-val if_type = thm "if_type";
-val mem_asym = thm "mem_asym";
-val mem_irrefl = thm "mem_irrefl";
-val mem_not_refl = thm "mem_not_refl";
-val mem_imp_not_eq = thm "mem_imp_not_eq";
-val succ_iff = thm "succ_iff";
-val succI1 = thm "succI1";
-val succI2 = thm "succI2";
-val succE = thm "succE";
-val succCI = thm "succCI";
-val succ_not_0 = thm "succ_not_0";
-val succ_neq_0 = thm "succ_neq_0";
-val succ_subsetD = thm "succ_subsetD";
-val succ_neq_self = thm "succ_neq_self";
-val succ_inject_iff = thm "succ_inject_iff";
-val succ_inject = thm "succ_inject";
-
-val split_ifs = thms "split_ifs";
-val ball_simps = thms "ball_simps";
-val bex_simps = thms "bex_simps";
-
-val ball_conj_distrib = thm "ball_conj_distrib";
-val bex_disj_distrib = thm "bex_disj_distrib";
-val bex_triv_one_point1 = thm "bex_triv_one_point1";
-val bex_triv_one_point2 = thm "bex_triv_one_point2";
-val bex_one_point1 = thm "bex_one_point1";
-val bex_one_point2 = thm "bex_one_point2";
-val ball_one_point1 = thm "ball_one_point1";
-val ball_one_point2 = thm "ball_one_point2";
-
-val Rep_simps = thms "Rep_simps";
-val misc_simps = thms "misc_simps";
-
-val UN_simps = thms "UN_simps";
-val INT_simps = thms "INT_simps";
-
-val UN_extend_simps = thms "UN_extend_simps";
-val INT_extend_simps = thms "INT_extend_simps";
-*}
-
 end