merged
authorhuffman
Tue, 06 Sep 2011 22:41:35 -0700
changeset 44767 233f30abb040
parent 44766 d4d33a4d7548 (current diff)
parent 44763 b50d5d694838 (diff)
child 44768 a7bc1bdb8bb4
child 44837 859fc95860c5
merged
--- a/src/HOL/Import/HOL/HOL4Base.thy	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/HOL4Base.thy	Tue Sep 06 22:41:35 2011 -0700
@@ -4,277 +4,225 @@
 
 ;setup_theory bool
 
-definition ARB :: "'a" where 
-  "ARB == SOME x::'a::type. True"
-
-lemma ARB_DEF: "ARB = (SOME x::'a::type. True)"
-  by (import bool ARB_DEF)
-
-definition IN :: "'a => ('a => bool) => bool" where 
-  "IN == %(x::'a::type) f::'a::type => bool. f x"
-
-lemma IN_DEF: "IN = (%(x::'a::type) f::'a::type => bool. f x)"
-  by (import bool IN_DEF)
-
-definition RES_FORALL :: "('a => bool) => ('a => bool) => bool" where 
-  "RES_FORALL ==
-%(p::'a::type => bool) m::'a::type => bool. ALL x::'a::type. IN x p --> m x"
-
-lemma RES_FORALL_DEF: "RES_FORALL =
-(%(p::'a::type => bool) m::'a::type => bool.
-    ALL x::'a::type. IN x p --> m x)"
-  by (import bool RES_FORALL_DEF)
-
-definition RES_EXISTS :: "('a => bool) => ('a => bool) => bool" where 
-  "RES_EXISTS ==
-%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x"
-
-lemma RES_EXISTS_DEF: "RES_EXISTS =
-(%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x)"
-  by (import bool RES_EXISTS_DEF)
-
-definition RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" where 
+definition
+  ARB :: "'a"  where
+  "ARB == SOME x. True"
+
+lemma ARB_DEF: "ARB = (SOME x. True)"
+  sorry
+
+definition
+  IN :: "'a => ('a => bool) => bool"  where
+  "IN == %x f. f x"
+
+lemma IN_DEF: "IN = (%x f. f x)"
+  sorry
+
+definition
+  RES_FORALL :: "('a => bool) => ('a => bool) => bool"  where
+  "RES_FORALL == %p m. ALL x. IN x p --> m x"
+
+lemma RES_FORALL_DEF: "RES_FORALL = (%p m. ALL x. IN x p --> m x)"
+  sorry
+
+definition
+  RES_EXISTS :: "('a => bool) => ('a => bool) => bool"  where
+  "RES_EXISTS == %p m. EX x. IN x p & m x"
+
+lemma RES_EXISTS_DEF: "RES_EXISTS = (%p m. EX x. IN x p & m x)"
+  sorry
+
+definition
+  RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool"  where
   "RES_EXISTS_UNIQUE ==
-%(p::'a::type => bool) m::'a::type => bool.
-   RES_EXISTS p m &
-   RES_FORALL p
-    (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y))"
+%p m. RES_EXISTS p m &
+      RES_FORALL p (%x. RES_FORALL p (%y. m x & m y --> x = y))"
 
 lemma RES_EXISTS_UNIQUE_DEF: "RES_EXISTS_UNIQUE =
-(%(p::'a::type => bool) m::'a::type => bool.
-    RES_EXISTS p m &
-    RES_FORALL p
-     (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y)))"
-  by (import bool RES_EXISTS_UNIQUE_DEF)
-
-definition RES_SELECT :: "('a => bool) => ('a => bool) => 'a" where 
-  "RES_SELECT ==
-%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x"
-
-lemma RES_SELECT_DEF: "RES_SELECT =
-(%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x)"
-  by (import bool RES_SELECT_DEF)
-
-lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t"
-  by (import bool EXCLUDED_MIDDLE)
-
-lemma FORALL_THM: "All (f::'a::type => bool) = All f"
-  by (import bool FORALL_THM)
-
-lemma EXISTS_THM: "Ex (f::'a::type => bool) = Ex f"
-  by (import bool EXISTS_THM)
-
-lemma F_IMP: "ALL t::bool. ~ t --> t --> False"
-  by (import bool F_IMP)
-
-lemma NOT_AND: "~ ((t::bool) & ~ t)"
-  by (import bool NOT_AND)
-
-lemma AND_CLAUSES: "ALL t::bool.
-   (True & t) = t &
-   (t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
-  by (import bool AND_CLAUSES)
-
-lemma OR_CLAUSES: "ALL t::bool.
-   (True | t) = True &
-   (t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
-  by (import bool OR_CLAUSES)
-
-lemma IMP_CLAUSES: "ALL t::bool.
-   (True --> t) = t &
-   (t --> True) = True &
-   (False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
-  by (import bool IMP_CLAUSES)
-
-lemma NOT_CLAUSES: "(ALL t::bool. (~ ~ t) = t) & (~ True) = False & (~ False) = True"
-  by (import bool NOT_CLAUSES)
+(%p m. RES_EXISTS p m &
+       RES_FORALL p (%x. RES_FORALL p (%y. m x & m y --> x = y)))"
+  sorry
+
+definition
+  RES_SELECT :: "('a => bool) => ('a => bool) => 'a"  where
+  "RES_SELECT == %p m. SOME x. IN x p & m x"
+
+lemma RES_SELECT_DEF: "RES_SELECT = (%p m. SOME x. IN x p & m x)"
+  sorry
+
+lemma EXCLUDED_MIDDLE: "t | ~ t"
+  sorry
+
+lemma FORALL_THM: "All f = All f"
+  sorry
+
+lemma EXISTS_THM: "Ex f = Ex f"
+  sorry
+
+lemma F_IMP: "[| ~ t; t |] ==> False"
+  sorry
+
+lemma NOT_AND: "~ (t & ~ t)"
+  sorry
+
+lemma AND_CLAUSES: "(True & t) = t &
+(t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
+  sorry
+
+lemma OR_CLAUSES: "(True | t) = True &
+(t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
+  sorry
+
+lemma IMP_CLAUSES: "(True --> t) = t &
+(t --> True) = True &
+(False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
+  sorry
+
+lemma NOT_CLAUSES: "(ALL t. (~ ~ t) = t) & (~ True) = False & (~ False) = True"
+  sorry
 
 lemma BOOL_EQ_DISTINCT: "True ~= False & False ~= True"
-  by (import bool BOOL_EQ_DISTINCT)
-
-lemma EQ_CLAUSES: "ALL t::bool.
-   (True = t) = t &
-   (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
-  by (import bool EQ_CLAUSES)
-
-lemma COND_CLAUSES: "ALL (t1::'a::type) t2::'a::type.
-   (if True then t1 else t2) = t1 & (if False then t1 else t2) = t2"
-  by (import bool COND_CLAUSES)
-
-lemma SELECT_UNIQUE: "ALL (P::'a::type => bool) x::'a::type.
-   (ALL y::'a::type. P y = (y = x)) --> Eps P = x"
-  by (import bool SELECT_UNIQUE)
-
-lemma BOTH_EXISTS_AND_THM: "ALL (P::bool) Q::bool.
-   (EX x::'a::type. P & Q) = ((EX x::'a::type. P) & (EX x::'a::type. Q))"
-  by (import bool BOTH_EXISTS_AND_THM)
-
-lemma BOTH_FORALL_OR_THM: "ALL (P::bool) Q::bool.
-   (ALL x::'a::type. P | Q) = ((ALL x::'a::type. P) | (ALL x::'a::type. Q))"
-  by (import bool BOTH_FORALL_OR_THM)
-
-lemma BOTH_FORALL_IMP_THM: "ALL (P::bool) Q::bool.
-   (ALL x::'a::type. P --> Q) =
-   ((EX x::'a::type. P) --> (ALL x::'a::type. Q))"
-  by (import bool BOTH_FORALL_IMP_THM)
-
-lemma BOTH_EXISTS_IMP_THM: "ALL (P::bool) Q::bool.
-   (EX x::'a::type. P --> Q) =
-   ((ALL x::'a::type. P) --> (EX x::'a::type. Q))"
-  by (import bool BOTH_EXISTS_IMP_THM)
-
-lemma OR_IMP_THM: "ALL (A::bool) B::bool. (A = (B | A)) = (B --> A)"
-  by (import bool OR_IMP_THM)
-
-lemma DE_MORGAN_THM: "ALL (A::bool) B::bool. (~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)"
-  by (import bool DE_MORGAN_THM)
-
-lemma IMP_F_EQ_F: "ALL t::bool. (t --> False) = (t = False)"
-  by (import bool IMP_F_EQ_F)
-
-lemma EQ_EXPAND: "ALL (t1::bool) t2::bool. (t1 = t2) = (t1 & t2 | ~ t1 & ~ t2)"
-  by (import bool EQ_EXPAND)
-
-lemma COND_RATOR: "ALL (b::bool) (f::'a::type => 'b::type) (g::'a::type => 'b::type)
-   x::'a::type. (if b then f else g) x = (if b then f x else g x)"
-  by (import bool COND_RATOR)
-
-lemma COND_ABS: "ALL (b::bool) (f::'a::type => 'b::type) g::'a::type => 'b::type.
-   (%x::'a::type. if b then f x else g x) = (if b then f else g)"
-  by (import bool COND_ABS)
-
-lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool.
-   (if b then t1 else t2) = ((~ b | t1) & (b | t2))"
-  by (import bool COND_EXPAND)
-
-lemma ONE_ONE_THM: "ALL f::'a::type => 'b::type.
-   inj f = (ALL (x1::'a::type) x2::'a::type. f x1 = f x2 --> x1 = x2)"
-  by (import bool ONE_ONE_THM)
-
-lemma ABS_REP_THM: "(All::(('a::type => bool) => bool) => bool)
- (%P::'a::type => bool.
-     (op -->::bool => bool => bool)
-      ((Ex::(('b::type => 'a::type) => bool) => bool)
-        ((TYPE_DEFINITION::('a::type => bool)
-                           => ('b::type => 'a::type) => bool)
-          P))
-      ((Ex::(('b::type => 'a::type) => bool) => bool)
-        (%x::'b::type => 'a::type.
-            (Ex::(('a::type => 'b::type) => bool) => bool)
-             (%abs::'a::type => 'b::type.
-                 (op &::bool => bool => bool)
-                  ((All::('b::type => bool) => bool)
-                    (%a::'b::type.
-                        (op =::'b::type => 'b::type => bool) (abs (x a)) a))
-                  ((All::('a::type => bool) => bool)
-                    (%r::'a::type.
-                        (op =::bool => bool => bool) (P r)
-                         ((op =::'a::type => 'a::type => bool) (x (abs r))
-                           r)))))))"
-  by (import bool ABS_REP_THM)
-
-lemma LET_RAND: "(P::'b::type => bool) (Let (M::'a::type) (N::'a::type => 'b::type)) =
-(let x::'a::type = M in P (N x))"
-  by (import bool LET_RAND)
-
-lemma LET_RATOR: "Let (M::'a::type) (N::'a::type => 'b::type => 'c::type) (b::'b::type) =
-(let x::'a::type = M in N x b)"
-  by (import bool LET_RATOR)
-
-lemma SWAP_FORALL_THM: "ALL P::'a::type => 'b::type => bool.
-   (ALL x::'a::type. All (P x)) = (ALL (y::'b::type) x::'a::type. P x y)"
-  by (import bool SWAP_FORALL_THM)
-
-lemma SWAP_EXISTS_THM: "ALL P::'a::type => 'b::type => bool.
-   (EX x::'a::type. Ex (P x)) = (EX (y::'b::type) x::'a::type. P x y)"
-  by (import bool SWAP_EXISTS_THM)
-
-lemma AND_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool.
-   (Q --> P = P') & (P' --> Q = Q') --> (P & Q) = (P' & Q')"
-  by (import bool AND_CONG)
-
-lemma OR_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool.
-   (~ Q --> P = P') & (~ P' --> Q = Q') --> (P | Q) = (P' | Q')"
-  by (import bool OR_CONG)
-
-lemma COND_CONG: "ALL (P::bool) (Q::bool) (x::'a::type) (x'::'a::type) (y::'a::type)
-   y'::'a::type.
-   P = Q & (Q --> x = x') & (~ Q --> y = y') -->
-   (if P then x else y) = (if Q then x' else y')"
-  by (import bool COND_CONG)
-
-lemma MONO_COND: "((x::bool) --> (y::bool)) -->
-((z::bool) --> (w::bool)) -->
-(if b::bool then x else z) --> (if b then y else w)"
-  by (import bool MONO_COND)
-
-lemma SKOLEM_THM: "ALL P::'a::type => 'b::type => bool.
-   (ALL x::'a::type. Ex (P x)) =
-   (EX f::'a::type => 'b::type. ALL x::'a::type. P x (f x))"
-  by (import bool SKOLEM_THM)
-
-lemma bool_case_thm: "(ALL (e0::'a::type) e1::'a::type.
-    (case True of True => e0 | False => e1) = e0) &
-(ALL (e0::'a::type) e1::'a::type.
-    (case False of True => e0 | False => e1) = e1)"
-  by (import bool bool_case_thm)
-
-lemma bool_case_ID: "ALL (x::'a::type) b::bool. (case b of True => x | _ => x) = x"
-  by (import bool bool_case_ID)
-
-lemma boolAxiom: "ALL (e0::'a::type) e1::'a::type.
-   EX x::bool => 'a::type. x True = e0 & x False = e1"
-  by (import bool boolAxiom)
-
-lemma UEXISTS_OR_THM: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   (EX! x::'a::type. P x | Q x) --> Ex1 P | Ex1 Q"
-  by (import bool UEXISTS_OR_THM)
-
-lemma UEXISTS_SIMP: "(EX! x::'a::type. (t::bool)) = (t & (ALL x::'a::type. All (op = x)))"
-  by (import bool UEXISTS_SIMP)
+  sorry
+
+lemma EQ_CLAUSES: "(True = t) = t & (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
+  sorry
+
+lemma COND_CLAUSES: "(if True then t1 else t2) = t1 & (if False then t1 else t2) = t2"
+  sorry
+
+lemma SELECT_UNIQUE: "(!!y. P y = (y = x)) ==> Eps P = x"
+  sorry
+
+lemma BOTH_EXISTS_AND_THM: "(EX x::'a. (P::bool) & (Q::bool)) = ((EX x::'a. P) & (EX x::'a. Q))"
+  sorry
+
+lemma BOTH_FORALL_OR_THM: "(ALL x::'a. (P::bool) | (Q::bool)) = ((ALL x::'a. P) | (ALL x::'a. Q))"
+  sorry
+
+lemma BOTH_FORALL_IMP_THM: "(ALL x::'a. (P::bool) --> (Q::bool)) = ((EX x::'a. P) --> (ALL x::'a. Q))"
+  sorry
+
+lemma BOTH_EXISTS_IMP_THM: "(EX x::'a. (P::bool) --> (Q::bool)) = ((ALL x::'a. P) --> (EX x::'a. Q))"
+  sorry
+
+lemma OR_IMP_THM: "(A = (B | A)) = (B --> A)"
+  sorry
+
+lemma DE_MORGAN_THM: "(~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)"
+  sorry
+
+lemma IMP_F_EQ_F: "(t --> False) = (t = False)"
+  sorry
+
+lemma COND_RATOR: "(if b::bool then f::'a => 'b else (g::'a => 'b)) (x::'a) =
+(if b then f x else g x)"
+  sorry
+
+lemma COND_ABS: "(%x. if b then f x else g x) = (if b then f else g)"
+  sorry
+
+lemma COND_EXPAND: "(if b then t1 else t2) = ((~ b | t1) & (b | t2))"
+  sorry
+
+lemma ONE_ONE_THM: "inj f = (ALL x1 x2. f x1 = f x2 --> x1 = x2)"
+  sorry
+
+lemma ABS_REP_THM: "(op ==>::prop => prop => prop)
+ ((Trueprop::bool => prop)
+   ((Ex::(('b::type => 'a::type) => bool) => bool)
+     ((TYPE_DEFINITION::('a::type => bool)
+                        => ('b::type => 'a::type) => bool)
+       (P::'a::type => bool))))
+ ((Trueprop::bool => prop)
+   ((Ex::(('b::type => 'a::type) => bool) => bool)
+     (%x::'b::type => 'a::type.
+         (Ex::(('a::type => 'b::type) => bool) => bool)
+          (%abs::'a::type => 'b::type.
+              (op &::bool => bool => bool)
+               ((All::('b::type => bool) => bool)
+                 (%a::'b::type.
+                     (op =::'b::type => 'b::type => bool) (abs (x a)) a))
+               ((All::('a::type => bool) => bool)
+                 (%r::'a::type.
+                     (op =::bool => bool => bool) (P r)
+                      ((op =::'a::type => 'a::type => bool) (x (abs r))
+                        r)))))))"
+  sorry
+
+lemma LET_RAND: "(P::'b => bool) (Let (M::'a) (N::'a => 'b)) = (let x::'a = M in P (N x))"
+  sorry
+
+lemma LET_RATOR: "Let (M::'a) (N::'a => 'b => 'c) (b::'b) = (let x::'a = M in N x b)"
+  sorry
+
+lemma AND_CONG: "(Q --> P = P') & (P' --> Q = Q') ==> (P & Q) = (P' & Q')"
+  sorry
+
+lemma OR_CONG: "(~ Q --> P = P') & (~ P' --> Q = Q') ==> (P | Q) = (P' | Q')"
+  sorry
+
+lemma COND_CONG: "P = Q & (Q --> x = x') & (~ Q --> y = y')
+==> (if P then x else y) = (if Q then x' else y')"
+  sorry
+
+lemma MONO_COND: "[| x ==> y; z ==> w; if b then x else z |] ==> if b then y else w"
+  sorry
+
+lemma SKOLEM_THM: "(ALL x. Ex (P x)) = (EX f. ALL x. P x (f x))"
+  sorry
+
+lemma bool_case_thm: "(ALL (e0::'a) e1::'a. (case True of True => e0 | False => e1) = e0) &
+(ALL (e0::'a) e1::'a. (case False of True => e0 | False => e1) = e1)"
+  sorry
+
+lemma bool_case_ID: "(case b of True => x | _ => x) = x"
+  sorry
+
+lemma boolAxiom: "EX x. x True = e0 & x False = e1"
+  sorry
+
+lemma UEXISTS_OR_THM: "EX! x. P x | Q x ==> Ex1 P | Ex1 Q"
+  sorry
+
+lemma UEXISTS_SIMP: "(EX! x::'a. (t::bool)) = (t & (ALL x::'a. All (op = x)))"
+  sorry
 
 consts
   RES_ABSTRACT :: "('a => bool) => ('a => 'b) => 'a => 'b" 
 
-specification (RES_ABSTRACT) RES_ABSTRACT_DEF: "(ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type.
+specification (RES_ABSTRACT) RES_ABSTRACT_DEF: "(ALL (p::'a => bool) (m::'a => 'b) x::'a.
     IN x p --> RES_ABSTRACT p m x = m x) &
-(ALL (p::'a::type => bool) (m1::'a::type => 'b::type)
-    m2::'a::type => 'b::type.
-    (ALL x::'a::type. IN x p --> m1 x = m2 x) -->
+(ALL (p::'a => bool) (m1::'a => 'b) m2::'a => 'b.
+    (ALL x::'a. IN x p --> m1 x = m2 x) -->
     RES_ABSTRACT p m1 = RES_ABSTRACT p m2)"
-  by (import bool RES_ABSTRACT_DEF)
-
-lemma BOOL_FUN_CASES_THM: "ALL f::bool => bool.
-   f = (%b::bool. True) |
-   f = (%b::bool. False) | f = (%b::bool. b) | f = Not"
-  by (import bool BOOL_FUN_CASES_THM)
-
-lemma BOOL_FUN_INDUCT: "ALL P::(bool => bool) => bool.
-   P (%b::bool. True) & P (%b::bool. False) & P (%b::bool. b) & P Not -->
-   All P"
-  by (import bool BOOL_FUN_INDUCT)
+  sorry
+
+lemma BOOL_FUN_CASES_THM: "f = (%b. True) | f = (%b. False) | f = (%b. b) | f = Not"
+  sorry
+
+lemma BOOL_FUN_INDUCT: "P (%b. True) & P (%b. False) & P (%b. b) & P Not ==> P x"
+  sorry
 
 ;end_setup
 
 ;setup_theory combin
 
-definition K :: "'a => 'b => 'a" where 
-  "K == %(x::'a::type) y::'b::type. x"
-
-lemma K_DEF: "K = (%(x::'a::type) y::'b::type. x)"
-  by (import combin K_DEF)
-
-definition S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" where 
-  "S ==
-%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type)
-   x::'a::type. f x (g x)"
-
-lemma S_DEF: "S =
-(%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type)
-    x::'a::type. f x (g x))"
-  by (import combin S_DEF)
-
-definition I :: "'a => 'a" where 
+definition
+  K :: "'a => 'b => 'a"  where
+  "K == %x y. x"
+
+lemma K_DEF: "K = (%x y. x)"
+  sorry
+
+definition
+  S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c"  where
+  "S == %f g x. f x (g x)"
+
+lemma S_DEF: "S = (%f g x. f x (g x))"
+  sorry
+
+definition
+  I :: "'a => 'a"  where
   "(op ==::('a::type => 'a::type) => ('a::type => 'a::type) => prop)
  (I::'a::type => 'a::type)
  ((S::('a::type => ('a::type => 'a::type) => 'a::type)
@@ -288,47 +236,46 @@
       => ('a::type => 'a::type => 'a::type) => 'a::type => 'a::type)
    (K::'a::type => ('a::type => 'a::type) => 'a::type)
    (K::'a::type => 'a::type => 'a::type))"
-  by (import combin I_DEF)
-
-definition C :: "('a => 'b => 'c) => 'b => 'a => 'c" where 
-  "C == %(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x"
-
-lemma C_DEF: "C =
-(%(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x)"
-  by (import combin C_DEF)
-
-definition W :: "('a => 'a => 'b) => 'a => 'b" where 
-  "W == %(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x"
-
-lemma W_DEF: "W = (%(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x)"
-  by (import combin W_DEF)
-
-lemma I_THM: "ALL x::'a::type. I x = x"
-  by (import combin I_THM)
-
-lemma I_o_ID: "ALL f::'a::type => 'b::type. I o f = f & f o I = f"
-  by (import combin I_o_ID)
+  sorry
+
+definition
+  C :: "('a => 'b => 'c) => 'b => 'a => 'c"  where
+  "C == %f x y. f y x"
+
+lemma C_DEF: "C = (%f x y. f y x)"
+  sorry
+
+definition
+  W :: "('a => 'a => 'b) => 'a => 'b"  where
+  "W == %f x. f x x"
+
+lemma W_DEF: "W = (%f x. f x x)"
+  sorry
+
+lemma I_THM: "I x = x"
+  sorry
+
+lemma I_o_ID: "I o f = f & f o I = f"
+  sorry
 
 ;end_setup
 
 ;setup_theory sum
 
-lemma ISL_OR_ISR: "ALL x::'a::type + 'b::type. ISL x | ISR x"
-  by (import sum ISL_OR_ISR)
-
-lemma INL: "ALL x::'a::type + 'b::type. ISL x --> Inl (OUTL x) = x"
-  by (import sum INL)
-
-lemma INR: "ALL x::'a::type + 'b::type. ISR x --> Inr (OUTR x) = x"
-  by (import sum INR)
-
-lemma sum_case_cong: "ALL (M::'b::type + 'c::type) (M'::'b::type + 'c::type)
-   (f::'b::type => 'a::type) g::'c::type => 'a::type.
-   M = M' &
-   (ALL x::'b::type. M' = Inl x --> f x = (f'::'b::type => 'a::type) x) &
-   (ALL y::'c::type. M' = Inr y --> g y = (g'::'c::type => 'a::type) y) -->
-   sum_case f g M = sum_case f' g' M'"
-  by (import sum sum_case_cong)
+lemma ISL_OR_ISR: "ISL x | ISR x"
+  sorry
+
+lemma INL: "ISL x ==> Inl (OUTL x) = x"
+  sorry
+
+lemma INR: "ISR x ==> Inr (OUTR x) = x"
+  sorry
+
+lemma sum_case_cong: "(M::'b + 'c) = (M'::'b + 'c) &
+(ALL x::'b. M' = Inl x --> (f::'b => 'a) x = (f'::'b => 'a) x) &
+(ALL y::'c. M' = Inr y --> (g::'c => 'a) y = (g'::'c => 'a) y)
+==> sum_case f g M = sum_case f' g' M'"
+  sorry
 
 ;end_setup
 
@@ -345,34 +292,34 @@
         (%y::'a::type.
             (op =::bool => bool => bool)
              ((op =::'a::type option => 'a::type option => bool)
-               ((Some::'a::type ~=> 'a::type) x)
-               ((Some::'a::type ~=> 'a::type) y))
+               ((Some::'a::type => 'a::type option) x)
+               ((Some::'a::type => 'a::type option) y))
              ((op =::'a::type => 'a::type => bool) x y))))
  ((op &::bool => bool => bool)
    ((All::('a::type => bool) => bool)
      (%x::'a::type.
          (op =::'a::type => 'a::type => bool)
           ((the::'a::type option => 'a::type)
-            ((Some::'a::type ~=> 'a::type) x))
+            ((Some::'a::type => 'a::type option) x))
           x))
    ((op &::bool => bool => bool)
      ((All::('a::type => bool) => bool)
        (%x::'a::type.
-           (Not::bool => bool)
-            ((op =::'a::type option => 'a::type option => bool)
-              (None::'a::type option) ((Some::'a::type ~=> 'a::type) x))))
+           (op ~=::'a::type option => 'a::type option => bool)
+            (None::'a::type option)
+            ((Some::'a::type => 'a::type option) x)))
      ((op &::bool => bool => bool)
        ((All::('a::type => bool) => bool)
          (%x::'a::type.
-             (Not::bool => bool)
-              ((op =::'a::type option => 'a::type option => bool)
-                ((Some::'a::type ~=> 'a::type) x) (None::'a::type option))))
+             (op ~=::'a::type option => 'a::type option => bool)
+              ((Some::'a::type => 'a::type option) x)
+              (None::'a::type option)))
        ((op &::bool => bool => bool)
          ((All::('a::type => bool) => bool)
            (%x::'a::type.
                (op =::bool => bool => bool)
                 ((IS_SOME::'a::type option => bool)
-                  ((Some::'a::type ~=> 'a::type) x))
+                  ((Some::'a::type => 'a::type option) x))
                 (True::bool)))
          ((op &::bool => bool => bool)
            ((op =::bool => bool => bool)
@@ -399,7 +346,7 @@
                        (op -->::bool => bool => bool)
                         ((IS_SOME::'a::type option => bool) x)
                         ((op =::'a::type option => 'a::type option => bool)
-                          ((Some::'a::type ~=> 'a::type)
+                          ((Some::'a::type => 'a::type option)
                             ((the::'a::type option => 'a::type) x))
                           x)))
                  ((op &::bool => bool => bool)
@@ -407,9 +354,9 @@
                      (%x::'a::type option.
                          (op =::'a::type option => 'a::type option => bool)
                           ((option_case::'a::type option
-   => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
+   => ('a::type => 'a::type option) => 'a::type option => 'a::type option)
                             (None::'a::type option)
-                            (Some::'a::type ~=> 'a::type) x)
+                            (Some::'a::type => 'a::type option) x)
                           x))
                    ((op &::bool => bool => bool)
                      ((All::('a::type option => bool) => bool)
@@ -417,8 +364,8 @@
                            (op =::'a::type option
                                   => 'a::type option => bool)
                             ((option_case::'a::type option
-     => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
-                              x (Some::'a::type ~=> 'a::type) x)
+     => ('a::type => 'a::type option) => 'a::type option => 'a::type option)
+                              x (Some::'a::type => 'a::type option) x)
                             x))
                      ((op &::bool => bool => bool)
                        ((All::('a::type option => bool) => bool)
@@ -449,8 +396,9 @@
                                   ((op =::'a::type option
     => 'a::type option => bool)
                                     ((option_case::'a::type option
-             => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
-(ea::'a::type option) (Some::'a::type ~=> 'a::type) x)
+             => ('a::type => 'a::type option)
+                => 'a::type option => 'a::type option)
+(ea::'a::type option) (Some::'a::type => 'a::type option) x)
                                     x)))
                            ((op &::bool => bool => bool)
                              ((All::('b::type => bool) => bool)
@@ -475,7 +423,7 @@
           ((option_case::'b::type
                          => ('a::type => 'b::type)
                             => 'a::type option => 'b::type)
-            u f ((Some::'a::type ~=> 'a::type) x))
+            u f ((Some::'a::type => 'a::type option) x))
           (f x)))))
                                ((op &::bool => bool => bool)
                                  ((All::(('a::type => 'b::type) => bool)
@@ -484,51 +432,48 @@
  (All::('a::type => bool) => bool)
   (%x::'a::type.
       (op =::'b::type option => 'b::type option => bool)
-       ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type)
-         f ((Some::'a::type ~=> 'a::type) x))
-       ((Some::'b::type ~=> 'b::type) (f x)))))
+       ((Option.map::('a::type => 'b::type)
+                     => 'a::type option => 'b::type option)
+         f ((Some::'a::type => 'a::type option) x))
+       ((Some::'b::type => 'b::type option) (f x)))))
                                  ((op &::bool => bool => bool)
                                    ((All::(('a::type => 'b::type) => bool)
     => bool)
                                      (%f::'a::type => 'b::type.
    (op =::'b::type option => 'b::type option => bool)
-    ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type) f
-      (None::'a::type option))
+    ((Option.map::('a::type => 'b::type)
+                  => 'a::type option => 'b::type option)
+      f (None::'a::type option))
     (None::'b::type option)))
                                    ((op &::bool => bool => bool)
                                      ((op =::'a::type option
        => 'a::type option => bool)
- ((OPTION_JOIN::'a::type option option ~=> 'a::type)
+ ((OPTION_JOIN::'a::type option option => 'a::type option)
    (None::'a::type option option))
  (None::'a::type option))
                                      ((All::('a::type option => bool)
       => bool)
  (%x::'a::type option.
      (op =::'a::type option => 'a::type option => bool)
-      ((OPTION_JOIN::'a::type option option ~=> 'a::type)
-        ((Some::'a::type option ~=> 'a::type option) x))
+      ((OPTION_JOIN::'a::type option option => 'a::type option)
+        ((Some::'a::type option => 'a::type option option) x))
       x))))))))))))))))))))"
-  by (import option option_CLAUSES)
-
-lemma option_case_compute: "option_case (e::'b::type) (f::'a::type => 'b::type) (x::'a::type option) =
+  sorry
+
+lemma option_case_compute: "option_case (e::'b) (f::'a => 'b) (x::'a option) =
 (if IS_SOME x then f (the x) else e)"
-  by (import option option_case_compute)
-
-lemma OPTION_MAP_EQ_SOME: "ALL (f::'a::type => 'b::type) (x::'a::type option) y::'b::type.
-   (option_map f x = Some y) = (EX z::'a::type. x = Some z & y = f z)"
-  by (import option OPTION_MAP_EQ_SOME)
-
-lemma OPTION_JOIN_EQ_SOME: "ALL (x::'a::type option option) xa::'a::type.
-   (OPTION_JOIN x = Some xa) = (x = Some (Some xa))"
-  by (import option OPTION_JOIN_EQ_SOME)
-
-lemma option_case_cong: "ALL (M::'a::type option) (M'::'a::type option) (u::'b::type)
-   f::'a::type => 'b::type.
-   M = M' &
-   (M' = None --> u = (u'::'b::type)) &
-   (ALL x::'a::type. M' = Some x --> f x = (f'::'a::type => 'b::type) x) -->
-   option_case u f M = option_case u' f' M'"
-  by (import option option_case_cong)
+  sorry
+
+lemma OPTION_MAP_EQ_SOME: "(Option.map (f::'a => 'b) (x::'a option) = Some (y::'b)) =
+(EX z::'a. x = Some z & y = f z)"
+  sorry
+
+lemma OPTION_JOIN_EQ_SOME: "(OPTION_JOIN x = Some xa) = (x = Some (Some xa))"
+  sorry
+
+lemma option_case_cong: "M = M' & (M' = None --> u = u') & (ALL x. M' = Some x --> f x = f' x)
+==> option_case u f M = option_case u' f' M'"
+  sorry
 
 ;end_setup
 
@@ -538,531 +483,341 @@
   stmarker :: "'a => 'a" 
 
 defs
-  stmarker_primdef: "stmarker == %x::'a::type. x"
-
-lemma stmarker_def: "ALL x::'a::type. stmarker x = x"
-  by (import marker stmarker_def)
-
-lemma move_left_conj: "ALL (x::bool) (xa::bool) xb::bool.
-   (x & stmarker xb) = (stmarker xb & x) &
-   ((stmarker xb & x) & xa) = (stmarker xb & x & xa) &
-   (x & stmarker xb & xa) = (stmarker xb & x & xa)"
-  by (import marker move_left_conj)
-
-lemma move_right_conj: "ALL (x::bool) (xa::bool) xb::bool.
-   (stmarker xb & x) = (x & stmarker xb) &
-   (x & xa & stmarker xb) = ((x & xa) & stmarker xb) &
-   ((x & stmarker xb) & xa) = ((x & xa) & stmarker xb)"
-  by (import marker move_right_conj)
-
-lemma move_left_disj: "ALL (x::bool) (xa::bool) xb::bool.
-   (x | stmarker xb) = (stmarker xb | x) &
-   ((stmarker xb | x) | xa) = (stmarker xb | x | xa) &
-   (x | stmarker xb | xa) = (stmarker xb | x | xa)"
-  by (import marker move_left_disj)
-
-lemma move_right_disj: "ALL (x::bool) (xa::bool) xb::bool.
-   (stmarker xb | x) = (x | stmarker xb) &
-   (x | xa | stmarker xb) = ((x | xa) | stmarker xb) &
-   ((x | stmarker xb) | xa) = ((x | xa) | stmarker xb)"
-  by (import marker move_right_disj)
+  stmarker_primdef: "stmarker == %x. x"
+
+lemma stmarker_def: "stmarker x = x"
+  sorry
+
+lemma move_left_conj: "(x & stmarker xb) = (stmarker xb & x) &
+((stmarker xb & x) & xa) = (stmarker xb & x & xa) &
+(x & stmarker xb & xa) = (stmarker xb & x & xa)"
+  sorry
+
+lemma move_right_conj: "(stmarker xb & x) = (x & stmarker xb) &
+(x & xa & stmarker xb) = ((x & xa) & stmarker xb) &
+((x & stmarker xb) & xa) = ((x & xa) & stmarker xb)"
+  sorry
+
+lemma move_left_disj: "(x | stmarker xb) = (stmarker xb | x) &
+((stmarker xb | x) | xa) = (stmarker xb | x | xa) &
+(x | stmarker xb | xa) = (stmarker xb | x | xa)"
+  sorry
+
+lemma move_right_disj: "(stmarker xb | x) = (x | stmarker xb) &
+(x | xa | stmarker xb) = ((x | xa) | stmarker xb) &
+((x | stmarker xb) | xa) = ((x | xa) | stmarker xb)"
+  sorry
 
 ;end_setup
 
 ;setup_theory relation
 
-definition TC :: "('a => 'a => bool) => 'a => 'a => bool" where 
+definition
+  TC :: "('a => 'a => bool) => 'a => 'a => bool"  where
   "TC ==
-%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
-   ALL P::'a::type => 'a::type => bool.
-      (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
-      (ALL (x::'a::type) (y::'a::type) z::'a::type.
-          P x y & P y z --> P x z) -->
+%R a b.
+   ALL P.
+      (ALL x y. R x y --> P x y) & (ALL x y z. P x y & P y z --> P x z) -->
       P a b"
 
-lemma TC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
-   TC R a b =
-   (ALL P::'a::type => 'a::type => bool.
-       (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
-       (ALL (x::'a::type) (y::'a::type) z::'a::type.
-           P x y & P y z --> P x z) -->
-       P a b)"
-  by (import relation TC_DEF)
-
-definition RTC :: "('a => 'a => bool) => 'a => 'a => bool" where 
+lemma TC_DEF: "TC R a b =
+(ALL P.
+    (ALL x y. R x y --> P x y) & (ALL x y z. P x y & P y z --> P x z) -->
+    P a b)"
+  sorry
+
+definition
+  RTC :: "('a => 'a => bool) => 'a => 'a => bool"  where
   "RTC ==
-%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
-   ALL P::'a::type => 'a::type => bool.
-      (ALL x::'a::type. P x x) &
-      (ALL (x::'a::type) (y::'a::type) z::'a::type.
-          R x y & P y z --> P x z) -->
-      P a b"
-
-lemma RTC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
-   RTC R a b =
-   (ALL P::'a::type => 'a::type => bool.
-       (ALL x::'a::type. P x x) &
-       (ALL (x::'a::type) (y::'a::type) z::'a::type.
-           R x y & P y z --> P x z) -->
-       P a b)"
-  by (import relation RTC_DEF)
+%R a b.
+   ALL P. (ALL x. P x x) & (ALL x y z. R x y & P y z --> P x z) --> P a b"
+
+lemma RTC_DEF: "RTC R a b =
+(ALL P. (ALL x. P x x) & (ALL x y z. R x y & P y z --> P x z) --> P a b)"
+  sorry
 
 consts
   RC :: "('a => 'a => bool) => 'a => 'a => bool" 
 
 defs
-  RC_primdef: "RC ==
-%(R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. x = y | R x y"
-
-lemma RC_def: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   RC R x y = (x = y | R x y)"
-  by (import relation RC_def)
+  RC_primdef: "RC == %R x y. x = y | R x y"
+
+lemma RC_def: "RC R x y = (x = y | R x y)"
+  sorry
 
 consts
   transitive :: "('a => 'a => bool) => bool" 
 
 defs
-  transitive_primdef: "transitive ==
-%R::'a::type => 'a::type => bool.
-   ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z"
-
-lemma transitive_def: "ALL R::'a::type => 'a::type => bool.
-   transitive R =
-   (ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z)"
-  by (import relation transitive_def)
-
-definition pred_reflexive :: "('a => 'a => bool) => bool" where 
-  "pred_reflexive == %R::'a::type => 'a::type => bool. ALL x::'a::type. R x x"
-
-lemma reflexive_def: "ALL R::'a::type => 'a::type => bool.
-   pred_reflexive R = (ALL x::'a::type. R x x)"
-  by (import relation reflexive_def)
-
-lemma TC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (TC x)"
-  by (import relation TC_TRANSITIVE)
-
-lemma RTC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
-   (ALL x::'a::type. xa x x) &
-   (ALL (xb::'a::type) (y::'a::type) z::'a::type.
-       x xb y & xa y z --> xa xb z) -->
-   (ALL (xb::'a::type) xc::'a::type. RTC x xb xc --> xa xb xc)"
-  by (import relation RTC_INDUCT)
-
-lemma TC_RULES: "ALL x::'a::type => 'a::type => bool.
-   (ALL (xa::'a::type) xb::'a::type. x xa xb --> TC x xa xb) &
-   (ALL (xa::'a::type) (xb::'a::type) xc::'a::type.
-       TC x xa xb & TC x xb xc --> TC x xa xc)"
-  by (import relation TC_RULES)
-
-lemma RTC_RULES: "ALL x::'a::type => 'a::type => bool.
-   (ALL xa::'a::type. RTC x xa xa) &
-   (ALL (xa::'a::type) (xb::'a::type) xc::'a::type.
-       x xa xb & RTC x xb xc --> RTC x xa xc)"
-  by (import relation RTC_RULES)
-
-lemma RTC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
-   (ALL x::'a::type. P x x) &
-   (ALL (x::'a::type) (y::'a::type) z::'a::type.
-       R x y & RTC R y z & P y z --> P x z) -->
-   (ALL (x::'a::type) y::'a::type. RTC R x y --> P x y)"
-  by (import relation RTC_STRONG_INDUCT)
-
-lemma RTC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   RTC R x y --> (ALL z::'a::type. RTC R y z --> RTC R x z)"
-  by (import relation RTC_RTC)
-
-lemma RTC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (RTC x)"
-  by (import relation RTC_TRANSITIVE)
-
-lemma RTC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RTC R)"
-  by (import relation RTC_REFLEXIVE)
-
-lemma RC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RC R)"
-  by (import relation RC_REFLEXIVE)
-
-lemma TC_SUBSET: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
-   x xa xb --> TC x xa xb"
-  by (import relation TC_SUBSET)
-
-lemma RTC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   R x y --> RTC R x y"
-  by (import relation RTC_SUBSET)
-
-lemma RC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   R x y --> RC R x y"
-  by (import relation RC_SUBSET)
-
-lemma RC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   RC R x y --> RTC R x y"
-  by (import relation RC_RTC)
-
-lemma TC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
-   (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) &
-   (ALL (x::'a::type) (y::'a::type) z::'a::type.
-       xa x y & xa y z --> xa x z) -->
-   (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)"
-  by (import relation TC_INDUCT)
-
-lemma TC_INDUCT_LEFT1: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
-   (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) &
-   (ALL (xb::'a::type) (y::'a::type) z::'a::type.
-       x xb y & xa y z --> xa xb z) -->
-   (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)"
-  by (import relation TC_INDUCT_LEFT1)
-
-lemma TC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
-   (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
-   (ALL (x::'a::type) (y::'a::type) z::'a::type.
-       P x y & P y z & TC R x y & TC R y z --> P x z) -->
-   (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)"
-  by (import relation TC_STRONG_INDUCT)
-
-lemma TC_STRONG_INDUCT_LEFT1: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
-   (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
-   (ALL (x::'a::type) (y::'a::type) z::'a::type.
-       R x y & P y z & TC R y z --> P x z) -->
-   (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)"
-  by (import relation TC_STRONG_INDUCT_LEFT1)
-
-lemma TC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   TC R x y --> RTC R x y"
-  by (import relation TC_RTC)
-
-lemma RTC_TC_RC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
-   RTC R x y --> RC R x y | TC R x y"
-  by (import relation RTC_TC_RC)
-
-lemma TC_RC_EQNS: "ALL R::'a::type => 'a::type => bool. RC (TC R) = RTC R & TC (RC R) = RTC R"
-  by (import relation TC_RC_EQNS)
-
-lemma RC_IDEM: "ALL R::'a::type => 'a::type => bool. RC (RC R) = RC R"
-  by (import relation RC_IDEM)
-
-lemma TC_IDEM: "ALL R::'a::type => 'a::type => bool. TC (TC R) = TC R"
-  by (import relation TC_IDEM)
-
-lemma RTC_IDEM: "ALL R::'a::type => 'a::type => bool. RTC (RTC R) = RTC R"
-  by (import relation RTC_IDEM)
-
-lemma RTC_CASES1: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
-   RTC x xa xb = (xa = xb | (EX u::'a::type. x xa u & RTC x u xb))"
-  by (import relation RTC_CASES1)
-
-lemma RTC_CASES2: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
-   RTC x xa xb = (xa = xb | (EX u::'a::type. RTC x xa u & x u xb))"
-  by (import relation RTC_CASES2)
-
-lemma RTC_CASES_RTC_TWICE: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
-   RTC x xa xb = (EX u::'a::type. RTC x xa u & RTC x u xb)"
-  by (import relation RTC_CASES_RTC_TWICE)
-
-lemma TC_CASES1: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type.
-   TC R x z --> R x z | (EX y::'a::type. R x y & TC R y z)"
-  by (import relation TC_CASES1)
-
-lemma TC_CASES2: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type.
-   TC R x z --> R x z | (EX y::'a::type. TC R x y & R y z)"
-  by (import relation TC_CASES2)
-
-lemma TC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool.
-   (ALL (x::'a::type) y::'a::type. R x y --> Q x y) -->
-   (ALL (x::'a::type) y::'a::type. TC R x y --> TC Q x y)"
-  by (import relation TC_MONOTONE)
-
-lemma RTC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool.
-   (ALL (x::'a::type) y::'a::type. R x y --> Q x y) -->
-   (ALL (x::'a::type) y::'a::type. RTC R x y --> RTC Q x y)"
-  by (import relation RTC_MONOTONE)
-
-definition WF :: "('a => 'a => bool) => bool" where 
-  "WF ==
-%R::'a::type => 'a::type => bool.
-   ALL B::'a::type => bool.
-      Ex B -->
-      (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b))"
-
-lemma WF_DEF: "ALL R::'a::type => 'a::type => bool.
-   WF R =
-   (ALL B::'a::type => bool.
-       Ex B -->
-       (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b)))"
-  by (import relation WF_DEF)
-
-lemma WF_INDUCTION_THM: "ALL R::'a::type => 'a::type => bool.
-   WF R -->
-   (ALL P::'a::type => bool.
-       (ALL x::'a::type. (ALL y::'a::type. R y x --> P y) --> P x) -->
-       All P)"
-  by (import relation WF_INDUCTION_THM)
-
-lemma WF_NOT_REFL: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
-   WF x --> x xa xb --> xa ~= xb"
-  by (import relation WF_NOT_REFL)
-
-definition EMPTY_REL :: "'a => 'a => bool" where 
-  "EMPTY_REL == %(x::'a::type) y::'a::type. False"
-
-lemma EMPTY_REL_DEF: "ALL (x::'a::type) y::'a::type. EMPTY_REL x y = False"
-  by (import relation EMPTY_REL_DEF)
+  transitive_primdef: "transitive == %R. ALL x y z. R x y & R y z --> R x z"
+
+lemma transitive_def: "transitive R = (ALL x y z. R x y & R y z --> R x z)"
+  sorry
+
+definition
+  pred_reflexive :: "('a => 'a => bool) => bool"  where
+  "pred_reflexive == %R. ALL x. R x x"
+
+lemma reflexive_def: "pred_reflexive R = (ALL x. R x x)"
+  sorry
+
+lemma TC_TRANSITIVE: "transitive (TC x)"
+  sorry
+
+lemma RTC_INDUCT: "[| (ALL x. xa x x) & (ALL xb y z. x xb y & xa y z --> xa xb z);
+   RTC x xb xc |]
+==> xa xb xc"
+  sorry
+
+lemma TC_RULES: "(ALL xa xb. x xa xb --> TC x xa xb) &
+(ALL xa xb xc. TC x xa xb & TC x xb xc --> TC x xa xc)"
+  sorry
+
+lemma RTC_RULES: "(ALL xa. RTC x xa xa) &
+(ALL xa xb xc. x xa xb & RTC x xb xc --> RTC x xa xc)"
+  sorry
+
+lemma RTC_STRONG_INDUCT: "[| (ALL x. P x x) & (ALL x y z. R x y & RTC R y z & P y z --> P x z);
+   RTC R x y |]
+==> P x y"
+  sorry
+
+lemma RTC_RTC: "[| RTC R x y; RTC R y z |] ==> RTC R x z"
+  sorry
+
+lemma RTC_TRANSITIVE: "transitive (RTC x)"
+  sorry
+
+lemma RTC_REFLEXIVE: "pred_reflexive (RTC R)"
+  sorry
+
+lemma RC_REFLEXIVE: "pred_reflexive (RC R)"
+  sorry
+
+lemma TC_SUBSET: "x xa xb ==> TC x xa xb"
+  sorry
+
+lemma RTC_SUBSET: "R x y ==> RTC R x y"
+  sorry
+
+lemma RC_SUBSET: "R x y ==> RC R x y"
+  sorry
+
+lemma RC_RTC: "RC R x y ==> RTC R x y"
+  sorry
+
+lemma TC_INDUCT: "[| (ALL xb y. x xb y --> xa xb y) & (ALL x y z. xa x y & xa y z --> xa x z);
+   TC x xb xc |]
+==> xa xb xc"
+  sorry
+
+lemma TC_INDUCT_LEFT1: "[| (ALL xb y. x xb y --> xa xb y) &
+   (ALL xb y z. x xb y & xa y z --> xa xb z);
+   TC x xb xc |]
+==> xa xb xc"
+  sorry
+
+lemma TC_STRONG_INDUCT: "[| (ALL x y. R x y --> P x y) &
+   (ALL x y z. P x y & P y z & TC R x y & TC R y z --> P x z);
+   TC R u v |]
+==> P u v"
+  sorry
+
+lemma TC_STRONG_INDUCT_LEFT1: "[| (ALL x y. R x y --> P x y) &
+   (ALL x y z. R x y & P y z & TC R y z --> P x z);
+   TC R u v |]
+==> P u v"
+  sorry
+
+lemma TC_RTC: "TC R x y ==> RTC R x y"
+  sorry
+
+lemma RTC_TC_RC: "RTC R x y ==> RC R x y | TC R x y"
+  sorry
+
+lemma TC_RC_EQNS: "RC (TC R) = RTC R & TC (RC R) = RTC R"
+  sorry
+
+lemma RC_IDEM: "RC (RC R) = RC R"
+  sorry
+
+lemma TC_IDEM: "TC (TC R) = TC R"
+  sorry
+
+lemma RTC_IDEM: "RTC (RTC R) = RTC R"
+  sorry
+
+lemma RTC_CASES1: "RTC x xa xb = (xa = xb | (EX u. x xa u & RTC x u xb))"
+  sorry
+
+lemma RTC_CASES2: "RTC x xa xb = (xa = xb | (EX u. RTC x xa u & x u xb))"
+  sorry
+
+lemma RTC_CASES_RTC_TWICE: "RTC x xa xb = (EX u. RTC x xa u & RTC x u xb)"
+  sorry
+
+lemma TC_CASES1: "TC R x z ==> R x z | (EX y. R x y & TC R y z)"
+  sorry
+
+lemma TC_CASES2: "TC R x z ==> R x z | (EX y. TC R x y & R y z)"
+  sorry
+
+lemma TC_MONOTONE: "[| !!x y. R x y ==> Q x y; TC R x y |] ==> TC Q x y"
+  sorry
+
+lemma RTC_MONOTONE: "[| !!x y. R x y ==> Q x y; RTC R x y |] ==> RTC Q x y"
+  sorry
+
+definition
+  WF :: "('a => 'a => bool) => bool"  where
+  "WF == %R. ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b))"
+
+lemma WF_DEF: "WF R = (ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b)))"
+  sorry
+
+lemma WF_INDUCTION_THM: "[| WF R; !!x. (!!y. R y x ==> P y) ==> P x |] ==> P x"
+  sorry
+
+lemma WF_NOT_REFL: "[| WF x; x xa xb |] ==> xa ~= xb"
+  sorry
+
+definition
+  EMPTY_REL :: "'a => 'a => bool"  where
+  "EMPTY_REL == %x y. False"
+
+lemma EMPTY_REL_DEF: "EMPTY_REL x y = False"
+  sorry
 
 lemma WF_EMPTY_REL: "WF EMPTY_REL"
-  by (import relation WF_EMPTY_REL)
-
-lemma WF_SUBSET: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
-   WF x & (ALL (xb::'a::type) y::'a::type. xa xb y --> x xb y) --> WF xa"
-  by (import relation WF_SUBSET)
-
-lemma WF_TC: "ALL R::'a::type => 'a::type => bool. WF R --> WF (TC R)"
-  by (import relation WF_TC)
+  sorry
+
+lemma WF_SUBSET: "WF x & (ALL xb y. xa xb y --> x xb y) ==> WF xa"
+  sorry
+
+lemma WF_TC: "WF R ==> WF (TC R)"
+  sorry
 
 consts
   inv_image :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" 
 
 defs
   inv_image_primdef: "relation.inv_image ==
-%(R::'b::type => 'b::type => bool) (f::'a::type => 'b::type) (x::'a::type)
-   y::'a::type. R (f x) (f y)"
-
-lemma inv_image_def: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type.
-   relation.inv_image R f = (%(x::'a::type) y::'a::type. R (f x) (f y))"
-  by (import relation inv_image_def)
-
-lemma WF_inv_image: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type.
-   WF R --> WF (relation.inv_image R f)"
-  by (import relation WF_inv_image)
-
-definition RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" where 
-  "RESTRICT ==
-%(f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) (x::'a::type)
-   y::'a::type. if R y x then f y else ARB"
-
-lemma RESTRICT_DEF: "ALL (f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) x::'a::type.
-   RESTRICT f R x = (%y::'a::type. if R y x then f y else ARB)"
-  by (import relation RESTRICT_DEF)
-
-lemma RESTRICT_LEMMA: "ALL (x::'a::type => 'b::type) (xa::'a::type => 'a::type => bool)
-   (xb::'a::type) xc::'a::type. xa xb xc --> RESTRICT x xa xc xb = x xb"
-  by (import relation RESTRICT_LEMMA)
+%(R::'b => 'b => bool) (f::'a => 'b) (x::'a) y::'a. R (f x) (f y)"
+
+lemma inv_image_def: "relation.inv_image R f = (%x y. R (f x) (f y))"
+  sorry
+
+lemma WF_inv_image: "WF (R::'b => 'b => bool) ==> WF (relation.inv_image R (f::'a => 'b))"
+  sorry
+
+definition
+  RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b"  where
+  "RESTRICT == %f R x y. if R y x then f y else ARB"
+
+lemma RESTRICT_DEF: "RESTRICT f R x = (%y. if R y x then f y else ARB)"
+  sorry
+
+lemma RESTRICT_LEMMA: "xa xb xc ==> RESTRICT x xa xc xb = x xb"
+  sorry
 
 consts
   approx :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => ('a => 'b) => bool" 
 
 defs
-  approx_primdef: "approx ==
-%(R::'a::type => 'a::type => bool)
-   (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type)
-   f::'a::type => 'b::type.
-   f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x"
-
-lemma approx_def: "ALL (R::'a::type => 'a::type => bool)
-   (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type)
-   f::'a::type => 'b::type.
-   approx R M x f = (f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x)"
-  by (import relation approx_def)
+  approx_primdef: "approx == %R M x f. f = RESTRICT (%y. M (RESTRICT f R y) y) R x"
+
+lemma approx_def: "approx R M x f = (f = RESTRICT (%y. M (RESTRICT f R y) y) R x)"
+  sorry
 
 consts
   the_fun :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'a => 'b" 
 
 defs
-  the_fun_primdef: "the_fun ==
-%(R::'a::type => 'a::type => bool)
-   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
-   Eps (approx R M x)"
-
-lemma the_fun_def: "ALL (R::'a::type => 'a::type => bool)
-   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
-   the_fun R M x = Eps (approx R M x)"
-  by (import relation the_fun_def)
-
-definition WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" where 
+  the_fun_primdef: "the_fun == %R M x. Eps (approx R M x)"
+
+lemma the_fun_def: "the_fun R M x = Eps (approx R M x)"
+  sorry
+
+definition
+  WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b"  where
   "WFREC ==
-%(R::'a::type => 'a::type => bool)
-   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
-   M (RESTRICT
-       (the_fun (TC R)
-         (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v) x)
-       R x)
-    x"
-
-lemma WFREC_DEF: "ALL (R::'a::type => 'a::type => bool)
-   M::('a::type => 'b::type) => 'a::type => 'b::type.
-   WFREC R M =
-   (%x::'a::type.
-       M (RESTRICT
-           (the_fun (TC R)
-             (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v)
-             x)
-           R x)
-        x)"
-  by (import relation WFREC_DEF)
-
-lemma WFREC_THM: "ALL (R::'a::type => 'a::type => bool)
-   M::('a::type => 'b::type) => 'a::type => 'b::type.
-   WF R --> (ALL x::'a::type. WFREC R M x = M (RESTRICT (WFREC R M) R x) x)"
-  by (import relation WFREC_THM)
-
-lemma WFREC_COROLLARY: "ALL (M::('a::type => 'b::type) => 'a::type => 'b::type)
-   (R::'a::type => 'a::type => bool) f::'a::type => 'b::type.
-   f = WFREC R M --> WF R --> (ALL x::'a::type. f x = M (RESTRICT f R x) x)"
-  by (import relation WFREC_COROLLARY)
-
-lemma WF_RECURSION_THM: "ALL R::'a::type => 'a::type => bool.
-   WF R -->
-   (ALL M::('a::type => 'b::type) => 'a::type => 'b::type.
-       EX! f::'a::type => 'b::type.
-          ALL x::'a::type. f x = M (RESTRICT f R x) x)"
-  by (import relation WF_RECURSION_THM)
+%R M x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x"
+
+lemma WFREC_DEF: "WFREC R M =
+(%x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x)"
+  sorry
+
+lemma WFREC_THM: "WF R ==> WFREC R M x = M (RESTRICT (WFREC R M) R x) x"
+  sorry
+
+lemma WFREC_COROLLARY: "[| f = WFREC R M; WF R |] ==> f x = M (RESTRICT f R x) x"
+  sorry
+
+lemma WF_RECURSION_THM: "WF R ==> EX! f. ALL x. f x = M (RESTRICT f R x) x"
+  sorry
 
 ;end_setup
 
 ;setup_theory pair
 
-lemma CURRY_ONE_ONE_THM: "(curry (f::'a::type * 'b::type => 'c::type) =
- curry (g::'a::type * 'b::type => 'c::type)) =
-(f = g)"
-  by (import pair CURRY_ONE_ONE_THM)
-
-lemma UNCURRY_ONE_ONE_THM: "(op =::bool => bool => bool)
- ((op =::('a::type * 'b::type => 'c::type)
-         => ('a::type * 'b::type => 'c::type) => bool)
-   ((split::('a::type => 'b::type => 'c::type)
-            => 'a::type * 'b::type => 'c::type)
-     (f::'a::type => 'b::type => 'c::type))
-   ((split::('a::type => 'b::type => 'c::type)
-            => 'a::type * 'b::type => 'c::type)
-     (g::'a::type => 'b::type => 'c::type)))
- ((op =::('a::type => 'b::type => 'c::type)
-         => ('a::type => 'b::type => 'c::type) => bool)
-   f g)"
-  by (import pair UNCURRY_ONE_ONE_THM)
-
-lemma pair_Axiom: "ALL f::'a::type => 'b::type => 'c::type.
-   EX x::'a::type * 'b::type => 'c::type.
-      ALL (xa::'a::type) y::'b::type. x (xa, y) = f xa y"
-  by (import pair pair_Axiom)
-
-lemma UNCURRY_CONG: "ALL (M::'a::type * 'b::type) (M'::'a::type * 'b::type)
-   f::'a::type => 'b::type => 'c::type.
-   M = M' &
-   (ALL (x::'a::type) y::'b::type.
-       M' = (x, y) -->
-       f x y = (f'::'a::type => 'b::type => 'c::type) x y) -->
-   split f M = split f' M'"
-  by (import pair UNCURRY_CONG)
-
-lemma ELIM_PEXISTS: "(EX p::'a::type * 'b::type.
-    (P::'a::type => 'b::type => bool) (fst p) (snd p)) =
-(EX p1::'a::type. Ex (P p1))"
-  by (import pair ELIM_PEXISTS)
-
-lemma ELIM_PFORALL: "(ALL p::'a::type * 'b::type.
-    (P::'a::type => 'b::type => bool) (fst p) (snd p)) =
-(ALL p1::'a::type. All (P p1))"
-  by (import pair ELIM_PFORALL)
-
-lemma PFORALL_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool)
- (%x::'a::type => 'b::type => bool.
-     (op =::bool => bool => bool)
-      ((All::('a::type => bool) => bool)
-        (%xa::'a::type. (All::('b::type => bool) => bool) (x xa)))
-      ((All::('a::type * 'b::type => bool) => bool)
-        ((split::('a::type => 'b::type => bool)
-                 => 'a::type * 'b::type => bool)
-          x)))"
-  by (import pair PFORALL_THM)
-
-lemma PEXISTS_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool)
- (%x::'a::type => 'b::type => bool.
-     (op =::bool => bool => bool)
-      ((Ex::('a::type => bool) => bool)
-        (%xa::'a::type. (Ex::('b::type => bool) => bool) (x xa)))
-      ((Ex::('a::type * 'b::type => bool) => bool)
-        ((split::('a::type => 'b::type => bool)
-                 => 'a::type * 'b::type => bool)
-          x)))"
-  by (import pair PEXISTS_THM)
-
-lemma LET2_RAND: "(All::(('c::type => 'd::type) => bool) => bool)
- (%x::'c::type => 'd::type.
-     (All::('a::type * 'b::type => bool) => bool)
-      (%xa::'a::type * 'b::type.
-          (All::(('a::type => 'b::type => 'c::type) => bool) => bool)
-           (%xb::'a::type => 'b::type => 'c::type.
-               (op =::'d::type => 'd::type => bool)
-                (x ((Let::'a::type * 'b::type
-                          => ('a::type * 'b::type => 'c::type) => 'c::type)
-                     xa ((split::('a::type => 'b::type => 'c::type)
-                                 => 'a::type * 'b::type => 'c::type)
-                          xb)))
-                ((Let::'a::type * 'b::type
-                       => ('a::type * 'b::type => 'd::type) => 'd::type)
-                  xa ((split::('a::type => 'b::type => 'd::type)
-                              => 'a::type * 'b::type => 'd::type)
-                       (%(xa::'a::type) y::'b::type. x (xb xa y)))))))"
-  by (import pair LET2_RAND)
-
-lemma LET2_RATOR: "(All::('a1::type * 'a2::type => bool) => bool)
- (%x::'a1::type * 'a2::type.
-     (All::(('a1::type => 'a2::type => 'b::type => 'c::type) => bool)
-           => bool)
-      (%xa::'a1::type => 'a2::type => 'b::type => 'c::type.
-          (All::('b::type => bool) => bool)
-           (%xb::'b::type.
-               (op =::'c::type => 'c::type => bool)
-                ((Let::'a1::type * 'a2::type
-                       => ('a1::type * 'a2::type => 'b::type => 'c::type)
-                          => 'b::type => 'c::type)
-                  x ((split::('a1::type
-                              => 'a2::type => 'b::type => 'c::type)
-                             => 'a1::type * 'a2::type
-                                => 'b::type => 'c::type)
-                      xa)
-                  xb)
-                ((Let::'a1::type * 'a2::type
-                       => ('a1::type * 'a2::type => 'c::type) => 'c::type)
-                  x ((split::('a1::type => 'a2::type => 'c::type)
-                             => 'a1::type * 'a2::type => 'c::type)
-                      (%(x::'a1::type) y::'a2::type. xa x y xb))))))"
-  by (import pair LET2_RATOR)
-
-lemma pair_case_cong: "ALL (x::'a::type * 'b::type) (xa::'a::type * 'b::type)
-   xb::'a::type => 'b::type => 'c::type.
-   x = xa &
-   (ALL (x::'a::type) y::'b::type.
-       xa = (x, y) -->
-       xb x y = (f'::'a::type => 'b::type => 'c::type) x y) -->
-   split xb x = split f' xa"
-  by (import pair pair_case_cong)
-
-definition LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where 
-  "LEX ==
-%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
-   (s::'a::type, t::'b::type) (u::'a::type, v::'b::type).
-   R1 s u | s = u & R2 t v"
-
-lemma LEX_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool.
-   LEX R1 R2 =
-   (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type).
-       R1 s u | s = u & R2 t v)"
-  by (import pair LEX_DEF)
-
-lemma WF_LEX: "ALL (x::'a::type => 'a::type => bool) xa::'b::type => 'b::type => bool.
-   WF x & WF xa --> WF (LEX x xa)"
-  by (import pair WF_LEX)
-
-definition RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where 
-  "RPROD ==
-%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
-   (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v"
-
-lemma RPROD_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool.
-   RPROD R1 R2 =
-   (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v)"
-  by (import pair RPROD_DEF)
-
-lemma WF_RPROD: "ALL (R::'a::type => 'a::type => bool) Q::'b::type => 'b::type => bool.
-   WF R & WF Q --> WF (RPROD R Q)"
-  by (import pair WF_RPROD)
+lemma CURRY_ONE_ONE_THM: "(curry f = curry g) = (f = g)"
+  sorry
+
+lemma UNCURRY_ONE_ONE_THM: "((%(x, y). f x y) = (%(x, y). g x y)) = (f = g)"
+  sorry
+
+lemma pair_Axiom: "EX x. ALL xa y. x (xa, y) = f xa y"
+  sorry
+
+lemma UNCURRY_CONG: "M = M' & (ALL x y. M' = (x, y) --> f x y = f' x y)
+==> prod_case f M = prod_case f' M'"
+  sorry
+
+lemma ELIM_PEXISTS: "(EX p. P (fst p) (snd p)) = (EX p1. Ex (P p1))"
+  sorry
+
+lemma ELIM_PFORALL: "(ALL p. P (fst p) (snd p)) = (ALL p1. All (P p1))"
+  sorry
+
+lemma PFORALL_THM: "(ALL xa. All (x xa)) = All (%(xa, y). x xa y)"
+  sorry
+
+lemma PEXISTS_THM: "(EX xa. Ex (x xa)) = Ex (%(xa, y). x xa y)"
+  sorry
+
+lemma LET2_RAND: "(x::'c => 'd)
+ (let (x::'a, y::'b) = xa::'a * 'b in (xb::'a => 'b => 'c) x y) =
+(let (xa::'a, y::'b) = xa in x (xb xa y))"
+  sorry
+
+lemma LET2_RATOR: "(let (x::'a1, y::'a2) = x::'a1 * 'a2 in (xa::'a1 => 'a2 => 'b => 'c) x y)
+ (xb::'b) =
+(let (x::'a1, y::'a2) = x in xa x y xb)"
+  sorry
+
+lemma pair_case_cong: "x = xa & (ALL x y. xa = (x, y) --> xb x y = f' x y)
+==> prod_case xb x = prod_case f' xa"
+  sorry
+
+definition
+  LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool"  where
+  "LEX == %R1 R2 (s, t) (u, v). R1 s u | s = u & R2 t v"
+
+lemma LEX_DEF: "LEX R1 R2 = (%(s, t) (u, v). R1 s u | s = u & R2 t v)"
+  sorry
+
+lemma WF_LEX: "WF x & WF xa ==> WF (LEX x xa)"
+  sorry
+
+definition
+  RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool"  where
+  "RPROD == %R1 R2 (s, t) (u, v). R1 s u & R2 t v"
+
+lemma RPROD_DEF: "RPROD R1 R2 = (%(s, t) (u, v). R1 s u & R2 t v)"
+  sorry
+
+lemma WF_RPROD: "WF R & WF Q ==> WF (RPROD R Q)"
+  sorry
 
 ;end_setup
 
@@ -1073,174 +828,113 @@
 ;setup_theory prim_rec
 
 lemma LESS_0_0: "0 < Suc 0"
-  by (import prim_rec LESS_0_0)
-
-lemma LESS_LEMMA1: "ALL (x::nat) xa::nat. x < Suc xa --> x = xa | x < xa"
-  by (import prim_rec LESS_LEMMA1)
-
-lemma LESS_LEMMA2: "ALL (m::nat) n::nat. m = n | m < n --> m < Suc n"
-  by (import prim_rec LESS_LEMMA2)
-
-lemma LESS_THM: "ALL (m::nat) n::nat. (m < Suc n) = (m = n | m < n)"
-  by (import prim_rec LESS_THM)
-
-lemma LESS_SUC_IMP: "ALL (x::nat) xa::nat. x < Suc xa --> x ~= xa --> x < xa"
-  by (import prim_rec LESS_SUC_IMP)
-
-lemma EQ_LESS: "ALL n::nat. Suc (m::nat) = n --> m < n"
-  by (import prim_rec EQ_LESS)
-
-lemma NOT_LESS_EQ: "ALL (m::nat) n::nat. m = n --> ~ m < n"
-  by (import prim_rec NOT_LESS_EQ)
-
-definition SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" where 
-  "(op ==::((nat => 'a::type)
-         => 'a::type => ('a::type => 'a::type) => nat => bool)
-        => ((nat => 'a::type)
-            => 'a::type => ('a::type => 'a::type) => nat => bool)
-           => prop)
- (SIMP_REC_REL::(nat => 'a::type)
-                => 'a::type => ('a::type => 'a::type) => nat => bool)
- (%(fun::nat => 'a::type) (x::'a::type) (f::'a::type => 'a::type) n::nat.
-     (op &::bool => bool => bool)
-      ((op =::'a::type => 'a::type => bool) (fun (0::nat)) x)
-      ((All::(nat => bool) => bool)
-        (%m::nat.
-            (op -->::bool => bool => bool) ((op <::nat => nat => bool) m n)
-             ((op =::'a::type => 'a::type => bool)
-               (fun ((Suc::nat => nat) m)) (f (fun m))))))"
-
-lemma SIMP_REC_REL: "(All::((nat => 'a::type) => bool) => bool)
- (%fun::nat => 'a::type.
-     (All::('a::type => bool) => bool)
-      (%x::'a::type.
-          (All::(('a::type => 'a::type) => bool) => bool)
-           (%f::'a::type => 'a::type.
-               (All::(nat => bool) => bool)
-                (%n::nat.
-                    (op =::bool => bool => bool)
-                     ((SIMP_REC_REL::(nat => 'a::type)
-                                     => 'a::type
-  => ('a::type => 'a::type) => nat => bool)
-                       fun x f n)
-                     ((op &::bool => bool => bool)
-                       ((op =::'a::type => 'a::type => bool) (fun (0::nat))
-                         x)
-                       ((All::(nat => bool) => bool)
-                         (%m::nat.
-                             (op -->::bool => bool => bool)
-                              ((op <::nat => nat => bool) m n)
-                              ((op =::'a::type => 'a::type => bool)
-                                (fun ((Suc::nat => nat) m))
-                                (f (fun m))))))))))"
-  by (import prim_rec SIMP_REC_REL)
-
-lemma SIMP_REC_EXISTS: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat.
-   EX fun::nat => 'a::type. SIMP_REC_REL fun x f n"
-  by (import prim_rec SIMP_REC_EXISTS)
-
-lemma SIMP_REC_REL_UNIQUE: "ALL (x::'a::type) (xa::'a::type => 'a::type) (xb::nat => 'a::type)
-   (xc::nat => 'a::type) (xd::nat) xe::nat.
-   SIMP_REC_REL xb x xa xd & SIMP_REC_REL xc x xa xe -->
-   (ALL n::nat. n < xd & n < xe --> xb n = xc n)"
-  by (import prim_rec SIMP_REC_REL_UNIQUE)
-
-lemma SIMP_REC_REL_UNIQUE_RESULT: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat.
-   EX! y::'a::type.
-      EX g::nat => 'a::type. SIMP_REC_REL g x f (Suc n) & y = g n"
-  by (import prim_rec SIMP_REC_REL_UNIQUE_RESULT)
+  sorry
+
+lemma LESS_LEMMA1: "x < Suc xa ==> x = xa | x < xa"
+  sorry
+
+lemma LESS_LEMMA2: "m = n | m < n ==> m < Suc n"
+  sorry
+
+lemma LESS_THM: "(m < Suc n) = (m = n | m < n)"
+  sorry
+
+lemma LESS_SUC_IMP: "[| x < Suc xa; x ~= xa |] ==> x < xa"
+  sorry
+
+lemma EQ_LESS: "Suc m = n ==> m < n"
+  sorry
+
+lemma NOT_LESS_EQ: "(m::nat) = (n::nat) ==> ~ m < n"
+  sorry
+
+definition
+  SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool"  where
+  "SIMP_REC_REL == %fun x f n. fun 0 = x & (ALL m<n. fun (Suc m) = f (fun m))"
+
+lemma SIMP_REC_REL: "SIMP_REC_REL fun x f n = (fun 0 = x & (ALL m<n. fun (Suc m) = f (fun m)))"
+  sorry
+
+lemma SIMP_REC_EXISTS: "EX fun. SIMP_REC_REL fun x f n"
+  sorry
+
+lemma SIMP_REC_REL_UNIQUE: "[| SIMP_REC_REL xb x xa xd & SIMP_REC_REL xc x xa xe; n < xd & n < xe |]
+==> xb n = xc n"
+  sorry
+
+lemma SIMP_REC_REL_UNIQUE_RESULT: "EX! y. EX g. SIMP_REC_REL g x f (Suc n) & y = g n"
+  sorry
 
 consts
   SIMP_REC :: "'a => ('a => 'a) => nat => 'a" 
 
-specification (SIMP_REC) SIMP_REC: "ALL (x::'a::type) (f'::'a::type => 'a::type) n::nat.
-   EX g::nat => 'a::type.
-      SIMP_REC_REL g x f' (Suc n) & SIMP_REC x f' n = g n"
-  by (import prim_rec SIMP_REC)
-
-lemma LESS_SUC_SUC: "ALL m::nat. m < Suc m & m < Suc (Suc m)"
-  by (import prim_rec LESS_SUC_SUC)
-
-lemma SIMP_REC_THM: "ALL (x::'a::type) f::'a::type => 'a::type.
-   SIMP_REC x f 0 = x &
-   (ALL m::nat. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))"
-  by (import prim_rec SIMP_REC_THM)
-
-definition PRE :: "nat => nat" where 
-  "PRE == %m::nat. if m = 0 then 0 else SOME n::nat. m = Suc n"
-
-lemma PRE_DEF: "ALL m::nat. PRE m = (if m = 0 then 0 else SOME n::nat. m = Suc n)"
-  by (import prim_rec PRE_DEF)
-
-lemma PRE: "PRE 0 = 0 & (ALL m::nat. PRE (Suc m) = m)"
-  by (import prim_rec PRE)
-
-definition PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" where 
-  "PRIM_REC_FUN ==
-%(x::'a::type) f::'a::type => nat => 'a::type.
-   SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)"
-
-lemma PRIM_REC_FUN: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
-   PRIM_REC_FUN x f =
-   SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)"
-  by (import prim_rec PRIM_REC_FUN)
-
-lemma PRIM_REC_EQN: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
-   (ALL n::nat. PRIM_REC_FUN x f 0 n = x) &
-   (ALL (m::nat) n::nat.
-       PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)"
-  by (import prim_rec PRIM_REC_EQN)
-
-definition PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" where 
-  "PRIM_REC ==
-%(x::'a::type) (f::'a::type => nat => 'a::type) m::nat.
-   PRIM_REC_FUN x f m (PRE m)"
-
-lemma PRIM_REC: "ALL (x::'a::type) (f::'a::type => nat => 'a::type) m::nat.
-   PRIM_REC x f m = PRIM_REC_FUN x f m (PRE m)"
-  by (import prim_rec PRIM_REC)
-
-lemma PRIM_REC_THM: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
-   PRIM_REC x f 0 = x &
-   (ALL m::nat. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)"
-  by (import prim_rec PRIM_REC_THM)
-
-lemma DC: "ALL (P::'a::type => bool) (R::'a::type => 'a::type => bool) a::'a::type.
-   P a & (ALL x::'a::type. P x --> (EX y::'a::type. P y & R x y)) -->
-   (EX x::nat => 'a::type.
-       x 0 = a & (ALL n::nat. P (x n) & R (x n) (x (Suc n))))"
-  by (import prim_rec DC)
-
-lemma num_Axiom_old: "ALL (e::'a::type) f::'a::type => nat => 'a::type.
-   EX! fn1::nat => 'a::type.
-      fn1 0 = e & (ALL n::nat. fn1 (Suc n) = f (fn1 n) n)"
-  by (import prim_rec num_Axiom_old)
-
-lemma num_Axiom: "ALL (e::'a::type) f::nat => 'a::type => 'a::type.
-   EX x::nat => 'a::type. x 0 = e & (ALL n::nat. x (Suc n) = f n (x n))"
-  by (import prim_rec num_Axiom)
+specification (SIMP_REC) SIMP_REC: "ALL x f' n. EX g. SIMP_REC_REL g x f' (Suc n) & SIMP_REC x f' n = g n"
+  sorry
+
+lemma LESS_SUC_SUC: "m < Suc m & m < Suc (Suc m)"
+  sorry
+
+lemma SIMP_REC_THM: "SIMP_REC x f 0 = x & (ALL m. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))"
+  sorry
+
+definition
+  PRE :: "nat => nat"  where
+  "PRE == %m. if m = 0 then 0 else SOME n. m = Suc n"
+
+lemma PRE_DEF: "PRE m = (if m = 0 then 0 else SOME n. m = Suc n)"
+  sorry
+
+lemma PRE: "PRE 0 = 0 & (ALL m. PRE (Suc m) = m)"
+  sorry
+
+definition
+  PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a"  where
+  "PRIM_REC_FUN == %x f. SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)"
+
+lemma PRIM_REC_FUN: "PRIM_REC_FUN x f = SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)"
+  sorry
+
+lemma PRIM_REC_EQN: "(ALL n. PRIM_REC_FUN x f 0 n = x) &
+(ALL m n. PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)"
+  sorry
+
+definition
+  PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a"  where
+  "PRIM_REC == %x f m. PRIM_REC_FUN x f m (PRE m)"
+
+lemma PRIM_REC: "PRIM_REC x f m = PRIM_REC_FUN x f m (PRE m)"
+  sorry
+
+lemma PRIM_REC_THM: "PRIM_REC x f 0 = x & (ALL m. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)"
+  sorry
+
+lemma DC: "P a & (ALL x. P x --> (EX y. P y & R x y))
+==> EX x. x 0 = a & (ALL n. P (x n) & R (x n) (x (Suc n)))"
+  sorry
+
+lemma num_Axiom_old: "EX! fn1. fn1 0 = e & (ALL n. fn1 (Suc n) = f (fn1 n) n)"
+  sorry
+
+lemma num_Axiom: "EX x. x 0 = e & (ALL n. x (Suc n) = f n (x n))"
+  sorry
 
 consts
   wellfounded :: "('a => 'a => bool) => bool" 
 
 defs
-  wellfounded_primdef: "wellfounded ==
-%R::'a::type => 'a::type => bool.
-   ~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n))"
-
-lemma wellfounded_def: "ALL R::'a::type => 'a::type => bool.
-   wellfounded R =
-   (~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n)))"
-  by (import prim_rec wellfounded_def)
-
-lemma WF_IFF_WELLFOUNDED: "ALL R::'a::type => 'a::type => bool. WF R = wellfounded R"
-  by (import prim_rec WF_IFF_WELLFOUNDED)
-
-lemma WF_PRED: "WF (%(x::nat) y::nat. y = Suc x)"
-  by (import prim_rec WF_PRED)
+  wellfounded_primdef: "wellfounded == %R. ~ (EX f. ALL n. R (f (Suc n)) (f n))"
+
+lemma wellfounded_def: "wellfounded R = (~ (EX f. ALL n. R (f (Suc n)) (f n)))"
+  sorry
+
+lemma WF_IFF_WELLFOUNDED: "WF R = wellfounded R"
+  sorry
+
+lemma WF_PRED: "WF (%x y. y = Suc x)"
+  sorry
 
 lemma WF_LESS: "(WF::(nat => nat => bool) => bool) (op <::nat => nat => bool)"
-  by (import prim_rec WF_LESS)
+  sorry
 
 consts
   measure :: "('a => nat) => 'a => 'a => bool" 
@@ -1249,616 +943,533 @@
   measure_primdef: "prim_rec.measure == relation.inv_image op <"
 
 lemma measure_def: "prim_rec.measure = relation.inv_image op <"
-  by (import prim_rec measure_def)
-
-lemma WF_measure: "ALL x::'a::type => nat. WF (prim_rec.measure x)"
-  by (import prim_rec WF_measure)
-
-lemma measure_thm: "ALL (x::'a::type => nat) (xa::'a::type) xb::'a::type.
-   prim_rec.measure x xa xb = (x xa < x xb)"
-  by (import prim_rec measure_thm)
+  sorry
+
+lemma WF_measure: "WF (prim_rec.measure x)"
+  sorry
+
+lemma measure_thm: "prim_rec.measure x xa xb = (x xa < x xb)"
+  sorry
 
 ;end_setup
 
 ;setup_theory arithmetic
 
-definition nat_elim__magic :: "nat => nat" where 
-  "nat_elim__magic == %n::nat. n"
-
-lemma nat_elim__magic: "ALL n::nat. nat_elim__magic n = n"
-  by (import arithmetic nat_elim__magic)
+definition
+  nat_elim__magic :: "nat => nat"  where
+  "nat_elim__magic == %n. n"
+
+lemma nat_elim__magic: "nat_elim__magic n = n"
+  sorry
 
 consts
   EVEN :: "nat => bool" 
 
-specification (EVEN) EVEN: "EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))"
-  by (import arithmetic EVEN)
+specification (EVEN) EVEN: "EVEN 0 = True & (ALL n. EVEN (Suc n) = (~ EVEN n))"
+  sorry
 
 consts
   ODD :: "nat => bool" 
 
-specification (ODD) ODD: "ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))"
-  by (import arithmetic ODD)
+specification (ODD) ODD: "ODD 0 = False & (ALL n. ODD (Suc n) = (~ ODD n))"
+  sorry
 
 lemma TWO: "2 = Suc 1"
-  by (import arithmetic TWO)
-
-lemma NORM_0: "(op =::nat => nat => bool) (0::nat) (0::nat)"
-  by (import arithmetic NORM_0)
-
-lemma num_case_compute: "ALL n::nat.
-   nat_case (f::'a::type) (g::nat => 'a::type) n =
-   (if n = 0 then f else g (PRE n))"
-  by (import arithmetic num_case_compute)
-
-lemma ADD_CLAUSES: "0 + (m::nat) = m &
-m + 0 = m & Suc m + (n::nat) = Suc (m + n) & m + Suc n = Suc (m + n)"
-  by (import arithmetic ADD_CLAUSES)
-
-lemma LESS_ADD: "ALL (m::nat) n::nat. n < m --> (EX p::nat. p + n = m)"
-  by (import arithmetic LESS_ADD)
-
-lemma LESS_ANTISYM: "ALL (m::nat) n::nat. ~ (m < n & n < m)"
-  by (import arithmetic LESS_ANTISYM)
-
-lemma LESS_LESS_SUC: "ALL (x::nat) xa::nat. ~ (x < xa & xa < Suc x)"
-  by (import arithmetic LESS_LESS_SUC)
-
-lemma FUN_EQ_LEMMA: "ALL (f::'a::type => bool) (x1::'a::type) x2::'a::type.
-   f x1 & ~ f x2 --> x1 ~= x2"
-  by (import arithmetic FUN_EQ_LEMMA)
-
-lemma LESS_NOT_SUC: "ALL (m::nat) n::nat. m < n & n ~= Suc m --> Suc m < n"
-  by (import arithmetic LESS_NOT_SUC)
-
-lemma LESS_0_CASES: "ALL m::nat. 0 = m | 0 < m"
-  by (import arithmetic LESS_0_CASES)
-
-lemma LESS_CASES_IMP: "ALL (m::nat) n::nat. ~ m < n & m ~= n --> n < m"
-  by (import arithmetic LESS_CASES_IMP)
-
-lemma LESS_CASES: "ALL (m::nat) n::nat. m < n | n <= m"
-  by (import arithmetic LESS_CASES)
-
-lemma LESS_EQ_SUC_REFL: "ALL m::nat. m <= Suc m"
-  by (import arithmetic LESS_EQ_SUC_REFL)
-
-lemma LESS_ADD_NONZERO: "ALL (m::nat) n::nat. n ~= 0 --> m < m + n"
-  by (import arithmetic LESS_ADD_NONZERO)
-
-lemma LESS_EQ_ANTISYM: "ALL (x::nat) xa::nat. ~ (x < xa & xa <= x)"
-  by (import arithmetic LESS_EQ_ANTISYM)
-
-lemma SUB_0: "ALL m::nat. 0 - m = 0 & m - 0 = m"
-  by (import arithmetic SUB_0)
-
-lemma SUC_SUB1: "ALL m::nat. Suc m - 1 = m"
-  by (import arithmetic SUC_SUB1)
-
-lemma PRE_SUB1: "ALL m::nat. PRE m = m - 1"
-  by (import arithmetic PRE_SUB1)
-
-lemma MULT_CLAUSES: "ALL (x::nat) xa::nat.
-   0 * x = 0 &
-   x * 0 = 0 &
-   1 * x = x &
-   x * 1 = x & Suc x * xa = x * xa + xa & x * Suc xa = x + x * xa"
-  by (import arithmetic MULT_CLAUSES)
-
-lemma PRE_SUB: "ALL (m::nat) n::nat. PRE (m - n) = PRE m - n"
-  by (import arithmetic PRE_SUB)
-
-lemma ADD_EQ_1: "ALL (m::nat) n::nat. (m + n = 1) = (m = 1 & n = 0 | m = 0 & n = 1)"
-  by (import arithmetic ADD_EQ_1)
-
-lemma ADD_INV_0_EQ: "ALL (m::nat) n::nat. (m + n = m) = (n = 0)"
-  by (import arithmetic ADD_INV_0_EQ)
-
-lemma PRE_SUC_EQ: "ALL (m::nat) n::nat. 0 < n --> (m = PRE n) = (Suc m = n)"
-  by (import arithmetic PRE_SUC_EQ)
-
-lemma INV_PRE_EQ: "ALL (m::nat) n::nat. 0 < m & 0 < n --> (PRE m = PRE n) = (m = n)"
-  by (import arithmetic INV_PRE_EQ)
-
-lemma LESS_SUC_NOT: "ALL (m::nat) n::nat. m < n --> ~ n < Suc m"
-  by (import arithmetic LESS_SUC_NOT)
-
-lemma ADD_EQ_SUB: "ALL (m::nat) (n::nat) p::nat. n <= p --> (m + n = p) = (m = p - n)"
-  by (import arithmetic ADD_EQ_SUB)
-
-lemma LESS_ADD_1: "ALL (x::nat) xa::nat. xa < x --> (EX xb::nat. x = xa + (xb + 1))"
-  by (import arithmetic LESS_ADD_1)
-
-lemma NOT_ODD_EQ_EVEN: "ALL (n::nat) m::nat. Suc (n + n) ~= m + m"
-  by (import arithmetic NOT_ODD_EQ_EVEN)
-
-lemma MULT_SUC_EQ: "ALL (p::nat) (m::nat) n::nat. (n * Suc p = m * Suc p) = (n = m)"
-  by (import arithmetic MULT_SUC_EQ)
-
-lemma MULT_EXP_MONO: "ALL (p::nat) (q::nat) (n::nat) m::nat.
-   (n * Suc q ^ p = m * Suc q ^ p) = (n = m)"
-  by (import arithmetic MULT_EXP_MONO)
-
-lemma LESS_ADD_SUC: "ALL (m::nat) n::nat. m < m + Suc n"
-  by (import arithmetic LESS_ADD_SUC)
-
-lemma LESS_OR_EQ_ADD: "ALL (n::nat) m::nat. n < m | (EX p::nat. n = p + m)"
-  by (import arithmetic LESS_OR_EQ_ADD)
-
-lemma WOP: "(All::((nat => bool) => bool) => bool)
- (%P::nat => bool.
-     (op -->::bool => bool => bool) ((Ex::(nat => bool) => bool) P)
-      ((Ex::(nat => bool) => bool)
-        (%n::nat.
-            (op &::bool => bool => bool) (P n)
-             ((All::(nat => bool) => bool)
-               (%m::nat.
-                   (op -->::bool => bool => bool)
-                    ((op <::nat => nat => bool) m n)
-                    ((Not::bool => bool) (P m)))))))"
-  by (import arithmetic WOP)
-
-lemma INV_PRE_LESS: "ALL m>0. ALL n::nat. (PRE m < PRE n) = (m < n)"
-  by (import arithmetic INV_PRE_LESS)
-
-lemma INV_PRE_LESS_EQ: "ALL n>0. ALL m::nat. (PRE m <= PRE n) = (m <= n)"
-  by (import arithmetic INV_PRE_LESS_EQ)
-
-lemma SUB_EQ_EQ_0: "ALL (m::nat) n::nat. (m - n = m) = (m = 0 | n = 0)"
-  by (import arithmetic SUB_EQ_EQ_0)
-
-lemma SUB_LESS_OR: "ALL (m::nat) n::nat. n < m --> n <= m - 1"
-  by (import arithmetic SUB_LESS_OR)
-
-lemma LESS_SUB_ADD_LESS: "ALL (n::nat) (m::nat) i::nat. i < n - m --> i + m < n"
-  by (import arithmetic LESS_SUB_ADD_LESS)
-
-lemma LESS_EQ_SUB_LESS: "ALL (x::nat) xa::nat. xa <= x --> (ALL c::nat. (x - xa < c) = (x < xa + c))"
-  by (import arithmetic LESS_EQ_SUB_LESS)
-
-lemma NOT_SUC_LESS_EQ: "ALL (x::nat) xa::nat. (~ Suc x <= xa) = (xa <= x)"
-  by (import arithmetic NOT_SUC_LESS_EQ)
-
-lemma SUB_LESS_EQ_ADD: "ALL (m::nat) p::nat. m <= p --> (ALL n::nat. (p - m <= n) = (p <= m + n))"
-  by (import arithmetic SUB_LESS_EQ_ADD)
-
-lemma SUB_CANCEL: "ALL (x::nat) (xa::nat) xb::nat.
-   xa <= x & xb <= x --> (x - xa = x - xb) = (xa = xb)"
-  by (import arithmetic SUB_CANCEL)
-
-lemma NOT_EXP_0: "ALL (m::nat) n::nat. Suc n ^ m ~= 0"
-  by (import arithmetic NOT_EXP_0)
-
-lemma ZERO_LESS_EXP: "ALL (m::nat) n::nat. 0 < Suc n ^ m"
-  by (import arithmetic ZERO_LESS_EXP)
-
-lemma ODD_OR_EVEN: "ALL x::nat. EX xa::nat. x = Suc (Suc 0) * xa | x = Suc (Suc 0) * xa + 1"
-  by (import arithmetic ODD_OR_EVEN)
-
-lemma LESS_EXP_SUC_MONO: "ALL (n::nat) m::nat. Suc (Suc m) ^ n < Suc (Suc m) ^ Suc n"
-  by (import arithmetic LESS_EXP_SUC_MONO)
-
-lemma LESS_LESS_CASES: "ALL (m::nat) n::nat. m = n | m < n | n < m"
-  by (import arithmetic LESS_LESS_CASES)
-
-lemma LESS_EQUAL_ADD: "ALL (m::nat) n::nat. m <= n --> (EX p::nat. n = m + p)"
-  by (import arithmetic LESS_EQUAL_ADD)
-
-lemma MULT_EQ_1: "ALL (x::nat) y::nat. (x * y = 1) = (x = 1 & y = 1)"
-  by (import arithmetic MULT_EQ_1)
+  sorry
+
+lemma NORM_0: "(0::nat) = (0::nat)"
+  sorry
+
+lemma num_case_compute: "nat_case f g n = (if n = 0 then f else g (PRE n))"
+  sorry
+
+lemma ADD_CLAUSES: "0 + m = m & m + 0 = m & Suc m + n = Suc (m + n) & m + Suc n = Suc (m + n)"
+  sorry
+
+lemma LESS_ADD: "(n::nat) < (m::nat) ==> EX p::nat. p + n = m"
+  sorry
+
+lemma LESS_ANTISYM: "~ ((m::nat) < (n::nat) & n < m)"
+  sorry
+
+lemma LESS_LESS_SUC: "~ (x < xa & xa < Suc x)"
+  sorry
+
+lemma FUN_EQ_LEMMA: "f x1 & ~ f x2 ==> x1 ~= x2"
+  sorry
+
+lemma LESS_NOT_SUC: "m < n & n ~= Suc m ==> Suc m < n"
+  sorry
+
+lemma LESS_0_CASES: "(0::nat) = (m::nat) | (0::nat) < m"
+  sorry
+
+lemma LESS_CASES_IMP: "~ (m::nat) < (n::nat) & m ~= n ==> n < m"
+  sorry
+
+lemma LESS_CASES: "(m::nat) < (n::nat) | n <= m"
+  sorry
+
+lemma LESS_EQ_SUC_REFL: "m <= Suc m"
+  sorry
+
+lemma LESS_ADD_NONZERO: "(n::nat) ~= (0::nat) ==> (m::nat) < m + n"
+  sorry
+
+lemma LESS_EQ_ANTISYM: "~ ((x::nat) < (xa::nat) & xa <= x)"
+  sorry
+
+lemma SUB_0: "(0::nat) - (m::nat) = (0::nat) & m - (0::nat) = m"
+  sorry
+
+lemma PRE_SUB1: "PRE m = m - 1"
+  sorry
+
+lemma MULT_CLAUSES: "0 * x = 0 &
+x * 0 = 0 &
+1 * x = x & x * 1 = x & Suc x * xa = x * xa + xa & x * Suc xa = x + x * xa"
+  sorry
+
+lemma PRE_SUB: "PRE (m - n) = PRE m - n"
+  sorry
+
+lemma ADD_EQ_1: "((m::nat) + (n::nat) = (1::nat)) =
+(m = (1::nat) & n = (0::nat) | m = (0::nat) & n = (1::nat))"
+  sorry
+
+lemma ADD_INV_0_EQ: "((m::nat) + (n::nat) = m) = (n = (0::nat))"
+  sorry
+
+lemma PRE_SUC_EQ: "0 < n ==> (m = PRE n) = (Suc m = n)"
+  sorry
+
+lemma INV_PRE_EQ: "0 < m & 0 < n ==> (PRE m = PRE n) = (m = n)"
+  sorry
+
+lemma LESS_SUC_NOT: "m < n ==> ~ n < Suc m"
+  sorry
+
+lemma ADD_EQ_SUB: "(n::nat) <= (p::nat) ==> ((m::nat) + n = p) = (m = p - n)"
+  sorry
+
+lemma LESS_ADD_1: "(xa::nat) < (x::nat) ==> EX xb::nat. x = xa + (xb + (1::nat))"
+  sorry
+
+lemma NOT_ODD_EQ_EVEN: "Suc (n + n) ~= m + m"
+  sorry
+
+lemma MULT_SUC_EQ: "(n * Suc p = m * Suc p) = (n = m)"
+  sorry
+
+lemma MULT_EXP_MONO: "(n * Suc q ^ p = m * Suc q ^ p) = (n = m)"
+  sorry
+
+lemma LESS_ADD_SUC: "m < m + Suc n"
+  sorry
+
+lemma LESS_OR_EQ_ADD: "(n::nat) < (m::nat) | (EX p::nat. n = p + m)"
+  sorry
+
+lemma WOP: "Ex (P::nat => bool) ==> EX n::nat. P n & (ALL m<n. ~ P m)"
+  sorry
+
+lemma INV_PRE_LESS: "0 < m ==> (PRE m < PRE n) = (m < n)"
+  sorry
+
+lemma INV_PRE_LESS_EQ: "0 < n ==> (PRE m <= PRE n) = (m <= n)"
+  sorry
+
+lemma SUB_EQ_EQ_0: "((m::nat) - (n::nat) = m) = (m = (0::nat) | n = (0::nat))"
+  sorry
+
+lemma SUB_LESS_OR: "(n::nat) < (m::nat) ==> n <= m - (1::nat)"
+  sorry
+
+lemma LESS_SUB_ADD_LESS: "(i::nat) < (n::nat) - (m::nat) ==> i + m < n"
+  sorry
+
+lemma LESS_EQ_SUB_LESS: "(xa::nat) <= (x::nat) ==> (x - xa < (c::nat)) = (x < xa + c)"
+  sorry
+
+lemma NOT_SUC_LESS_EQ: "(~ Suc x <= xa) = (xa <= x)"
+  sorry
+
+lemma SUB_LESS_EQ_ADD: "(m::nat) <= (p::nat) ==> (p - m <= (n::nat)) = (p <= m + n)"
+  sorry
+
+lemma SUB_CANCEL: "(xa::nat) <= (x::nat) & (xb::nat) <= x ==> (x - xa = x - xb) = (xa = xb)"
+  sorry
+
+lemma NOT_EXP_0: "Suc n ^ m ~= 0"
+  sorry
+
+lemma ZERO_LESS_EXP: "0 < Suc n ^ m"
+  sorry
+
+lemma ODD_OR_EVEN: "EX xa. x = Suc (Suc 0) * xa | x = Suc (Suc 0) * xa + 1"
+  sorry
+
+lemma LESS_EXP_SUC_MONO: "Suc (Suc m) ^ n < Suc (Suc m) ^ Suc n"
+  sorry
+
+lemma LESS_LESS_CASES: "(m::nat) = (n::nat) | m < n | n < m"
+  sorry
 
 consts
   FACT :: "nat => nat" 
 
-specification (FACT) FACT: "FACT 0 = 1 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)"
-  by (import arithmetic FACT)
-
-lemma FACT_LESS: "ALL n::nat. 0 < FACT n"
-  by (import arithmetic FACT_LESS)
-
-lemma EVEN_ODD: "ALL n::nat. EVEN n = (~ ODD n)"
-  by (import arithmetic EVEN_ODD)
-
-lemma ODD_EVEN: "ALL x::nat. ODD x = (~ EVEN x)"
-  by (import arithmetic ODD_EVEN)
-
-lemma EVEN_OR_ODD: "ALL x::nat. EVEN x | ODD x"
-  by (import arithmetic EVEN_OR_ODD)
-
-lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)"
-  by (import arithmetic EVEN_AND_ODD)
-
-lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)"
-  by (import arithmetic EVEN_ADD)
-
-lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)"
-  by (import arithmetic EVEN_MULT)
-
-lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)"
-  by (import arithmetic ODD_ADD)
-
-lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)"
-  by (import arithmetic ODD_MULT)
-
-lemma EVEN_DOUBLE: "ALL n::nat. EVEN (2 * n)"
-  by (import arithmetic EVEN_DOUBLE)
-
-lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (2 * x))"
-  by (import arithmetic ODD_DOUBLE)
-
-lemma EVEN_ODD_EXISTS: "ALL x::nat.
-   (EVEN x --> (EX m::nat. x = 2 * m)) &
-   (ODD x --> (EX m::nat. x = Suc (2 * m)))"
-  by (import arithmetic EVEN_ODD_EXISTS)
-
-lemma EVEN_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = 2 * m)"
-  by (import arithmetic EVEN_EXISTS)
-
-lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc (2 * m))"
-  by (import arithmetic ODD_EXISTS)
-
-lemma NOT_SUC_LESS_EQ_0: "ALL x::nat. ~ Suc x <= 0"
-  by (import arithmetic NOT_SUC_LESS_EQ_0)
-
-lemma NOT_LEQ: "ALL (x::nat) xa::nat. (~ x <= xa) = (Suc xa <= x)"
-  by (import arithmetic NOT_LEQ)
-
-lemma NOT_NUM_EQ: "ALL (x::nat) xa::nat. (x ~= xa) = (Suc x <= xa | Suc xa <= x)"
-  by (import arithmetic NOT_NUM_EQ)
-
-lemma NOT_GREATER_EQ: "ALL (x::nat) xa::nat. (~ xa <= x) = (Suc x <= xa)"
-  by (import arithmetic NOT_GREATER_EQ)
-
-lemma SUC_ADD_SYM: "ALL (m::nat) n::nat. Suc (m + n) = Suc n + m"
-  by (import arithmetic SUC_ADD_SYM)
-
-lemma NOT_SUC_ADD_LESS_EQ: "ALL (m::nat) n::nat. ~ Suc (m + n) <= m"
-  by (import arithmetic NOT_SUC_ADD_LESS_EQ)
-
-lemma SUB_LEFT_ADD: "ALL (m::nat) (n::nat) p::nat.
-   m + (n - p) = (if n <= p then m else m + n - p)"
-  by (import arithmetic SUB_LEFT_ADD)
-
-lemma SUB_RIGHT_ADD: "ALL (m::nat) (n::nat) p::nat. m - n + p = (if m <= n then p else m + p - n)"
-  by (import arithmetic SUB_RIGHT_ADD)
-
-lemma SUB_LEFT_SUB: "ALL (m::nat) (n::nat) p::nat.
-   m - (n - p) = (if n <= p then m else m + p - n)"
-  by (import arithmetic SUB_LEFT_SUB)
-
-lemma SUB_LEFT_SUC: "ALL (m::nat) n::nat. Suc (m - n) = (if m <= n then Suc 0 else Suc m - n)"
-  by (import arithmetic SUB_LEFT_SUC)
-
-lemma SUB_LEFT_LESS_EQ: "ALL (m::nat) (n::nat) p::nat. (m <= n - p) = (m + p <= n | m <= 0)"
-  by (import arithmetic SUB_LEFT_LESS_EQ)
-
-lemma SUB_RIGHT_LESS_EQ: "ALL (m::nat) (n::nat) p::nat. (m - n <= p) = (m <= n + p)"
-  by (import arithmetic SUB_RIGHT_LESS_EQ)
-
-lemma SUB_RIGHT_LESS: "ALL (m::nat) (n::nat) p::nat. (m - n < p) = (m < n + p & 0 < p)"
-  by (import arithmetic SUB_RIGHT_LESS)
-
-lemma SUB_RIGHT_GREATER_EQ: "ALL (m::nat) (n::nat) p::nat. (p <= m - n) = (n + p <= m | p <= 0)"
-  by (import arithmetic SUB_RIGHT_GREATER_EQ)
-
-lemma SUB_LEFT_GREATER: "ALL (m::nat) (n::nat) p::nat. (n - p < m) = (n < m + p & 0 < m)"
-  by (import arithmetic SUB_LEFT_GREATER)
-
-lemma SUB_RIGHT_GREATER: "ALL (m::nat) (n::nat) p::nat. (p < m - n) = (n + p < m)"
-  by (import arithmetic SUB_RIGHT_GREATER)
-
-lemma SUB_LEFT_EQ: "ALL (m::nat) (n::nat) p::nat. (m = n - p) = (m + p = n | m <= 0 & n <= p)"
-  by (import arithmetic SUB_LEFT_EQ)
-
-lemma SUB_RIGHT_EQ: "ALL (m::nat) (n::nat) p::nat. (m - n = p) = (m = n + p | m <= n & p <= 0)"
-  by (import arithmetic SUB_RIGHT_EQ)
-
-lemma LE: "(ALL n::nat. (n <= 0) = (n = 0)) &
+specification (FACT) FACT: "FACT 0 = 1 & (ALL n. FACT (Suc n) = Suc n * FACT n)"
+  sorry
+
+lemma FACT_LESS: "0 < FACT n"
+  sorry
+
+lemma EVEN_ODD: "EVEN n = (~ ODD n)"
+  sorry
+
+lemma ODD_EVEN: "ODD x = (~ EVEN x)"
+  sorry
+
+lemma EVEN_OR_ODD: "EVEN x | ODD x"
+  sorry
+
+lemma EVEN_AND_ODD: "~ (EVEN x & ODD x)"
+  sorry
+
+lemma EVEN_ADD: "EVEN (m + n) = (EVEN m = EVEN n)"
+  sorry
+
+lemma EVEN_MULT: "EVEN (m * n) = (EVEN m | EVEN n)"
+  sorry
+
+lemma ODD_ADD: "ODD (m + n) = (ODD m ~= ODD n)"
+  sorry
+
+lemma ODD_MULT: "ODD (m * n) = (ODD m & ODD n)"
+  sorry
+
+lemma EVEN_DOUBLE: "EVEN (2 * n)"
+  sorry
+
+lemma ODD_DOUBLE: "ODD (Suc (2 * x))"
+  sorry
+
+lemma EVEN_ODD_EXISTS: "(EVEN x --> (EX m. x = 2 * m)) & (ODD x --> (EX m. x = Suc (2 * m)))"
+  sorry
+
+lemma EVEN_EXISTS: "EVEN n = (EX m. n = 2 * m)"
+  sorry
+
+lemma ODD_EXISTS: "ODD n = (EX m. n = Suc (2 * m))"
+  sorry
+
+lemma NOT_SUC_LESS_EQ_0: "~ Suc x <= 0"
+  sorry
+
+lemma NOT_NUM_EQ: "(x ~= xa) = (Suc x <= xa | Suc xa <= x)"
+  sorry
+
+lemma SUC_ADD_SYM: "Suc (m + n) = Suc n + m"
+  sorry
+
+lemma NOT_SUC_ADD_LESS_EQ: "~ Suc (m + n) <= m"
+  sorry
+
+lemma SUB_LEFT_ADD: "(m::nat) + ((n::nat) - (p::nat)) = (if n <= p then m else m + n - p)"
+  sorry
+
+lemma SUB_RIGHT_ADD: "(m::nat) - (n::nat) + (p::nat) = (if m <= n then p else m + p - n)"
+  sorry
+
+lemma SUB_LEFT_SUB: "(m::nat) - ((n::nat) - (p::nat)) = (if n <= p then m else m + p - n)"
+  sorry
+
+lemma SUB_LEFT_SUC: "Suc (m - n) = (if m <= n then Suc 0 else Suc m - n)"
+  sorry
+
+lemma SUB_LEFT_LESS_EQ: "((m::nat) <= (n::nat) - (p::nat)) = (m + p <= n | m <= (0::nat))"
+  sorry
+
+lemma SUB_RIGHT_LESS_EQ: "((m::nat) - (n::nat) <= (p::nat)) = (m <= n + p)"
+  sorry
+
+lemma SUB_RIGHT_LESS: "((m::nat) - (n::nat) < (p::nat)) = (m < n + p & (0::nat) < p)"
+  sorry
+
+lemma SUB_RIGHT_GREATER_EQ: "((p::nat) <= (m::nat) - (n::nat)) = (n + p <= m | p <= (0::nat))"
+  sorry
+
+lemma SUB_LEFT_GREATER: "((n::nat) - (p::nat) < (m::nat)) = (n < m + p & (0::nat) < m)"
+  sorry
+
+lemma SUB_RIGHT_GREATER: "((p::nat) < (m::nat) - (n::nat)) = (n + p < m)"
+  sorry
+
+lemma SUB_LEFT_EQ: "((m::nat) = (n::nat) - (p::nat)) = (m + p = n | m <= (0::nat) & n <= p)"
+  sorry
+
+lemma SUB_RIGHT_EQ: "((m::nat) - (n::nat) = (p::nat)) = (m = n + p | m <= n & p <= (0::nat))"
+  sorry
+
+lemma LE: "(ALL n::nat. (n <= (0::nat)) = (n = (0::nat))) &
 (ALL (m::nat) n::nat. (m <= Suc n) = (m = Suc n | m <= n))"
-  by (import arithmetic LE)
-
-lemma DA: "ALL (k::nat) n::nat. 0 < n --> (EX (x::nat) q::nat. k = q * n + x & x < n)"
-  by (import arithmetic DA)
-
-lemma DIV_LESS_EQ: "ALL n>0. ALL k::nat. k div n <= k"
-  by (import arithmetic DIV_LESS_EQ)
-
-lemma DIV_UNIQUE: "ALL (n::nat) (k::nat) q::nat.
-   (EX r::nat. k = q * n + r & r < n) --> k div n = q"
-  by (import arithmetic DIV_UNIQUE)
-
-lemma MOD_UNIQUE: "ALL (n::nat) (k::nat) r::nat.
-   (EX q::nat. k = q * n + r & r < n) --> k mod n = r"
-  by (import arithmetic MOD_UNIQUE)
-
-lemma DIV_MULT: "ALL (n::nat) r::nat. r < n --> (ALL q::nat. (q * n + r) div n = q)"
-  by (import arithmetic DIV_MULT)
-
-lemma MOD_EQ_0: "ALL n>0. ALL k::nat. k * n mod n = 0"
-  by (import arithmetic MOD_EQ_0)
-
-lemma ZERO_MOD: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
-      ((op =::nat => nat => bool) ((op mod::nat => nat => nat) (0::nat) n)
-        (0::nat)))"
-  by (import arithmetic ZERO_MOD)
-
-lemma ZERO_DIV: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
-      ((op =::nat => nat => bool) ((op div::nat => nat => nat) (0::nat) n)
-        (0::nat)))"
-  by (import arithmetic ZERO_DIV)
-
-lemma MOD_MULT: "ALL (n::nat) r::nat. r < n --> (ALL q::nat. (q * n + r) mod n = r)"
-  by (import arithmetic MOD_MULT)
-
-lemma MOD_TIMES: "ALL n>0. ALL (q::nat) r::nat. (q * n + r) mod n = r mod n"
-  by (import arithmetic MOD_TIMES)
-
-lemma MOD_PLUS: "ALL n>0. ALL (j::nat) k::nat. (j mod n + k mod n) mod n = (j + k) mod n"
-  by (import arithmetic MOD_PLUS)
-
-lemma MOD_MOD: "ALL n>0. ALL k::nat. k mod n mod n = k mod n"
-  by (import arithmetic MOD_MOD)
-
-lemma ADD_DIV_ADD_DIV: "ALL x>0. ALL (xa::nat) r::nat. (xa * x + r) div x = xa + r div x"
-  by (import arithmetic ADD_DIV_ADD_DIV)
-
-lemma MOD_MULT_MOD: "ALL (m::nat) n::nat.
-   0 < n & 0 < m --> (ALL x::nat. x mod (n * m) mod n = x mod n)"
-  by (import arithmetic MOD_MULT_MOD)
-
-lemma DIVMOD_ID: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
-      ((op &::bool => bool => bool)
-        ((op =::nat => nat => bool) ((op div::nat => nat => nat) n n)
-          (1::nat))
-        ((op =::nat => nat => bool) ((op mod::nat => nat => nat) n n)
-          (0::nat))))"
-  by (import arithmetic DIVMOD_ID)
-
-lemma DIV_DIV_DIV_MULT: "ALL (x::nat) xa::nat.
-   0 < x & 0 < xa --> (ALL xb::nat. xb div x div xa = xb div (x * xa))"
-  by (import arithmetic DIV_DIV_DIV_MULT)
-
-lemma DIV_P: "ALL (P::nat => bool) (p::nat) q::nat.
-   0 < q --> P (p div q) = (EX (k::nat) r::nat. p = k * q + r & r < q & P k)"
-  by (import arithmetic DIV_P)
-
-lemma MOD_P: "ALL (P::nat => bool) (p::nat) q::nat.
-   0 < q --> P (p mod q) = (EX (k::nat) r::nat. p = k * q + r & r < q & P r)"
-  by (import arithmetic MOD_P)
-
-lemma MOD_TIMES2: "ALL n>0. ALL (j::nat) k::nat. j mod n * (k mod n) mod n = j * k mod n"
-  by (import arithmetic MOD_TIMES2)
-
-lemma MOD_COMMON_FACTOR: "ALL (n::nat) (p::nat) q::nat.
-   0 < n & 0 < q --> n * (p mod q) = n * p mod (n * q)"
-  by (import arithmetic MOD_COMMON_FACTOR)
-
-lemma num_case_cong: "ALL (M::nat) (M'::nat) (b::'a::type) f::nat => 'a::type.
-   M = M' &
-   (M' = 0 --> b = (b'::'a::type)) &
-   (ALL n::nat. M' = Suc n --> f n = (f'::nat => 'a::type) n) -->
-   nat_case b f M = nat_case b' f' M'"
-  by (import arithmetic num_case_cong)
-
-lemma SUC_ELIM_THM: "ALL P::nat => nat => bool.
-   (ALL n::nat. P (Suc n) n) = (ALL n>0. P n (n - 1))"
-  by (import arithmetic SUC_ELIM_THM)
+  sorry
+
+lemma DA: "(0::nat) < (n::nat) ==> EX (x::nat) q::nat. (k::nat) = q * n + x & x < n"
+  sorry
+
+lemma DIV_LESS_EQ: "(0::nat) < (n::nat) ==> (k::nat) div n <= k"
+  sorry
+
+lemma DIV_UNIQUE: "EX r::nat. (k::nat) = (q::nat) * (n::nat) + r & r < n ==> k div n = q"
+  sorry
+
+lemma MOD_UNIQUE: "EX q::nat. (k::nat) = q * (n::nat) + (r::nat) & r < n ==> k mod n = r"
+  sorry
+
+lemma DIV_MULT: "(r::nat) < (n::nat) ==> ((q::nat) * n + r) div n = q"
+  sorry
+
+lemma MOD_EQ_0: "(0::nat) < (n::nat) ==> (k::nat) * n mod n = (0::nat)"
+  sorry
+
+lemma ZERO_MOD: "(0::nat) < (n::nat) ==> (0::nat) mod n = (0::nat)"
+  sorry
+
+lemma ZERO_DIV: "(0::nat) < (n::nat) ==> (0::nat) div n = (0::nat)"
+  sorry
+
+lemma MOD_MULT: "(r::nat) < (n::nat) ==> ((q::nat) * n + r) mod n = r"
+  sorry
+
+lemma MOD_TIMES: "(0::nat) < (n::nat) ==> ((q::nat) * n + (r::nat)) mod n = r mod n"
+  sorry
+
+lemma MOD_PLUS: "(0::nat) < (n::nat)
+==> ((j::nat) mod n + (k::nat) mod n) mod n = (j + k) mod n"
+  sorry
+
+lemma MOD_MOD: "(0::nat) < (n::nat) ==> (k::nat) mod n mod n = k mod n"
+  sorry
+
+lemma ADD_DIV_ADD_DIV: "(0::nat) < (x::nat) ==> ((xa::nat) * x + (r::nat)) div x = xa + r div x"
+  sorry
+
+lemma MOD_MULT_MOD: "(0::nat) < (n::nat) & (0::nat) < (m::nat)
+==> (x::nat) mod (n * m) mod n = x mod n"
+  sorry
+
+lemma DIVMOD_ID: "(0::nat) < (n::nat) ==> n div n = (1::nat) & n mod n = (0::nat)"
+  sorry
+
+lemma DIV_DIV_DIV_MULT: "(0::nat) < (x::nat) & (0::nat) < (xa::nat)
+==> (xb::nat) div x div xa = xb div (x * xa)"
+  sorry
+
+lemma DIV_P: "(0::nat) < (q::nat)
+==> (P::nat => bool) ((p::nat) div q) =
+    (EX (k::nat) r::nat. p = k * q + r & r < q & P k)"
+  sorry
+
+lemma MOD_P: "(0::nat) < (q::nat)
+==> (P::nat => bool) ((p::nat) mod q) =
+    (EX (k::nat) r::nat. p = k * q + r & r < q & P r)"
+  sorry
+
+lemma MOD_TIMES2: "(0::nat) < (n::nat)
+==> (j::nat) mod n * ((k::nat) mod n) mod n = j * k mod n"
+  sorry
+
+lemma MOD_COMMON_FACTOR: "(0::nat) < (n::nat) & (0::nat) < (q::nat)
+==> n * ((p::nat) mod q) = n * p mod (n * q)"
+  sorry
+
+lemma num_case_cong: "M = M' & (M' = 0 --> b = b') & (ALL n. M' = Suc n --> f n = f' n)
+==> nat_case b f M = nat_case b' f' M'"
+  sorry
+
+lemma SUC_ELIM_THM: "(ALL n. P (Suc n) n) = (ALL n>0. P n (n - 1))"
+  sorry
 
 lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
-(ALL x::nat. (b = a + x --> P 0) & (a = b + x --> P x))"
-  by (import arithmetic SUB_ELIM_THM)
-
-lemma PRE_ELIM_THM: "(P::nat => bool) (PRE (n::nat)) =
-(ALL m::nat. (n = 0 --> P 0) & (n = Suc m --> P m))"
-  by (import arithmetic PRE_ELIM_THM)
-
-lemma MULT_INCREASES: "ALL (m::nat) n::nat. 1 < m & 0 < n --> Suc n <= m * n"
-  by (import arithmetic MULT_INCREASES)
-
-lemma EXP_ALWAYS_BIG_ENOUGH: "ALL b>1. ALL n::nat. EX m::nat. n <= b ^ m"
-  by (import arithmetic EXP_ALWAYS_BIG_ENOUGH)
-
-lemma EXP_EQ_0: "ALL (n::nat) m::nat. (n ^ m = 0) = (n = 0 & 0 < m)"
-  by (import arithmetic EXP_EQ_0)
-
-lemma EXP_1: "(All::(nat => bool) => bool)
- (%x::nat.
-     (op &::bool => bool => bool)
-      ((op =::nat => nat => bool) ((op ^::nat => nat => nat) (1::nat) x)
-        (1::nat))
-      ((op =::nat => nat => bool) ((op ^::nat => nat => nat) x (1::nat)) x))"
-  by (import arithmetic EXP_1)
-
-lemma EXP_EQ_1: "ALL (n::nat) m::nat. (n ^ m = 1) = (n = 1 | m = 0)"
-  by (import arithmetic EXP_EQ_1)
-
-lemma MIN_MAX_EQ: "ALL (x::nat) xa::nat. (min x xa = max x xa) = (x = xa)"
-  by (import arithmetic MIN_MAX_EQ)
-
-lemma MIN_MAX_LT: "ALL (x::nat) xa::nat. (min x xa < max x xa) = (x ~= xa)"
-  by (import arithmetic MIN_MAX_LT)
-
-lemma MIN_MAX_PRED: "ALL (P::nat => bool) (m::nat) n::nat.
-   P m & P n --> P (min m n) & P (max m n)"
-  by (import arithmetic MIN_MAX_PRED)
-
-lemma MIN_LT: "ALL (x::nat) xa::nat.
-   (min xa x < xa) = (xa ~= x & min xa x = x) &
-   (min xa x < x) = (xa ~= x & min xa x = xa) &
-   (xa < min xa x) = False & (x < min xa x) = False"
-  by (import arithmetic MIN_LT)
-
-lemma MAX_LT: "ALL (x::nat) xa::nat.
-   (xa < max xa x) = (xa ~= x & max xa x = x) &
-   (x < max xa x) = (xa ~= x & max xa x = xa) &
-   (max xa x < xa) = False & (max xa x < x) = False"
-  by (import arithmetic MAX_LT)
-
-lemma MIN_LE: "ALL (x::nat) xa::nat. min xa x <= xa & min xa x <= x"
-  by (import arithmetic MIN_LE)
-
-lemma MAX_LE: "ALL (x::nat) xa::nat. xa <= max xa x & x <= max xa x"
-  by (import arithmetic MAX_LE)
-
-lemma MIN_0: "ALL x::nat. min x 0 = 0 & min 0 x = 0"
-  by (import arithmetic MIN_0)
-
-lemma MAX_0: "ALL x::nat. max x 0 = x & max 0 x = x"
-  by (import arithmetic MAX_0)
-
-lemma EXISTS_GREATEST: "ALL P::nat => bool.
-   (Ex P & (EX x::nat. ALL y::nat. x < y --> ~ P y)) =
-   (EX x::nat. P x & (ALL y::nat. x < y --> ~ P y))"
-  by (import arithmetic EXISTS_GREATEST)
+(ALL x::nat. (b = a + x --> P (0::nat)) & (a = b + x --> P x))"
+  sorry
+
+lemma PRE_ELIM_THM: "P (PRE n) = (ALL m. (n = 0 --> P 0) & (n = Suc m --> P m))"
+  sorry
+
+lemma MULT_INCREASES: "1 < m & 0 < n ==> Suc n <= m * n"
+  sorry
+
+lemma EXP_ALWAYS_BIG_ENOUGH: "(1::nat) < (b::nat) ==> EX m::nat. (n::nat) <= b ^ m"
+  sorry
+
+lemma EXP_EQ_0: "((n::nat) ^ (m::nat) = (0::nat)) = (n = (0::nat) & (0::nat) < m)"
+  sorry
+
+lemma EXP_1: "(1::nat) ^ (x::nat) = (1::nat) & x ^ (1::nat) = x"
+  sorry
+
+lemma MIN_MAX_EQ: "(min (x::nat) (xa::nat) = max x xa) = (x = xa)"
+  sorry
+
+lemma MIN_MAX_LT: "(min (x::nat) (xa::nat) < max x xa) = (x ~= xa)"
+  sorry
+
+lemma MIN_MAX_PRED: "(P::nat => bool) (m::nat) & P (n::nat) ==> P (min m n) & P (max m n)"
+  sorry
+
+lemma MIN_LT: "(min (xa::nat) (x::nat) < xa) = (xa ~= x & min xa x = x) &
+(min xa x < x) = (xa ~= x & min xa x = xa) &
+(xa < min xa x) = False & (x < min xa x) = False"
+  sorry
+
+lemma MAX_LT: "((xa::nat) < max xa (x::nat)) = (xa ~= x & max xa x = x) &
+(x < max xa x) = (xa ~= x & max xa x = xa) &
+(max xa x < xa) = False & (max xa x < x) = False"
+  sorry
+
+lemma MIN_LE: "min (xa::nat) (x::nat) <= xa & min xa x <= x"
+  sorry
+
+lemma MAX_LE: "(xa::nat) <= max xa (x::nat) & x <= max xa x"
+  sorry
+
+lemma MIN_0: "min (x::nat) (0::nat) = (0::nat) & min (0::nat) x = (0::nat)"
+  sorry
+
+lemma MAX_0: "max (x::nat) (0::nat) = x & max (0::nat) x = x"
+  sorry
+
+lemma EXISTS_GREATEST: "(Ex (P::nat => bool) & (EX x::nat. ALL y>x. ~ P y)) =
+(EX x::nat. P x & (ALL y>x. ~ P y))"
+  sorry
 
 ;end_setup
 
 ;setup_theory hrat
 
-definition trat_1 :: "nat * nat" where 
+definition
+  trat_1 :: "nat * nat"  where
   "trat_1 == (0, 0)"
 
 lemma trat_1: "trat_1 = (0, 0)"
-  by (import hrat trat_1)
-
-definition trat_inv :: "nat * nat => nat * nat" where 
-  "trat_inv == %(x::nat, y::nat). (y, x)"
-
-lemma trat_inv: "ALL (x::nat) y::nat. trat_inv (x, y) = (y, x)"
-  by (import hrat trat_inv)
-
-definition trat_add :: "nat * nat => nat * nat => nat * nat" where 
+  sorry
+
+definition
+  trat_inv :: "nat * nat => nat * nat"  where
+  "trat_inv == %(x, y). (y, x)"
+
+lemma trat_inv: "trat_inv (x, y) = (y, x)"
+  sorry
+
+definition
+  trat_add :: "nat * nat => nat * nat => nat * nat"  where
   "trat_add ==
-%(x::nat, y::nat) (x'::nat, y'::nat).
+%(x, y) (x', y').
    (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
 
-lemma trat_add: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
-   trat_add (x, y) (x', y') =
-   (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
-  by (import hrat trat_add)
-
-definition trat_mul :: "nat * nat => nat * nat => nat * nat" where 
-  "trat_mul ==
-%(x::nat, y::nat) (x'::nat, y'::nat).
-   (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
-
-lemma trat_mul: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
-   trat_mul (x, y) (x', y') = (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
-  by (import hrat trat_mul)
+lemma trat_add: "trat_add (x, y) (x', y') =
+(PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
+  sorry
+
+definition
+  trat_mul :: "nat * nat => nat * nat => nat * nat"  where
+  "trat_mul == %(x, y) (x', y'). (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
+
+lemma trat_mul: "trat_mul (x, y) (x', y') = (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
+  sorry
 
 consts
   trat_sucint :: "nat => nat * nat" 
 
 specification (trat_sucint) trat_sucint: "trat_sucint 0 = trat_1 &
-(ALL n::nat. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)"
-  by (import hrat trat_sucint)
-
-definition trat_eq :: "nat * nat => nat * nat => bool" where 
-  "trat_eq ==
-%(x::nat, y::nat) (x'::nat, y'::nat). Suc x * Suc y' = Suc x' * Suc y"
-
-lemma trat_eq: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
-   trat_eq (x, y) (x', y') = (Suc x * Suc y' = Suc x' * Suc y)"
-  by (import hrat trat_eq)
-
-lemma TRAT_EQ_REFL: "ALL p::nat * nat. trat_eq p p"
-  by (import hrat TRAT_EQ_REFL)
-
-lemma TRAT_EQ_SYM: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = trat_eq q p"
-  by (import hrat TRAT_EQ_SYM)
-
-lemma TRAT_EQ_TRANS: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
-   trat_eq p q & trat_eq q r --> trat_eq p r"
-  by (import hrat TRAT_EQ_TRANS)
-
-lemma TRAT_EQ_AP: "ALL (p::nat * nat) q::nat * nat. p = q --> trat_eq p q"
-  by (import hrat TRAT_EQ_AP)
-
-lemma TRAT_ADD_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_add h i = trat_add i h"
-  by (import hrat TRAT_ADD_SYM_EQ)
-
-lemma TRAT_MUL_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_mul h i = trat_mul i h"
-  by (import hrat TRAT_MUL_SYM_EQ)
-
-lemma TRAT_INV_WELLDEFINED: "ALL (p::nat * nat) q::nat * nat.
-   trat_eq p q --> trat_eq (trat_inv p) (trat_inv q)"
-  by (import hrat TRAT_INV_WELLDEFINED)
-
-lemma TRAT_ADD_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
-   trat_eq p q --> trat_eq (trat_add p r) (trat_add q r)"
-  by (import hrat TRAT_ADD_WELLDEFINED)
-
-lemma TRAT_ADD_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat.
-   trat_eq p1 p2 & trat_eq q1 q2 -->
-   trat_eq (trat_add p1 q1) (trat_add p2 q2)"
-  by (import hrat TRAT_ADD_WELLDEFINED2)
-
-lemma TRAT_MUL_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
-   trat_eq p q --> trat_eq (trat_mul p r) (trat_mul q r)"
-  by (import hrat TRAT_MUL_WELLDEFINED)
-
-lemma TRAT_MUL_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat.
-   trat_eq p1 p2 & trat_eq q1 q2 -->
-   trat_eq (trat_mul p1 q1) (trat_mul p2 q2)"
-  by (import hrat TRAT_MUL_WELLDEFINED2)
-
-lemma TRAT_ADD_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_add h i) (trat_add i h)"
-  by (import hrat TRAT_ADD_SYM)
-
-lemma TRAT_ADD_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
-   trat_eq (trat_add h (trat_add i j)) (trat_add (trat_add h i) j)"
-  by (import hrat TRAT_ADD_ASSOC)
-
-lemma TRAT_MUL_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_mul h i) (trat_mul i h)"
-  by (import hrat TRAT_MUL_SYM)
-
-lemma TRAT_MUL_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
-   trat_eq (trat_mul h (trat_mul i j)) (trat_mul (trat_mul h i) j)"
-  by (import hrat TRAT_MUL_ASSOC)
-
-lemma TRAT_LDISTRIB: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
-   trat_eq (trat_mul h (trat_add i j))
-    (trat_add (trat_mul h i) (trat_mul h j))"
-  by (import hrat TRAT_LDISTRIB)
-
-lemma TRAT_MUL_LID: "ALL h::nat * nat. trat_eq (trat_mul trat_1 h) h"
-  by (import hrat TRAT_MUL_LID)
-
-lemma TRAT_MUL_LINV: "ALL h::nat * nat. trat_eq (trat_mul (trat_inv h) h) trat_1"
-  by (import hrat TRAT_MUL_LINV)
-
-lemma TRAT_NOZERO: "ALL (h::nat * nat) i::nat * nat. ~ trat_eq (trat_add h i) h"
-  by (import hrat TRAT_NOZERO)
-
-lemma TRAT_ADD_TOTAL: "ALL (h::nat * nat) i::nat * nat.
-   trat_eq h i |
-   (EX d::nat * nat. trat_eq h (trat_add i d)) |
-   (EX d::nat * nat. trat_eq i (trat_add h d))"
-  by (import hrat TRAT_ADD_TOTAL)
-
-lemma TRAT_SUCINT_0: "ALL n::nat. trat_eq (trat_sucint n) (n, 0)"
-  by (import hrat TRAT_SUCINT_0)
-
-lemma TRAT_ARCH: "ALL h::nat * nat.
-   EX (n::nat) d::nat * nat. trat_eq (trat_sucint n) (trat_add h d)"
-  by (import hrat TRAT_ARCH)
+(ALL n. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)"
+  sorry
+
+definition
+  trat_eq :: "nat * nat => nat * nat => bool"  where
+  "trat_eq == %(x, y) (x', y'). Suc x * Suc y' = Suc x' * Suc y"
+
+lemma trat_eq: "trat_eq (x, y) (x', y') = (Suc x * Suc y' = Suc x' * Suc y)"
+  sorry
+
+lemma TRAT_EQ_REFL: "trat_eq p p"
+  sorry
+
+lemma TRAT_EQ_SYM: "trat_eq p q = trat_eq q p"
+  sorry
+
+lemma TRAT_EQ_TRANS: "trat_eq p q & trat_eq q r ==> trat_eq p r"
+  sorry
+
+lemma TRAT_EQ_AP: "p = q ==> trat_eq p q"
+  sorry
+
+lemma TRAT_ADD_SYM_EQ: "trat_add h i = trat_add i h"
+  sorry
+
+lemma TRAT_MUL_SYM_EQ: "trat_mul h i = trat_mul i h"
+  sorry
+
+lemma TRAT_INV_WELLDEFINED: "trat_eq p q ==> trat_eq (trat_inv p) (trat_inv q)"
+  sorry
+
+lemma TRAT_ADD_WELLDEFINED: "trat_eq p q ==> trat_eq (trat_add p r) (trat_add q r)"
+  sorry
+
+lemma TRAT_ADD_WELLDEFINED2: "trat_eq p1 p2 & trat_eq q1 q2 ==> trat_eq (trat_add p1 q1) (trat_add p2 q2)"
+  sorry
+
+lemma TRAT_MUL_WELLDEFINED: "trat_eq p q ==> trat_eq (trat_mul p r) (trat_mul q r)"
+  sorry
+
+lemma TRAT_MUL_WELLDEFINED2: "trat_eq p1 p2 & trat_eq q1 q2 ==> trat_eq (trat_mul p1 q1) (trat_mul p2 q2)"
+  sorry
+
+lemma TRAT_ADD_SYM: "trat_eq (trat_add h i) (trat_add i h)"
+  sorry
+
+lemma TRAT_ADD_ASSOC: "trat_eq (trat_add h (trat_add i j)) (trat_add (trat_add h i) j)"
+  sorry
+
+lemma TRAT_MUL_SYM: "trat_eq (trat_mul h i) (trat_mul i h)"
+  sorry
+
+lemma TRAT_MUL_ASSOC: "trat_eq (trat_mul h (trat_mul i j)) (trat_mul (trat_mul h i) j)"
+  sorry
+
+lemma TRAT_LDISTRIB: "trat_eq (trat_mul h (trat_add i j)) (trat_add (trat_mul h i) (trat_mul h j))"
+  sorry
+
+lemma TRAT_MUL_LID: "trat_eq (trat_mul trat_1 h) h"
+  sorry
+
+lemma TRAT_MUL_LINV: "trat_eq (trat_mul (trat_inv h) h) trat_1"
+  sorry
+
+lemma TRAT_NOZERO: "~ trat_eq (trat_add h i) h"
+  sorry
+
+lemma TRAT_ADD_TOTAL: "trat_eq h i |
+(EX d. trat_eq h (trat_add i d)) | (EX d. trat_eq i (trat_add h d))"
+  sorry
+
+lemma TRAT_SUCINT_0: "trat_eq (trat_sucint n) (n, 0)"
+  sorry
+
+lemma TRAT_ARCH: "EX n d. trat_eq (trat_sucint n) (trat_add h d)"
+  sorry
 
 lemma TRAT_SUCINT: "trat_eq (trat_sucint 0) trat_1 &
-(ALL n::nat.
-    trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))"
-  by (import hrat TRAT_SUCINT)
-
-lemma TRAT_EQ_EQUIV: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = (trat_eq p = trat_eq q)"
-  by (import hrat TRAT_EQ_EQUIV)
-
-typedef (open) hrat = "{x::nat * nat => bool. EX xa::nat * nat. x = trat_eq xa}" 
-  by (rule typedef_helper,import hrat hrat_TY_DEF)
+(ALL n. trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))"
+  sorry
+
+lemma TRAT_EQ_EQUIV: "trat_eq p q = (trat_eq p = trat_eq q)"
+  sorry
+
+typedef (open) hrat = "{x. EX xa. x = trat_eq xa}" 
+  sorry
 
 lemmas hrat_TY_DEF = typedef_hol2hol4 [OF type_definition_hrat]
 
@@ -1866,227 +1477,213 @@
   mk_hrat :: "(nat * nat => bool) => hrat" 
   dest_hrat :: "hrat => nat * nat => bool" 
 
-specification (dest_hrat mk_hrat) hrat_tybij: "(ALL a::hrat. mk_hrat (dest_hrat a) = a) &
-(ALL r::nat * nat => bool.
-    (EX x::nat * nat. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))"
-  by (import hrat hrat_tybij)
-
-definition hrat_1 :: "hrat" where 
+specification (dest_hrat mk_hrat) hrat_tybij: "(ALL a. mk_hrat (dest_hrat a) = a) &
+(ALL r. (EX x. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))"
+  sorry
+
+definition
+  hrat_1 :: "hrat"  where
   "hrat_1 == mk_hrat (trat_eq trat_1)"
 
 lemma hrat_1: "hrat_1 = mk_hrat (trat_eq trat_1)"
-  by (import hrat hrat_1)
-
-definition hrat_inv :: "hrat => hrat" where 
-  "hrat_inv == %T1::hrat. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
-
-lemma hrat_inv: "ALL T1::hrat.
-   hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
-  by (import hrat hrat_inv)
-
-definition hrat_add :: "hrat => hrat => hrat" where 
+  sorry
+
+definition
+  hrat_inv :: "hrat => hrat"  where
+  "hrat_inv == %T1. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
+
+lemma hrat_inv: "hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
+  sorry
+
+definition
+  hrat_add :: "hrat => hrat => hrat"  where
   "hrat_add ==
-%(T1::hrat) T2::hrat.
+%T1 T2.
    mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
 
-lemma hrat_add: "ALL (T1::hrat) T2::hrat.
-   hrat_add T1 T2 =
-   mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
-  by (import hrat hrat_add)
-
-definition hrat_mul :: "hrat => hrat => hrat" where 
+lemma hrat_add: "hrat_add T1 T2 =
+mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+  sorry
+
+definition
+  hrat_mul :: "hrat => hrat => hrat"  where
   "hrat_mul ==
-%(T1::hrat) T2::hrat.
-   mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
-
-lemma hrat_mul: "ALL (T1::hrat) T2::hrat.
-   hrat_mul T1 T2 =
+%T1 T2.
    mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
-  by (import hrat hrat_mul)
-
-definition hrat_sucint :: "nat => hrat" where 
-  "hrat_sucint == %T1::nat. mk_hrat (trat_eq (trat_sucint T1))"
-
-lemma hrat_sucint: "ALL T1::nat. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))"
-  by (import hrat hrat_sucint)
-
-lemma HRAT_ADD_SYM: "ALL (h::hrat) i::hrat. hrat_add h i = hrat_add i h"
-  by (import hrat HRAT_ADD_SYM)
-
-lemma HRAT_ADD_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat.
-   hrat_add h (hrat_add i j) = hrat_add (hrat_add h i) j"
-  by (import hrat HRAT_ADD_ASSOC)
-
-lemma HRAT_MUL_SYM: "ALL (h::hrat) i::hrat. hrat_mul h i = hrat_mul i h"
-  by (import hrat HRAT_MUL_SYM)
-
-lemma HRAT_MUL_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat.
-   hrat_mul h (hrat_mul i j) = hrat_mul (hrat_mul h i) j"
-  by (import hrat HRAT_MUL_ASSOC)
-
-lemma HRAT_LDISTRIB: "ALL (h::hrat) (i::hrat) j::hrat.
-   hrat_mul h (hrat_add i j) = hrat_add (hrat_mul h i) (hrat_mul h j)"
-  by (import hrat HRAT_LDISTRIB)
-
-lemma HRAT_MUL_LID: "ALL h::hrat. hrat_mul hrat_1 h = h"
-  by (import hrat HRAT_MUL_LID)
-
-lemma HRAT_MUL_LINV: "ALL h::hrat. hrat_mul (hrat_inv h) h = hrat_1"
-  by (import hrat HRAT_MUL_LINV)
-
-lemma HRAT_NOZERO: "ALL (h::hrat) i::hrat. hrat_add h i ~= h"
-  by (import hrat HRAT_NOZERO)
-
-lemma HRAT_ADD_TOTAL: "ALL (h::hrat) i::hrat.
-   h = i | (EX x::hrat. h = hrat_add i x) | (EX x::hrat. i = hrat_add h x)"
-  by (import hrat HRAT_ADD_TOTAL)
-
-lemma HRAT_ARCH: "ALL h::hrat. EX (x::nat) xa::hrat. hrat_sucint x = hrat_add h xa"
-  by (import hrat HRAT_ARCH)
+
+lemma hrat_mul: "hrat_mul T1 T2 =
+mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+  sorry
+
+definition
+  hrat_sucint :: "nat => hrat"  where
+  "hrat_sucint == %T1. mk_hrat (trat_eq (trat_sucint T1))"
+
+lemma hrat_sucint: "hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))"
+  sorry
+
+lemma HRAT_ADD_SYM: "hrat_add h i = hrat_add i h"
+  sorry
+
+lemma HRAT_ADD_ASSOC: "hrat_add h (hrat_add i j) = hrat_add (hrat_add h i) j"
+  sorry
+
+lemma HRAT_MUL_SYM: "hrat_mul h i = hrat_mul i h"
+  sorry
+
+lemma HRAT_MUL_ASSOC: "hrat_mul h (hrat_mul i j) = hrat_mul (hrat_mul h i) j"
+  sorry
+
+lemma HRAT_LDISTRIB: "hrat_mul h (hrat_add i j) = hrat_add (hrat_mul h i) (hrat_mul h j)"
+  sorry
+
+lemma HRAT_MUL_LID: "hrat_mul hrat_1 h = h"
+  sorry
+
+lemma HRAT_MUL_LINV: "hrat_mul (hrat_inv h) h = hrat_1"
+  sorry
+
+lemma HRAT_NOZERO: "hrat_add h i ~= h"
+  sorry
+
+lemma HRAT_ADD_TOTAL: "h = i | (EX x. h = hrat_add i x) | (EX x. i = hrat_add h x)"
+  sorry
+
+lemma HRAT_ARCH: "EX x xa. hrat_sucint x = hrat_add h xa"
+  sorry
 
 lemma HRAT_SUCINT: "hrat_sucint 0 = hrat_1 &
-(ALL x::nat. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)"
-  by (import hrat HRAT_SUCINT)
+(ALL x. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)"
+  sorry
 
 ;end_setup
 
 ;setup_theory hreal
 
-definition hrat_lt :: "hrat => hrat => bool" where 
-  "hrat_lt == %(x::hrat) y::hrat. EX d::hrat. y = hrat_add x d"
-
-lemma hrat_lt: "ALL (x::hrat) y::hrat. hrat_lt x y = (EX d::hrat. y = hrat_add x d)"
-  by (import hreal hrat_lt)
-
-lemma HRAT_LT_REFL: "ALL x::hrat. ~ hrat_lt x x"
-  by (import hreal HRAT_LT_REFL)
-
-lemma HRAT_LT_TRANS: "ALL (x::hrat) (y::hrat) z::hrat. hrat_lt x y & hrat_lt y z --> hrat_lt x z"
-  by (import hreal HRAT_LT_TRANS)
-
-lemma HRAT_LT_ANTISYM: "ALL (x::hrat) y::hrat. ~ (hrat_lt x y & hrat_lt y x)"
-  by (import hreal HRAT_LT_ANTISYM)
-
-lemma HRAT_LT_TOTAL: "ALL (x::hrat) y::hrat. x = y | hrat_lt x y | hrat_lt y x"
-  by (import hreal HRAT_LT_TOTAL)
-
-lemma HRAT_MUL_RID: "ALL x::hrat. hrat_mul x hrat_1 = x"
-  by (import hreal HRAT_MUL_RID)
-
-lemma HRAT_MUL_RINV: "ALL x::hrat. hrat_mul x (hrat_inv x) = hrat_1"
-  by (import hreal HRAT_MUL_RINV)
-
-lemma HRAT_RDISTRIB: "ALL (x::hrat) (y::hrat) z::hrat.
-   hrat_mul (hrat_add x y) z = hrat_add (hrat_mul x z) (hrat_mul y z)"
-  by (import hreal HRAT_RDISTRIB)
-
-lemma HRAT_LT_ADDL: "ALL (x::hrat) y::hrat. hrat_lt x (hrat_add x y)"
-  by (import hreal HRAT_LT_ADDL)
-
-lemma HRAT_LT_ADDR: "ALL (x::hrat) xa::hrat. hrat_lt xa (hrat_add x xa)"
-  by (import hreal HRAT_LT_ADDR)
-
-lemma HRAT_LT_GT: "ALL (x::hrat) y::hrat. hrat_lt x y --> ~ hrat_lt y x"
-  by (import hreal HRAT_LT_GT)
-
-lemma HRAT_LT_NE: "ALL (x::hrat) y::hrat. hrat_lt x y --> x ~= y"
-  by (import hreal HRAT_LT_NE)
-
-lemma HRAT_EQ_LADD: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_add x y = hrat_add x z) = (y = z)"
-  by (import hreal HRAT_EQ_LADD)
-
-lemma HRAT_EQ_LMUL: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_mul x y = hrat_mul x z) = (y = z)"
-  by (import hreal HRAT_EQ_LMUL)
-
-lemma HRAT_LT_ADD2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat.
-   hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_add u v) (hrat_add x y)"
-  by (import hreal HRAT_LT_ADD2)
-
-lemma HRAT_LT_LADD: "ALL (x::hrat) (y::hrat) z::hrat.
-   hrat_lt (hrat_add z x) (hrat_add z y) = hrat_lt x y"
-  by (import hreal HRAT_LT_LADD)
-
-lemma HRAT_LT_RADD: "ALL (x::hrat) (y::hrat) z::hrat.
-   hrat_lt (hrat_add x z) (hrat_add y z) = hrat_lt x y"
-  by (import hreal HRAT_LT_RADD)
-
-lemma HRAT_LT_MUL2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat.
-   hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_mul u v) (hrat_mul x y)"
-  by (import hreal HRAT_LT_MUL2)
-
-lemma HRAT_LT_LMUL: "ALL (x::hrat) (y::hrat) z::hrat.
-   hrat_lt (hrat_mul z x) (hrat_mul z y) = hrat_lt x y"
-  by (import hreal HRAT_LT_LMUL)
-
-lemma HRAT_LT_RMUL: "ALL (x::hrat) (y::hrat) z::hrat.
-   hrat_lt (hrat_mul x z) (hrat_mul y z) = hrat_lt x y"
-  by (import hreal HRAT_LT_RMUL)
-
-lemma HRAT_LT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1"
-  by (import hreal HRAT_LT_LMUL1)
-
-lemma HRAT_LT_RMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1"
-  by (import hreal HRAT_LT_RMUL1)
-
-lemma HRAT_GT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x"
-  by (import hreal HRAT_GT_LMUL1)
-
-lemma HRAT_LT_L1: "ALL (x::hrat) y::hrat.
-   hrat_lt (hrat_mul (hrat_inv x) y) hrat_1 = hrat_lt y x"
-  by (import hreal HRAT_LT_L1)
-
-lemma HRAT_LT_R1: "ALL (x::hrat) y::hrat.
-   hrat_lt (hrat_mul x (hrat_inv y)) hrat_1 = hrat_lt x y"
-  by (import hreal HRAT_LT_R1)
-
-lemma HRAT_GT_L1: "ALL (x::hrat) y::hrat.
-   hrat_lt hrat_1 (hrat_mul (hrat_inv x) y) = hrat_lt x y"
-  by (import hreal HRAT_GT_L1)
-
-lemma HRAT_INV_MUL: "ALL (x::hrat) y::hrat.
-   hrat_inv (hrat_mul x y) = hrat_mul (hrat_inv x) (hrat_inv y)"
-  by (import hreal HRAT_INV_MUL)
-
-lemma HRAT_UP: "ALL x::hrat. Ex (hrat_lt x)"
-  by (import hreal HRAT_UP)
-
-lemma HRAT_DOWN: "ALL x::hrat. EX xa::hrat. hrat_lt xa x"
-  by (import hreal HRAT_DOWN)
-
-lemma HRAT_DOWN2: "ALL (x::hrat) y::hrat. EX xa::hrat. hrat_lt xa x & hrat_lt xa y"
-  by (import hreal HRAT_DOWN2)
-
-lemma HRAT_MEAN: "ALL (x::hrat) y::hrat.
-   hrat_lt x y --> (EX xa::hrat. hrat_lt x xa & hrat_lt xa y)"
-  by (import hreal HRAT_MEAN)
-
-definition isacut :: "(hrat => bool) => bool" where 
+definition
+  hrat_lt :: "hrat => hrat => bool"  where
+  "hrat_lt == %x y. EX d. y = hrat_add x d"
+
+lemma hrat_lt: "hrat_lt x y = (EX d. y = hrat_add x d)"
+  sorry
+
+lemma HRAT_LT_REFL: "~ hrat_lt x x"
+  sorry
+
+lemma HRAT_LT_TRANS: "hrat_lt x y & hrat_lt y z ==> hrat_lt x z"
+  sorry
+
+lemma HRAT_LT_ANTISYM: "~ (hrat_lt x y & hrat_lt y x)"
+  sorry
+
+lemma HRAT_LT_TOTAL: "x = y | hrat_lt x y | hrat_lt y x"
+  sorry
+
+lemma HRAT_MUL_RID: "hrat_mul x hrat_1 = x"
+  sorry
+
+lemma HRAT_MUL_RINV: "hrat_mul x (hrat_inv x) = hrat_1"
+  sorry
+
+lemma HRAT_RDISTRIB: "hrat_mul (hrat_add x y) z = hrat_add (hrat_mul x z) (hrat_mul y z)"
+  sorry
+
+lemma HRAT_LT_ADDL: "hrat_lt x (hrat_add x y)"
+  sorry
+
+lemma HRAT_LT_ADDR: "hrat_lt xa (hrat_add x xa)"
+  sorry
+
+lemma HRAT_LT_GT: "hrat_lt x y ==> ~ hrat_lt y x"
+  sorry
+
+lemma HRAT_LT_NE: "hrat_lt x y ==> x ~= y"
+  sorry
+
+lemma HRAT_EQ_LADD: "(hrat_add x y = hrat_add x z) = (y = z)"
+  sorry
+
+lemma HRAT_EQ_LMUL: "(hrat_mul x y = hrat_mul x z) = (y = z)"
+  sorry
+
+lemma HRAT_LT_ADD2: "hrat_lt u x & hrat_lt v y ==> hrat_lt (hrat_add u v) (hrat_add x y)"
+  sorry
+
+lemma HRAT_LT_LADD: "hrat_lt (hrat_add z x) (hrat_add z y) = hrat_lt x y"
+  sorry
+
+lemma HRAT_LT_RADD: "hrat_lt (hrat_add x z) (hrat_add y z) = hrat_lt x y"
+  sorry
+
+lemma HRAT_LT_MUL2: "hrat_lt u x & hrat_lt v y ==> hrat_lt (hrat_mul u v) (hrat_mul x y)"
+  sorry
+
+lemma HRAT_LT_LMUL: "hrat_lt (hrat_mul z x) (hrat_mul z y) = hrat_lt x y"
+  sorry
+
+lemma HRAT_LT_RMUL: "hrat_lt (hrat_mul x z) (hrat_mul y z) = hrat_lt x y"
+  sorry
+
+lemma HRAT_LT_LMUL1: "hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1"
+  sorry
+
+lemma HRAT_LT_RMUL1: "hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1"
+  sorry
+
+lemma HRAT_GT_LMUL1: "hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x"
+  sorry
+
+lemma HRAT_LT_L1: "hrat_lt (hrat_mul (hrat_inv x) y) hrat_1 = hrat_lt y x"
+  sorry
+
+lemma HRAT_LT_R1: "hrat_lt (hrat_mul x (hrat_inv y)) hrat_1 = hrat_lt x y"
+  sorry
+
+lemma HRAT_GT_L1: "hrat_lt hrat_1 (hrat_mul (hrat_inv x) y) = hrat_lt x y"
+  sorry
+
+lemma HRAT_INV_MUL: "hrat_inv (hrat_mul x y) = hrat_mul (hrat_inv x) (hrat_inv y)"
+  sorry
+
+lemma HRAT_UP: "Ex (hrat_lt x)"
+  sorry
+
+lemma HRAT_DOWN: "EX xa. hrat_lt xa x"
+  sorry
+
+lemma HRAT_DOWN2: "EX xa. hrat_lt xa x & hrat_lt xa y"
+  sorry
+
+lemma HRAT_MEAN: "hrat_lt x y ==> EX xa. hrat_lt x xa & hrat_lt xa y"
+  sorry
+
+definition
+  isacut :: "(hrat => bool) => bool"  where
   "isacut ==
-%C::hrat => bool.
-   Ex C &
-   (EX x::hrat. ~ C x) &
-   (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) &
-   (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y))"
-
-lemma isacut: "ALL C::hrat => bool.
-   isacut C =
-   (Ex C &
-    (EX x::hrat. ~ C x) &
-    (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) &
-    (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y)))"
-  by (import hreal isacut)
-
-definition cut_of_hrat :: "hrat => hrat => bool" where 
-  "cut_of_hrat == %(x::hrat) y::hrat. hrat_lt y x"
-
-lemma cut_of_hrat: "ALL x::hrat. cut_of_hrat x = (%y::hrat. hrat_lt y x)"
-  by (import hreal cut_of_hrat)
-
-lemma ISACUT_HRAT: "ALL h::hrat. isacut (cut_of_hrat h)"
-  by (import hreal ISACUT_HRAT)
+%C. Ex C &
+    (EX x. ~ C x) &
+    (ALL x y. C x & hrat_lt y x --> C y) &
+    (ALL x. C x --> (EX y. C y & hrat_lt x y))"
+
+lemma isacut: "isacut (CC::hrat => bool) =
+(Ex CC &
+ (EX x::hrat. ~ CC x) &
+ (ALL (x::hrat) y::hrat. CC x & hrat_lt y x --> CC y) &
+ (ALL x::hrat. CC x --> (EX y::hrat. CC y & hrat_lt x y)))"
+  sorry
+
+definition
+  cut_of_hrat :: "hrat => hrat => bool"  where
+  "cut_of_hrat == %x y. hrat_lt y x"
+
+lemma cut_of_hrat: "cut_of_hrat x = (%y. hrat_lt y x)"
+  sorry
+
+lemma ISACUT_HRAT: "isacut (cut_of_hrat h)"
+  sorry
 
 typedef (open) hreal = "Collect isacut" 
-  by (rule typedef_helper,import hreal hreal_TY_DEF)
+  sorry
 
 lemmas hreal_TY_DEF = typedef_hol2hol4 [OF type_definition_hreal]
 
@@ -2094,795 +1691,506 @@
   hreal :: "(hrat => bool) => hreal" 
   cut :: "hreal => hrat => bool" 
 
-specification (cut hreal) hreal_tybij: "(ALL a::hreal. hreal (hreal.cut a) = a) &
-(ALL r::hrat => bool. isacut r = (hreal.cut (hreal r) = r))"
-  by (import hreal hreal_tybij)
-
-lemma EQUAL_CUTS: "ALL (X::hreal) Y::hreal. hreal.cut X = hreal.cut Y --> X = Y"
-  by (import hreal EQUAL_CUTS)
-
-lemma CUT_ISACUT: "ALL x::hreal. isacut (hreal.cut x)"
-  by (import hreal CUT_ISACUT)
-
-lemma CUT_NONEMPTY: "ALL x::hreal. Ex (hreal.cut x)"
-  by (import hreal CUT_NONEMPTY)
-
-lemma CUT_BOUNDED: "ALL x::hreal. EX xa::hrat. ~ hreal.cut x xa"
-  by (import hreal CUT_BOUNDED)
-
-lemma CUT_DOWN: "ALL (x::hreal) (xa::hrat) xb::hrat.
-   hreal.cut x xa & hrat_lt xb xa --> hreal.cut x xb"
-  by (import hreal CUT_DOWN)
-
-lemma CUT_UP: "ALL (x::hreal) xa::hrat.
-   hreal.cut x xa --> (EX y::hrat. hreal.cut x y & hrat_lt xa y)"
-  by (import hreal CUT_UP)
-
-lemma CUT_UBOUND: "ALL (x::hreal) (xa::hrat) xb::hrat.
-   ~ hreal.cut x xa & hrat_lt xa xb --> ~ hreal.cut x xb"
-  by (import hreal CUT_UBOUND)
-
-lemma CUT_STRADDLE: "ALL (X::hreal) (x::hrat) y::hrat.
-   hreal.cut X x & ~ hreal.cut X y --> hrat_lt x y"
-  by (import hreal CUT_STRADDLE)
-
-lemma CUT_NEARTOP_ADD: "ALL (X::hreal) e::hrat.
-   EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_add x e)"
-  by (import hreal CUT_NEARTOP_ADD)
-
-lemma CUT_NEARTOP_MUL: "ALL (X::hreal) u::hrat.
-   hrat_lt hrat_1 u -->
-   (EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))"
-  by (import hreal CUT_NEARTOP_MUL)
-
-definition hreal_1 :: "hreal" where 
+specification (cut hreal) hreal_tybij: "(ALL a. hreal (cut a) = a) & (ALL r. isacut r = (cut (hreal r) = r))"
+  sorry
+
+lemma EQUAL_CUTS: "cut X = cut Y ==> X = Y"
+  sorry
+
+lemma CUT_ISACUT: "isacut (cut x)"
+  sorry
+
+lemma CUT_NONEMPTY: "Ex (cut x)"
+  sorry
+
+lemma CUT_BOUNDED: "EX xa. ~ cut x xa"
+  sorry
+
+lemma CUT_DOWN: "cut x xa & hrat_lt xb xa ==> cut x xb"
+  sorry
+
+lemma CUT_UP: "cut x xa ==> EX y. cut x y & hrat_lt xa y"
+  sorry
+
+lemma CUT_UBOUND: "~ cut x xa & hrat_lt xa xb ==> ~ cut x xb"
+  sorry
+
+lemma CUT_STRADDLE: "cut X x & ~ cut X y ==> hrat_lt x y"
+  sorry
+
+lemma CUT_NEARTOP_ADD: "EX x. cut X x & ~ cut X (hrat_add x e)"
+  sorry
+
+lemma CUT_NEARTOP_MUL: "hrat_lt hrat_1 u ==> EX x. cut X x & ~ cut X (hrat_mul u x)"
+  sorry
+
+definition
+  hreal_1 :: "hreal"  where
   "hreal_1 == hreal (cut_of_hrat hrat_1)"
 
 lemma hreal_1: "hreal_1 = hreal (cut_of_hrat hrat_1)"
-  by (import hreal hreal_1)
-
-definition hreal_add :: "hreal => hreal => hreal" where 
-  "hreal_add ==
-%(X::hreal) Y::hreal.
-   hreal
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
-
-lemma hreal_add: "ALL (X::hreal) Y::hreal.
-   hreal_add X Y =
-   hreal
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
-  by (import hreal hreal_add)
-
-definition hreal_mul :: "hreal => hreal => hreal" where 
-  "hreal_mul ==
-%(X::hreal) Y::hreal.
-   hreal
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
-
-lemma hreal_mul: "ALL (X::hreal) Y::hreal.
-   hreal_mul X Y =
-   hreal
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
-  by (import hreal hreal_mul)
-
-definition hreal_inv :: "hreal => hreal" where 
+  sorry
+
+definition
+  hreal_add :: "hreal => hreal => hreal"  where
+  "hreal_add == %X Y. hreal (%w. EX x y. w = hrat_add x y & cut X x & cut Y y)"
+
+lemma hreal_add: "hreal_add X Y = hreal (%w. EX x y. w = hrat_add x y & cut X x & cut Y y)"
+  sorry
+
+definition
+  hreal_mul :: "hreal => hreal => hreal"  where
+  "hreal_mul == %X Y. hreal (%w. EX x y. w = hrat_mul x y & cut X x & cut Y y)"
+
+lemma hreal_mul: "hreal_mul X Y = hreal (%w. EX x y. w = hrat_mul x y & cut X x & cut Y y)"
+  sorry
+
+definition
+  hreal_inv :: "hreal => hreal"  where
   "hreal_inv ==
-%X::hreal.
-   hreal
-    (%w::hrat.
-        EX d::hrat.
-           hrat_lt d hrat_1 &
-           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
-
-lemma hreal_inv: "ALL X::hreal.
-   hreal_inv X =
-   hreal
-    (%w::hrat.
-        EX d::hrat.
-           hrat_lt d hrat_1 &
-           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
-  by (import hreal hreal_inv)
-
-definition hreal_sup :: "(hreal => bool) => hreal" where 
-  "hreal_sup ==
-%P::hreal => bool. hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
-
-lemma hreal_sup: "ALL P::hreal => bool.
-   hreal_sup P = hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
-  by (import hreal hreal_sup)
-
-definition hreal_lt :: "hreal => hreal => bool" where 
-  "hreal_lt ==
-%(X::hreal) Y::hreal.
-   X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x)"
-
-lemma hreal_lt: "ALL (X::hreal) Y::hreal.
-   hreal_lt X Y = (X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x))"
-  by (import hreal hreal_lt)
-
-lemma HREAL_INV_ISACUT: "ALL X::hreal.
-   isacut
-    (%w::hrat.
-        EX d::hrat.
-           hrat_lt d hrat_1 &
-           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
-  by (import hreal HREAL_INV_ISACUT)
-
-lemma HREAL_ADD_ISACUT: "ALL (X::hreal) Y::hreal.
-   isacut
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
-  by (import hreal HREAL_ADD_ISACUT)
-
-lemma HREAL_MUL_ISACUT: "ALL (X::hreal) Y::hreal.
-   isacut
-    (%w::hrat.
-        EX (x::hrat) y::hrat.
-           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
-  by (import hreal HREAL_MUL_ISACUT)
-
-lemma HREAL_ADD_SYM: "ALL (X::hreal) Y::hreal. hreal_add X Y = hreal_add Y X"
-  by (import hreal HREAL_ADD_SYM)
-
-lemma HREAL_MUL_SYM: "ALL (X::hreal) Y::hreal. hreal_mul X Y = hreal_mul Y X"
-  by (import hreal HREAL_MUL_SYM)
-
-lemma HREAL_ADD_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal.
-   hreal_add X (hreal_add Y Z) = hreal_add (hreal_add X Y) Z"
-  by (import hreal HREAL_ADD_ASSOC)
-
-lemma HREAL_MUL_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal.
-   hreal_mul X (hreal_mul Y Z) = hreal_mul (hreal_mul X Y) Z"
-  by (import hreal HREAL_MUL_ASSOC)
-
-lemma HREAL_LDISTRIB: "ALL (X::hreal) (Y::hreal) Z::hreal.
-   hreal_mul X (hreal_add Y Z) = hreal_add (hreal_mul X Y) (hreal_mul X Z)"
-  by (import hreal HREAL_LDISTRIB)
-
-lemma HREAL_MUL_LID: "ALL X::hreal. hreal_mul hreal_1 X = X"
-  by (import hreal HREAL_MUL_LID)
-
-lemma HREAL_MUL_LINV: "ALL X::hreal. hreal_mul (hreal_inv X) X = hreal_1"
-  by (import hreal HREAL_MUL_LINV)
-
-lemma HREAL_NOZERO: "ALL (X::hreal) Y::hreal. hreal_add X Y ~= X"
-  by (import hreal HREAL_NOZERO)
-
-definition hreal_sub :: "hreal => hreal => hreal" where 
-  "hreal_sub ==
-%(Y::hreal) X::hreal.
-   hreal
-    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
-
-lemma hreal_sub: "ALL (Y::hreal) X::hreal.
-   hreal_sub Y X =
-   hreal
-    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
-  by (import hreal hreal_sub)
-
-lemma HREAL_LT_LEMMA: "ALL (X::hreal) Y::hreal.
-   hreal_lt X Y --> (EX x::hrat. ~ hreal.cut X x & hreal.cut Y x)"
-  by (import hreal HREAL_LT_LEMMA)
-
-lemma HREAL_SUB_ISACUT: "ALL (X::hreal) Y::hreal.
-   hreal_lt X Y -->
-   isacut
-    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
-  by (import hreal HREAL_SUB_ISACUT)
-
-lemma HREAL_SUB_ADD: "ALL (X::hreal) Y::hreal. hreal_lt X Y --> hreal_add (hreal_sub Y X) X = Y"
-  by (import hreal HREAL_SUB_ADD)
-
-lemma HREAL_LT_TOTAL: "ALL (X::hreal) Y::hreal. X = Y | hreal_lt X Y | hreal_lt Y X"
-  by (import hreal HREAL_LT_TOTAL)
-
-lemma HREAL_LT: "ALL (X::hreal) Y::hreal. hreal_lt X Y = (EX D::hreal. Y = hreal_add X D)"
-  by (import hreal HREAL_LT)
-
-lemma HREAL_ADD_TOTAL: "ALL (X::hreal) Y::hreal.
-   X = Y |
-   (EX D::hreal. Y = hreal_add X D) | (EX D::hreal. X = hreal_add Y D)"
-  by (import hreal HREAL_ADD_TOTAL)
-
-lemma HREAL_SUP_ISACUT: "ALL P::hreal => bool.
-   Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) -->
-   isacut (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
-  by (import hreal HREAL_SUP_ISACUT)
-
-lemma HREAL_SUP: "ALL P::hreal => bool.
-   Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) -->
-   (ALL Y::hreal.
-       (EX X::hreal. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P))"
-  by (import hreal HREAL_SUP)
+%X. hreal
+     (%w. EX d. hrat_lt d hrat_1 &
+                (ALL x. cut X x --> hrat_lt (hrat_mul w x) d))"
+
+lemma hreal_inv: "hreal_inv X =
+hreal
+ (%w. EX d. hrat_lt d hrat_1 &
+            (ALL x. cut X x --> hrat_lt (hrat_mul w x) d))"
+  sorry
+
+definition
+  hreal_sup :: "(hreal => bool) => hreal"  where
+  "hreal_sup == %P. hreal (%w. EX X. P X & cut X w)"
+
+lemma hreal_sup: "hreal_sup P = hreal (%w. EX X. P X & cut X w)"
+  sorry
+
+definition
+  hreal_lt :: "hreal => hreal => bool"  where
+  "hreal_lt == %X Y. X ~= Y & (ALL x. cut X x --> cut Y x)"
+
+lemma hreal_lt: "hreal_lt X Y = (X ~= Y & (ALL x. cut X x --> cut Y x))"
+  sorry
+
+lemma HREAL_INV_ISACUT: "isacut
+ (%w. EX d. hrat_lt d hrat_1 &
+            (ALL x. cut X x --> hrat_lt (hrat_mul w x) d))"
+  sorry
+
+lemma HREAL_ADD_ISACUT: "isacut (%w. EX x y. w = hrat_add x y & cut X x & cut Y y)"
+  sorry
+
+lemma HREAL_MUL_ISACUT: "isacut (%w. EX x y. w = hrat_mul x y & cut X x & cut Y y)"
+  sorry
+
+lemma HREAL_ADD_SYM: "hreal_add X Y = hreal_add Y X"
+  sorry
+
+lemma HREAL_MUL_SYM: "hreal_mul X Y = hreal_mul Y X"
+  sorry
+
+lemma HREAL_ADD_ASSOC: "hreal_add X (hreal_add Y Z) = hreal_add (hreal_add X Y) Z"
+  sorry
+
+lemma HREAL_MUL_ASSOC: "hreal_mul X (hreal_mul Y Z) = hreal_mul (hreal_mul X Y) Z"
+  sorry
+
+lemma HREAL_LDISTRIB: "hreal_mul X (hreal_add Y Z) = hreal_add (hreal_mul X Y) (hreal_mul X Z)"
+  sorry
+
+lemma HREAL_MUL_LID: "hreal_mul hreal_1 X = X"
+  sorry
+
+lemma HREAL_MUL_LINV: "hreal_mul (hreal_inv X) X = hreal_1"
+  sorry
+
+lemma HREAL_NOZERO: "hreal_add X Y ~= X"
+  sorry
+
+definition
+  hreal_sub :: "hreal => hreal => hreal"  where
+  "hreal_sub == %Y X. hreal (%w. EX x. ~ cut X x & cut Y (hrat_add x w))"
+
+lemma hreal_sub: "hreal_sub Y X = hreal (%w. EX x. ~ cut X x & cut Y (hrat_add x w))"
+  sorry
+
+lemma HREAL_LT_LEMMA: "hreal_lt X Y ==> EX x. ~ cut X x & cut Y x"
+  sorry
+
+lemma HREAL_SUB_ISACUT: "hreal_lt X Y ==> isacut (%w. EX x. ~ cut X x & cut Y (hrat_add x w))"
+  sorry
+
+lemma HREAL_SUB_ADD: "hreal_lt X Y ==> hreal_add (hreal_sub Y X) X = Y"
+  sorry
+
+lemma HREAL_LT_TOTAL: "X = Y | hreal_lt X Y | hreal_lt Y X"
+  sorry
+
+lemma HREAL_LT: "hreal_lt X Y = (EX D. Y = hreal_add X D)"
+  sorry
+
+lemma HREAL_ADD_TOTAL: "X = Y | (EX D. Y = hreal_add X D) | (EX D. X = hreal_add Y D)"
+  sorry
+
+lemma HREAL_SUP_ISACUT: "Ex P & (EX Y. ALL X. P X --> hreal_lt X Y)
+==> isacut (%w. EX X. P X & cut X w)"
+  sorry
+
+lemma HREAL_SUP: "Ex P & (EX Y. ALL X. P X --> hreal_lt X Y)
+==> (EX X. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P)"
+  sorry
 
 ;end_setup
 
 ;setup_theory numeral
 
 lemma numeral_suc: "Suc ALT_ZERO = NUMERAL_BIT1 ALT_ZERO &
-(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) &
-(ALL x::nat. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))"
-  by (import numeral numeral_suc)
-
-definition iZ :: "nat => nat" where 
-  "iZ == %x::nat. x"
-
-lemma iZ: "ALL x::nat. iZ x = x"
-  by (import numeral iZ)
-
-definition iiSUC :: "nat => nat" where 
-  "iiSUC == %n::nat. Suc (Suc n)"
-
-lemma iiSUC: "ALL n::nat. iiSUC n = Suc (Suc n)"
-  by (import numeral iiSUC)
-
-lemma numeral_distrib: "(op &::bool => bool => bool)
- ((All::(nat => bool) => bool)
-   (%x::nat.
-       (op =::nat => nat => bool) ((op +::nat => nat => nat) (0::nat) x) x))
- ((op &::bool => bool => bool)
-   ((All::(nat => bool) => bool)
-     (%x::nat.
-         (op =::nat => nat => bool) ((op +::nat => nat => nat) x (0::nat))
-          x))
-   ((op &::bool => bool => bool)
-     ((All::(nat => bool) => bool)
-       (%x::nat.
-           (All::(nat => bool) => bool)
-            (%xa::nat.
-                (op =::nat => nat => bool)
-                 ((op +::nat => nat => nat) ((NUMERAL::nat => nat) x)
-                   ((NUMERAL::nat => nat) xa))
-                 ((NUMERAL::nat => nat)
-                   ((iZ::nat => nat) ((op +::nat => nat => nat) x xa))))))
-     ((op &::bool => bool => bool)
-       ((All::(nat => bool) => bool)
-         (%x::nat.
-             (op =::nat => nat => bool)
-              ((op *::nat => nat => nat) (0::nat) x) (0::nat)))
-       ((op &::bool => bool => bool)
-         ((All::(nat => bool) => bool)
-           (%x::nat.
-               (op =::nat => nat => bool)
-                ((op *::nat => nat => nat) x (0::nat)) (0::nat)))
-         ((op &::bool => bool => bool)
-           ((All::(nat => bool) => bool)
-             (%x::nat.
-                 (All::(nat => bool) => bool)
-                  (%xa::nat.
-                      (op =::nat => nat => bool)
-                       ((op *::nat => nat => nat) ((NUMERAL::nat => nat) x)
-                         ((NUMERAL::nat => nat) xa))
-                       ((NUMERAL::nat => nat)
-                         ((op *::nat => nat => nat) x xa)))))
-           ((op &::bool => bool => bool)
-             ((All::(nat => bool) => bool)
-               (%x::nat.
-                   (op =::nat => nat => bool)
-                    ((op -::nat => nat => nat) (0::nat) x) (0::nat)))
-             ((op &::bool => bool => bool)
-               ((All::(nat => bool) => bool)
-                 (%x::nat.
-                     (op =::nat => nat => bool)
-                      ((op -::nat => nat => nat) x (0::nat)) x))
-               ((op &::bool => bool => bool)
-                 ((All::(nat => bool) => bool)
-                   (%x::nat.
-                       (All::(nat => bool) => bool)
-                        (%xa::nat.
-                            (op =::nat => nat => bool)
-                             ((op -::nat => nat => nat)
-                               ((NUMERAL::nat => nat) x)
-                               ((NUMERAL::nat => nat) xa))
-                             ((NUMERAL::nat => nat)
-                               ((op -::nat => nat => nat) x xa)))))
-                 ((op &::bool => bool => bool)
-                   ((All::(nat => bool) => bool)
-                     (%x::nat.
-                         (op =::nat => nat => bool)
-                          ((op ^::nat => nat => nat) (0::nat)
-                            ((NUMERAL::nat => nat)
-                              ((NUMERAL_BIT1::nat => nat) x)))
-                          (0::nat)))
-                   ((op &::bool => bool => bool)
-                     ((All::(nat => bool) => bool)
-                       (%x::nat.
-                           (op =::nat => nat => bool)
-                            ((op ^::nat => nat => nat) (0::nat)
-                              ((NUMERAL::nat => nat)
-                                ((NUMERAL_BIT2::nat => nat) x)))
-                            (0::nat)))
-                     ((op &::bool => bool => bool)
-                       ((All::(nat => bool) => bool)
-                         (%x::nat.
-                             (op =::nat => nat => bool)
-                              ((op ^::nat => nat => nat) x (0::nat))
-                              (1::nat)))
-                       ((op &::bool => bool => bool)
-                         ((All::(nat => bool) => bool)
-                           (%x::nat.
-                               (All::(nat => bool) => bool)
-                                (%xa::nat.
-                                    (op =::nat => nat => bool)
-                                     ((op ^::nat => nat => nat)
- ((NUMERAL::nat => nat) x) ((NUMERAL::nat => nat) xa))
-                                     ((NUMERAL::nat => nat)
- ((op ^::nat => nat => nat) x xa)))))
-                         ((op &::bool => bool => bool)
-                           ((op =::nat => nat => bool)
-                             ((Suc::nat => nat) (0::nat)) (1::nat))
-                           ((op &::bool => bool => bool)
-                             ((All::(nat => bool) => bool)
-                               (%x::nat.
-                                   (op =::nat => nat => bool)
-                                    ((Suc::nat => nat)
-((NUMERAL::nat => nat) x))
-                                    ((NUMERAL::nat => nat)
-((Suc::nat => nat) x))))
-                             ((op &::bool => bool => bool)
-                               ((op =::nat => nat => bool)
-                                 ((PRE::nat => nat) (0::nat)) (0::nat))
-                               ((op &::bool => bool => bool)
-                                 ((All::(nat => bool) => bool)
-                                   (%x::nat.
- (op =::nat => nat => bool) ((PRE::nat => nat) ((NUMERAL::nat => nat) x))
-  ((NUMERAL::nat => nat) ((PRE::nat => nat) x))))
-                                 ((op &::bool => bool => bool)
-                                   ((All::(nat => bool) => bool)
-                                     (%x::nat.
-   (op =::bool => bool => bool)
-    ((op =::nat => nat => bool) ((NUMERAL::nat => nat) x) (0::nat))
-    ((op =::nat => nat => bool) x (ALT_ZERO::nat))))
-                                   ((op &::bool => bool => bool)
-                                     ((All::(nat => bool) => bool)
- (%x::nat.
-     (op =::bool => bool => bool)
-      ((op =::nat => nat => bool) (0::nat) ((NUMERAL::nat => nat) x))
-      ((op =::nat => nat => bool) x (ALT_ZERO::nat))))
-                                     ((op &::bool => bool => bool)
- ((All::(nat => bool) => bool)
-   (%x::nat.
-       (All::(nat => bool) => bool)
-        (%xa::nat.
-            (op =::bool => bool => bool)
-             ((op =::nat => nat => bool) ((NUMERAL::nat => nat) x)
-               ((NUMERAL::nat => nat) xa))
-             ((op =::nat => nat => bool) x xa))))
- ((op &::bool => bool => bool)
-   ((All::(nat => bool) => bool)
-     (%x::nat.
-         (op =::bool => bool => bool)
-          ((op <::nat => nat => bool) x (0::nat)) (False::bool)))
-   ((op &::bool => bool => bool)
-     ((All::(nat => bool) => bool)
-       (%x::nat.
-           (op =::bool => bool => bool)
-            ((op <::nat => nat => bool) (0::nat) ((NUMERAL::nat => nat) x))
-            ((op <::nat => nat => bool) (ALT_ZERO::nat) x)))
-     ((op &::bool => bool => bool)
-       ((All::(nat => bool) => bool)
-         (%x::nat.
-             (All::(nat => bool) => bool)
-              (%xa::nat.
-                  (op =::bool => bool => bool)
-                   ((op <::nat => nat => bool) ((NUMERAL::nat => nat) x)
-                     ((NUMERAL::nat => nat) xa))
-                   ((op <::nat => nat => bool) x xa))))
-       ((op &::bool => bool => bool)
-         ((All::(nat => bool) => bool)
-           (%x::nat.
-               (op =::bool => bool => bool)
-                ((op <::nat => nat => bool) x (0::nat)) (False::bool)))
-         ((op &::bool => bool => bool)
-           ((All::(nat => bool) => bool)
-             (%x::nat.
-                 (op =::bool => bool => bool)
-                  ((op <::nat => nat => bool) (0::nat)
-                    ((NUMERAL::nat => nat) x))
-                  ((op <::nat => nat => bool) (ALT_ZERO::nat) x)))
-           ((op &::bool => bool => bool)
-             ((All::(nat => bool) => bool)
-               (%x::nat.
-                   (All::(nat => bool) => bool)
-                    (%xa::nat.
-                        (op =::bool => bool => bool)
-                         ((op <::nat => nat => bool)
-                           ((NUMERAL::nat => nat) xa)
-                           ((NUMERAL::nat => nat) x))
-                         ((op <::nat => nat => bool) xa x))))
-             ((op &::bool => bool => bool)
-               ((All::(nat => bool) => bool)
-                 (%x::nat.
-                     (op =::bool => bool => bool)
-                      ((op <=::nat => nat => bool) (0::nat) x)
-                      (True::bool)))
-               ((op &::bool => bool => bool)
-                 ((All::(nat => bool) => bool)
-                   (%x::nat.
-                       (op =::bool => bool => bool)
-                        ((op <=::nat => nat => bool)
-                          ((NUMERAL::nat => nat) x) (0::nat))
-                        ((op <=::nat => nat => bool) x (ALT_ZERO::nat))))
-                 ((op &::bool => bool => bool)
-                   ((All::(nat => bool) => bool)
-                     (%x::nat.
-                         (All::(nat => bool) => bool)
-                          (%xa::nat.
-                              (op =::bool => bool => bool)
-                               ((op <=::nat => nat => bool)
-                                 ((NUMERAL::nat => nat) x)
-                                 ((NUMERAL::nat => nat) xa))
-                               ((op <=::nat => nat => bool) x xa))))
-                   ((op &::bool => bool => bool)
-                     ((All::(nat => bool) => bool)
-                       (%x::nat.
-                           (op =::bool => bool => bool)
-                            ((op <=::nat => nat => bool) (0::nat) x)
-                            (True::bool)))
-                     ((op &::bool => bool => bool)
-                       ((All::(nat => bool) => bool)
-                         (%x::nat.
-                             (op =::bool => bool => bool)
-                              ((op <=::nat => nat => bool) x (0::nat))
-                              ((op =::nat => nat => bool) x (0::nat))))
-                       ((op &::bool => bool => bool)
-                         ((All::(nat => bool) => bool)
-                           (%x::nat.
-                               (All::(nat => bool) => bool)
-                                (%xa::nat.
-                                    (op =::bool => bool => bool)
-                                     ((op <=::nat => nat => bool)
- ((NUMERAL::nat => nat) xa) ((NUMERAL::nat => nat) x))
-                                     ((op <=::nat => nat => bool) xa x))))
-                         ((op &::bool => bool => bool)
-                           ((All::(nat => bool) => bool)
-                             (%x::nat.
-                                 (op =::bool => bool => bool)
-                                  ((ODD::nat => bool)
-                                    ((NUMERAL::nat => nat) x))
-                                  ((ODD::nat => bool) x)))
-                           ((op &::bool => bool => bool)
-                             ((All::(nat => bool) => bool)
-                               (%x::nat.
-                                   (op =::bool => bool => bool)
-                                    ((EVEN::nat => bool)
-((NUMERAL::nat => nat) x))
-                                    ((EVEN::nat => bool) x)))
-                             ((op &::bool => bool => bool)
-                               ((Not::bool => bool)
-                                 ((ODD::nat => bool) (0::nat)))
-                               ((EVEN::nat => bool)
-                                 (0::nat))))))))))))))))))))))))))))))))))))"
-  by (import numeral numeral_distrib)
+(ALL x. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) &
+(ALL x. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))"
+  sorry
+
+definition
+  iZ :: "nat => nat"  where
+  "iZ == %x. x"
+
+lemma iZ: "iZ x = x"
+  sorry
+
+definition
+  iiSUC :: "nat => nat"  where
+  "iiSUC == %n. Suc (Suc n)"
+
+lemma iiSUC: "iiSUC n = Suc (Suc n)"
+  sorry
+
+lemma numeral_distrib: "(ALL x::nat. (0::nat) + x = x) &
+(ALL x::nat. x + (0::nat) = x) &
+(ALL (x::nat) xa::nat. NUMERAL x + NUMERAL xa = NUMERAL (iZ (x + xa))) &
+(ALL x::nat. (0::nat) * x = (0::nat)) &
+(ALL x::nat. x * (0::nat) = (0::nat)) &
+(ALL (x::nat) xa::nat. NUMERAL x * NUMERAL xa = NUMERAL (x * xa)) &
+(ALL x::nat. (0::nat) - x = (0::nat)) &
+(ALL x::nat. x - (0::nat) = x) &
+(ALL (x::nat) xa::nat. NUMERAL x - NUMERAL xa = NUMERAL (x - xa)) &
+(ALL x::nat. (0::nat) ^ NUMERAL (NUMERAL_BIT1 x) = (0::nat)) &
+(ALL x::nat. (0::nat) ^ NUMERAL (NUMERAL_BIT2 x) = (0::nat)) &
+(ALL x::nat. x ^ (0::nat) = (1::nat)) &
+(ALL (x::nat) xa::nat. NUMERAL x ^ NUMERAL xa = NUMERAL (x ^ xa)) &
+Suc (0::nat) = (1::nat) &
+(ALL x::nat. Suc (NUMERAL x) = NUMERAL (Suc x)) &
+PRE (0::nat) = (0::nat) &
+(ALL x::nat. PRE (NUMERAL x) = NUMERAL (PRE x)) &
+(ALL x::nat. (NUMERAL x = (0::nat)) = (x = ALT_ZERO)) &
+(ALL x::nat. ((0::nat) = NUMERAL x) = (x = ALT_ZERO)) &
+(ALL (x::nat) xa::nat. (NUMERAL x = NUMERAL xa) = (x = xa)) &
+(ALL x::nat. (x < (0::nat)) = False) &
+(ALL x::nat. ((0::nat) < NUMERAL x) = (ALT_ZERO < x)) &
+(ALL (x::nat) xa::nat. (NUMERAL x < NUMERAL xa) = (x < xa)) &
+(ALL x::nat. (x < (0::nat)) = False) &
+(ALL x::nat. ((0::nat) < NUMERAL x) = (ALT_ZERO < x)) &
+(ALL (x::nat) xa::nat. (NUMERAL xa < NUMERAL x) = (xa < x)) &
+(ALL x::nat. ((0::nat) <= x) = True) &
+(ALL x::nat. (NUMERAL x <= (0::nat)) = (x <= ALT_ZERO)) &
+(ALL (x::nat) xa::nat. (NUMERAL x <= NUMERAL xa) = (x <= xa)) &
+(ALL x::nat. ((0::nat) <= x) = True) &
+(ALL x::nat. (x <= (0::nat)) = (x = (0::nat))) &
+(ALL (x::nat) xa::nat. (NUMERAL xa <= NUMERAL x) = (xa <= x)) &
+(ALL x::nat. ODD (NUMERAL x) = ODD x) &
+(ALL x::nat. EVEN (NUMERAL x) = EVEN x) & ~ ODD (0::nat) & EVEN (0::nat)"
+  sorry
 
 lemma numeral_iisuc: "iiSUC ALT_ZERO = NUMERAL_BIT2 ALT_ZERO &
-iiSUC (NUMERAL_BIT1 (n::nat)) = NUMERAL_BIT1 (Suc n) &
+iiSUC (NUMERAL_BIT1 n) = NUMERAL_BIT1 (Suc n) &
 iiSUC (NUMERAL_BIT2 n) = NUMERAL_BIT2 (Suc n)"
-  by (import numeral numeral_iisuc)
-
-lemma numeral_add: "ALL (x::nat) xa::nat.
-   iZ (ALT_ZERO + x) = x &
-   iZ (x + ALT_ZERO) = x &
-   iZ (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (iZ (x + xa)) &
-   iZ (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
-   iZ (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
-   iZ (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
-   Suc (ALT_ZERO + x) = Suc x &
-   Suc (x + ALT_ZERO) = Suc x &
-   Suc (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
-   Suc (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
-   Suc (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
-   Suc (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (iiSUC (x + xa)) &
-   iiSUC (ALT_ZERO + x) = iiSUC x &
-   iiSUC (x + ALT_ZERO) = iiSUC x &
-   iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
-   iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) =
-   NUMERAL_BIT1 (iiSUC (x + xa)) &
-   iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) =
-   NUMERAL_BIT1 (iiSUC (x + xa)) &
-   iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (iiSUC (x + xa))"
-  by (import numeral numeral_add)
-
-lemma numeral_eq: "ALL (x::nat) xa::nat.
-   (ALT_ZERO = NUMERAL_BIT1 x) = False &
-   (NUMERAL_BIT1 x = ALT_ZERO) = False &
-   (ALT_ZERO = NUMERAL_BIT2 x) = False &
-   (NUMERAL_BIT2 x = ALT_ZERO) = False &
-   (NUMERAL_BIT1 x = NUMERAL_BIT2 xa) = False &
-   (NUMERAL_BIT2 x = NUMERAL_BIT1 xa) = False &
-   (NUMERAL_BIT1 x = NUMERAL_BIT1 xa) = (x = xa) &
-   (NUMERAL_BIT2 x = NUMERAL_BIT2 xa) = (x = xa)"
-  by (import numeral numeral_eq)
-
-lemma numeral_lt: "ALL (x::nat) xa::nat.
-   (ALT_ZERO < NUMERAL_BIT1 x) = True &
-   (ALT_ZERO < NUMERAL_BIT2 x) = True &
-   (x < ALT_ZERO) = False &
-   (NUMERAL_BIT1 x < NUMERAL_BIT1 xa) = (x < xa) &
-   (NUMERAL_BIT2 x < NUMERAL_BIT2 xa) = (x < xa) &
-   (NUMERAL_BIT1 x < NUMERAL_BIT2 xa) = (~ xa < x) &
-   (NUMERAL_BIT2 x < NUMERAL_BIT1 xa) = (x < xa)"
-  by (import numeral numeral_lt)
-
-lemma numeral_lte: "ALL (x::nat) xa::nat.
-   (ALT_ZERO <= x) = True &
-   (NUMERAL_BIT1 x <= ALT_ZERO) = False &
-   (NUMERAL_BIT2 x <= ALT_ZERO) = False &
-   (NUMERAL_BIT1 x <= NUMERAL_BIT1 xa) = (x <= xa) &
-   (NUMERAL_BIT1 x <= NUMERAL_BIT2 xa) = (x <= xa) &
-   (NUMERAL_BIT2 x <= NUMERAL_BIT1 xa) = (~ xa <= x) &
-   (NUMERAL_BIT2 x <= NUMERAL_BIT2 xa) = (x <= xa)"
-  by (import numeral numeral_lte)
+  sorry
+
+lemma numeral_add: "iZ (ALT_ZERO + x) = x &
+iZ (x + ALT_ZERO) = x &
+iZ (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (iZ (x + xa)) &
+iZ (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
+iZ (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
+iZ (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
+Suc (ALT_ZERO + x) = Suc x &
+Suc (x + ALT_ZERO) = Suc x &
+Suc (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
+Suc (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
+Suc (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
+Suc (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (iiSUC (x + xa)) &
+iiSUC (ALT_ZERO + x) = iiSUC x &
+iiSUC (x + ALT_ZERO) = iiSUC x &
+iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
+iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (iiSUC (x + xa)) &
+iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (iiSUC (x + xa)) &
+iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (iiSUC (x + xa))"
+  sorry
+
+lemma numeral_eq: "(ALT_ZERO = NUMERAL_BIT1 x) = False &
+(NUMERAL_BIT1 x = ALT_ZERO) = False &
+(ALT_ZERO = NUMERAL_BIT2 x) = False &
+(NUMERAL_BIT2 x = ALT_ZERO) = False &
+(NUMERAL_BIT1 x = NUMERAL_BIT2 xa) = False &
+(NUMERAL_BIT2 x = NUMERAL_BIT1 xa) = False &
+(NUMERAL_BIT1 x = NUMERAL_BIT1 xa) = (x = xa) &
+(NUMERAL_BIT2 x = NUMERAL_BIT2 xa) = (x = xa)"
+  sorry
+
+lemma numeral_lt: "(ALT_ZERO < NUMERAL_BIT1 x) = True &
+(ALT_ZERO < NUMERAL_BIT2 x) = True &
+(x < ALT_ZERO) = False &
+(NUMERAL_BIT1 x < NUMERAL_BIT1 xa) = (x < xa) &
+(NUMERAL_BIT2 x < NUMERAL_BIT2 xa) = (x < xa) &
+(NUMERAL_BIT1 x < NUMERAL_BIT2 xa) = (~ xa < x) &
+(NUMERAL_BIT2 x < NUMERAL_BIT1 xa) = (x < xa)"
+  sorry
+
+lemma numeral_lte: "(ALT_ZERO <= x) = True &
+(NUMERAL_BIT1 x <= ALT_ZERO) = False &
+(NUMERAL_BIT2 x <= ALT_ZERO) = False &
+(NUMERAL_BIT1 x <= NUMERAL_BIT1 xa) = (x <= xa) &
+(NUMERAL_BIT1 x <= NUMERAL_BIT2 xa) = (x <= xa) &
+(NUMERAL_BIT2 x <= NUMERAL_BIT1 xa) = (~ xa <= x) &
+(NUMERAL_BIT2 x <= NUMERAL_BIT2 xa) = (x <= xa)"
+  sorry
 
 lemma numeral_pre: "PRE ALT_ZERO = ALT_ZERO &
 PRE (NUMERAL_BIT1 ALT_ZERO) = ALT_ZERO &
-(ALL x::nat.
+(ALL x.
     PRE (NUMERAL_BIT1 (NUMERAL_BIT1 x)) =
     NUMERAL_BIT2 (PRE (NUMERAL_BIT1 x))) &
-(ALL x::nat.
+(ALL x.
     PRE (NUMERAL_BIT1 (NUMERAL_BIT2 x)) = NUMERAL_BIT2 (NUMERAL_BIT1 x)) &
-(ALL x::nat. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)"
-  by (import numeral numeral_pre)
-
-lemma bit_initiality: "ALL (zf::'a::type) (b1f::nat => 'a::type => 'a::type)
-   b2f::nat => 'a::type => 'a::type.
-   EX x::nat => 'a::type.
-      x ALT_ZERO = zf &
-      (ALL n::nat. x (NUMERAL_BIT1 n) = b1f n (x n)) &
-      (ALL n::nat. x (NUMERAL_BIT2 n) = b2f n (x n))"
-  by (import numeral bit_initiality)
+(ALL x. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)"
+  sorry
+
+lemma bit_initiality: "EX x. x ALT_ZERO = zf &
+      (ALL n. x (NUMERAL_BIT1 n) = b1f n (x n)) &
+      (ALL n. x (NUMERAL_BIT2 n) = b2f n (x n))"
+  sorry
 
 consts
   iBIT_cases :: "nat => 'a => (nat => 'a) => (nat => 'a) => 'a" 
 
-specification (iBIT_cases) iBIT_cases: "(ALL (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
+specification (iBIT_cases) iBIT_cases: "(ALL (zf::'a) (bf1::nat => 'a) bf2::nat => 'a.
     iBIT_cases ALT_ZERO zf bf1 bf2 = zf) &
-(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
+(ALL (n::nat) (zf::'a) (bf1::nat => 'a) bf2::nat => 'a.
     iBIT_cases (NUMERAL_BIT1 n) zf bf1 bf2 = bf1 n) &
-(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
+(ALL (n::nat) (zf::'a) (bf1::nat => 'a) bf2::nat => 'a.
     iBIT_cases (NUMERAL_BIT2 n) zf bf1 bf2 = bf2 n)"
-  by (import numeral iBIT_cases)
-
-definition iDUB :: "nat => nat" where 
-  "iDUB == %x::nat. x + x"
-
-lemma iDUB: "ALL x::nat. iDUB x = x + x"
-  by (import numeral iDUB)
+  sorry
+
+definition
+  iDUB :: "nat => nat"  where
+  "iDUB == %x. x + x"
+
+lemma iDUB: "iDUB x = x + x"
+  sorry
 
 consts
   iSUB :: "bool => nat => nat => nat" 
 
-specification (iSUB) iSUB_DEF: "(ALL (b::bool) x::nat. iSUB b ALT_ZERO x = ALT_ZERO) &
-(ALL (b::bool) (n::nat) x::nat.
+specification (iSUB) iSUB_DEF: "(ALL b x. iSUB b ALT_ZERO x = ALT_ZERO) &
+(ALL b n x.
     iSUB b (NUMERAL_BIT1 n) x =
     (if b
-     then iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m))
-           (%m::nat. NUMERAL_BIT1 (iSUB False n m))
-     else iBIT_cases x (iDUB n) (%m::nat. NUMERAL_BIT1 (iSUB False n m))
-           (%m::nat. iDUB (iSUB False n m)))) &
-(ALL (b::bool) (n::nat) x::nat.
+     then iBIT_cases x (NUMERAL_BIT1 n) (%m. iDUB (iSUB True n m))
+           (%m. NUMERAL_BIT1 (iSUB False n m))
+     else iBIT_cases x (iDUB n) (%m. NUMERAL_BIT1 (iSUB False n m))
+           (%m. iDUB (iSUB False n m)))) &
+(ALL b n x.
     iSUB b (NUMERAL_BIT2 n) x =
     (if b
-     then iBIT_cases x (NUMERAL_BIT2 n)
-           (%m::nat. NUMERAL_BIT1 (iSUB True n m))
-           (%m::nat. iDUB (iSUB True n m))
-     else iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m))
-           (%m::nat. NUMERAL_BIT1 (iSUB False n m))))"
-  by (import numeral iSUB_DEF)
-
-lemma bit_induction: "ALL P::nat => bool.
-   P ALT_ZERO &
-   (ALL n::nat. P n --> P (NUMERAL_BIT1 n)) &
-   (ALL n::nat. P n --> P (NUMERAL_BIT2 n)) -->
-   All P"
-  by (import numeral bit_induction)
-
-lemma iSUB_THM: "ALL (xa::bool) (xb::nat) xc::nat.
-   iSUB xa ALT_ZERO (x::nat) = ALT_ZERO &
-   iSUB True xb ALT_ZERO = xb &
-   iSUB False (NUMERAL_BIT1 xb) ALT_ZERO = iDUB xb &
-   iSUB True (NUMERAL_BIT1 xb) (NUMERAL_BIT1 xc) = iDUB (iSUB True xb xc) &
-   iSUB False (NUMERAL_BIT1 xb) (NUMERAL_BIT1 xc) =
-   NUMERAL_BIT1 (iSUB False xb xc) &
-   iSUB True (NUMERAL_BIT1 xb) (NUMERAL_BIT2 xc) =
-   NUMERAL_BIT1 (iSUB False xb xc) &
-   iSUB False (NUMERAL_BIT1 xb) (NUMERAL_BIT2 xc) =
-   iDUB (iSUB False xb xc) &
-   iSUB False (NUMERAL_BIT2 xb) ALT_ZERO = NUMERAL_BIT1 xb &
-   iSUB True (NUMERAL_BIT2 xb) (NUMERAL_BIT1 xc) =
-   NUMERAL_BIT1 (iSUB True xb xc) &
-   iSUB False (NUMERAL_BIT2 xb) (NUMERAL_BIT1 xc) = iDUB (iSUB True xb xc) &
-   iSUB True (NUMERAL_BIT2 xb) (NUMERAL_BIT2 xc) = iDUB (iSUB True xb xc) &
-   iSUB False (NUMERAL_BIT2 xb) (NUMERAL_BIT2 xc) =
-   NUMERAL_BIT1 (iSUB False xb xc)"
-  by (import numeral iSUB_THM)
-
-lemma numeral_sub: "ALL (x::nat) xa::nat.
-   NUMERAL (x - xa) = (if xa < x then NUMERAL (iSUB True x xa) else 0)"
-  by (import numeral numeral_sub)
-
-lemma iDUB_removal: "ALL x::nat.
-   iDUB (NUMERAL_BIT1 x) = NUMERAL_BIT2 (iDUB x) &
-   iDUB (NUMERAL_BIT2 x) = NUMERAL_BIT2 (NUMERAL_BIT1 x) &
-   iDUB ALT_ZERO = ALT_ZERO"
-  by (import numeral iDUB_removal)
-
-lemma numeral_mult: "ALL (x::nat) xa::nat.
-   ALT_ZERO * x = ALT_ZERO &
-   x * ALT_ZERO = ALT_ZERO &
-   NUMERAL_BIT1 x * xa = iZ (iDUB (x * xa) + xa) &
-   NUMERAL_BIT2 x * xa = iDUB (iZ (x * xa + xa))"
-  by (import numeral numeral_mult)
-
-definition iSQR :: "nat => nat" where 
-  "iSQR == %x::nat. x * x"
-
-lemma iSQR: "ALL x::nat. iSQR x = x * x"
-  by (import numeral iSQR)
-
-lemma numeral_exp: "(ALL x::nat. x ^ ALT_ZERO = NUMERAL_BIT1 ALT_ZERO) &
-(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT1 xa = x * iSQR (x ^ xa)) &
-(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT2 xa = iSQR x * iSQR (x ^ xa))"
-  by (import numeral numeral_exp)
-
-lemma numeral_evenodd: "ALL x::nat.
-   EVEN ALT_ZERO &
-   EVEN (NUMERAL_BIT2 x) &
-   ~ EVEN (NUMERAL_BIT1 x) &
-   ~ ODD ALT_ZERO & ~ ODD (NUMERAL_BIT2 x) & ODD (NUMERAL_BIT1 x)"
-  by (import numeral numeral_evenodd)
-
-lemma numeral_fact: "ALL n::nat. FACT n = (if n = 0 then 1 else n * FACT (PRE n))"
-  by (import numeral numeral_fact)
-
-lemma numeral_funpow: "ALL n::nat.
-   ((f::'a::type => 'a::type) ^^ n) (x::'a::type) =
-   (if n = 0 then x else (f ^^ (n - 1)) (f x))"
-  by (import numeral numeral_funpow)
+     then iBIT_cases x (NUMERAL_BIT2 n) (%m. NUMERAL_BIT1 (iSUB True n m))
+           (%m. iDUB (iSUB True n m))
+     else iBIT_cases x (NUMERAL_BIT1 n) (%m. iDUB (iSUB True n m))
+           (%m. NUMERAL_BIT1 (iSUB False n m))))"
+  sorry
+
+lemma bit_induction: "P ALT_ZERO &
+(ALL n. P n --> P (NUMERAL_BIT1 n)) & (ALL n. P n --> P (NUMERAL_BIT2 n))
+==> P x"
+  sorry
+
+lemma iSUB_THM: "iSUB (x::bool) ALT_ZERO (xn::nat) = ALT_ZERO &
+iSUB True (xa::nat) ALT_ZERO = xa &
+iSUB False (NUMERAL_BIT1 xa) ALT_ZERO = iDUB xa &
+iSUB True (NUMERAL_BIT1 xa) (NUMERAL_BIT1 (xb::nat)) =
+iDUB (iSUB True xa xb) &
+iSUB False (NUMERAL_BIT1 xa) (NUMERAL_BIT1 xb) =
+NUMERAL_BIT1 (iSUB False xa xb) &
+iSUB True (NUMERAL_BIT1 xa) (NUMERAL_BIT2 xb) =
+NUMERAL_BIT1 (iSUB False xa xb) &
+iSUB False (NUMERAL_BIT1 xa) (NUMERAL_BIT2 xb) = iDUB (iSUB False xa xb) &
+iSUB False (NUMERAL_BIT2 xa) ALT_ZERO = NUMERAL_BIT1 xa &
+iSUB True (NUMERAL_BIT2 xa) (NUMERAL_BIT1 xb) =
+NUMERAL_BIT1 (iSUB True xa xb) &
+iSUB False (NUMERAL_BIT2 xa) (NUMERAL_BIT1 xb) = iDUB (iSUB True xa xb) &
+iSUB True (NUMERAL_BIT2 xa) (NUMERAL_BIT2 xb) = iDUB (iSUB True xa xb) &
+iSUB False (NUMERAL_BIT2 xa) (NUMERAL_BIT2 xb) =
+NUMERAL_BIT1 (iSUB False xa xb)"
+  sorry
+
+lemma numeral_sub: "NUMERAL (x - xa) = (if xa < x then NUMERAL (iSUB True x xa) else 0)"
+  sorry
+
+lemma iDUB_removal: "iDUB (NUMERAL_BIT1 x) = NUMERAL_BIT2 (iDUB x) &
+iDUB (NUMERAL_BIT2 x) = NUMERAL_BIT2 (NUMERAL_BIT1 x) &
+iDUB ALT_ZERO = ALT_ZERO"
+  sorry
+
+lemma numeral_mult: "ALT_ZERO * x = ALT_ZERO &
+x * ALT_ZERO = ALT_ZERO &
+NUMERAL_BIT1 x * xa = iZ (iDUB (x * xa) + xa) &
+NUMERAL_BIT2 x * xa = iDUB (iZ (x * xa + xa))"
+  sorry
+
+definition
+  iSQR :: "nat => nat"  where
+  "iSQR == %x. x * x"
+
+lemma iSQR: "iSQR x = x * x"
+  sorry
+
+lemma numeral_exp: "(ALL x. x ^ ALT_ZERO = NUMERAL_BIT1 ALT_ZERO) &
+(ALL x xa. x ^ NUMERAL_BIT1 xa = x * iSQR (x ^ xa)) &
+(ALL x xa. x ^ NUMERAL_BIT2 xa = iSQR x * iSQR (x ^ xa))"
+  sorry
+
+lemma numeral_evenodd: "EVEN ALT_ZERO &
+EVEN (NUMERAL_BIT2 x) &
+~ EVEN (NUMERAL_BIT1 x) &
+~ ODD ALT_ZERO & ~ ODD (NUMERAL_BIT2 x) & ODD (NUMERAL_BIT1 x)"
+  sorry
+
+lemma numeral_fact: "FACT n = (if n = 0 then 1 else n * FACT (PRE n))"
+  sorry
+
+lemma numeral_funpow: "(f ^^ n) x = (if n = 0 then x else (f ^^ (n - 1)) (f x))"
+  sorry
 
 ;end_setup
 
 ;setup_theory ind_type
 
-lemma INJ_INVERSE2: "ALL P::'A::type => 'B::type => 'C::type.
-   (ALL (x1::'A::type) (y1::'B::type) (x2::'A::type) y2::'B::type.
-       (P x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2)) -->
-   (EX (x::'C::type => 'A::type) Y::'C::type => 'B::type.
-       ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)"
-  by (import ind_type INJ_INVERSE2)
-
-definition NUMPAIR :: "nat => nat => nat" where 
-  "NUMPAIR == %(x::nat) y::nat. 2 ^ x * (2 * y + 1)"
-
-lemma NUMPAIR: "ALL (x::nat) y::nat. NUMPAIR x y = 2 ^ x * (2 * y + 1)"
-  by (import ind_type NUMPAIR)
-
-lemma NUMPAIR_INJ_LEMMA: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat.
-   NUMPAIR x xa = NUMPAIR xb xc --> x = xb"
-  by (import ind_type NUMPAIR_INJ_LEMMA)
-
-lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat.
-   (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
-  by (import ind_type NUMPAIR_INJ)
+lemma INJ_INVERSE2: "(!!(x1::'A) (y1::'B) (x2::'A) y2::'B.
+    ((P::'A => 'B => 'C) x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2))
+==> EX (x::'C => 'A) Y::'C => 'B.
+       ALL (xa::'A) y::'B. x (P xa y) = xa & Y (P xa y) = y"
+  sorry
+
+definition
+  NUMPAIR :: "nat => nat => nat"  where
+  "NUMPAIR == %x y. 2 ^ x * (2 * y + 1)"
+
+lemma NUMPAIR: "NUMPAIR x y = 2 ^ x * (2 * y + 1)"
+  sorry
+
+lemma NUMPAIR_INJ_LEMMA: "NUMPAIR x xa = NUMPAIR xb xc ==> x = xb"
+  sorry
+
+lemma NUMPAIR_INJ: "(NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
+  sorry
 
 consts
   NUMSND :: "nat => nat" 
   NUMFST :: "nat => nat" 
 
-specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y"
-  by (import ind_type NUMPAIR_DEST)
-
-definition NUMSUM :: "bool => nat => nat" where 
-  "NUMSUM == %(b::bool) x::nat. if b then Suc (2 * x) else 2 * x"
-
-lemma NUMSUM: "ALL (b::bool) x::nat. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)"
-  by (import ind_type NUMSUM)
-
-lemma NUMSUM_INJ: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat.
-   (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
-  by (import ind_type NUMSUM_INJ)
+specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL x y. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y"
+  sorry
+
+definition
+  NUMSUM :: "bool => nat => nat"  where
+  "NUMSUM == %b x. if b then Suc (2 * x) else 2 * x"
+
+lemma NUMSUM: "NUMSUM b x = (if b then Suc (2 * x) else 2 * x)"
+  sorry
+
+lemma NUMSUM_INJ: "(NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
+  sorry
 
 consts
   NUMRIGHT :: "nat => nat" 
   NUMLEFT :: "nat => bool" 
 
-specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y"
-  by (import ind_type NUMSUM_DEST)
-
-definition INJN :: "nat => nat => 'a => bool" where 
-  "INJN == %(m::nat) (n::nat) a::'a::type. n = m"
-
-lemma INJN: "ALL m::nat. INJN m = (%(n::nat) a::'a::type. n = m)"
-  by (import ind_type INJN)
-
-lemma INJN_INJ: "ALL (n1::nat) n2::nat. (INJN n1 = INJN n2) = (n1 = n2)"
-  by (import ind_type INJN_INJ)
-
-definition INJA :: "'a => nat => 'a => bool" where 
-  "INJA == %(a::'a::type) (n::nat) b::'a::type. b = a"
-
-lemma INJA: "ALL a::'a::type. INJA a = (%(n::nat) b::'a::type. b = a)"
-  by (import ind_type INJA)
-
-lemma INJA_INJ: "ALL (a1::'a::type) a2::'a::type. (INJA a1 = INJA a2) = (a1 = a2)"
-  by (import ind_type INJA_INJ)
-
-definition INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" where 
-  "INJF == %(f::nat => nat => 'a::type => bool) n::nat. f (NUMFST n) (NUMSND n)"
-
-lemma INJF: "ALL f::nat => nat => 'a::type => bool.
-   INJF f = (%n::nat. f (NUMFST n) (NUMSND n))"
-  by (import ind_type INJF)
-
-lemma INJF_INJ: "ALL (f1::nat => nat => 'a::type => bool) f2::nat => nat => 'a::type => bool.
-   (INJF f1 = INJF f2) = (f1 = f2)"
-  by (import ind_type INJF_INJ)
-
-definition INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" where 
+specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL x y. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y"
+  sorry
+
+definition
+  INJN :: "nat => nat => 'a => bool"  where
+  "INJN == %m n a. n = m"
+
+lemma INJN: "INJN m = (%n a. n = m)"
+  sorry
+
+lemma INJN_INJ: "(INJN n1 = INJN n2) = (n1 = n2)"
+  sorry
+
+definition
+  INJA :: "'a => nat => 'a => bool"  where
+  "INJA == %a n b. b = a"
+
+lemma INJA: "INJA a = (%n b. b = a)"
+  sorry
+
+lemma INJA_INJ: "(INJA a1 = INJA a2) = (a1 = a2)"
+  sorry
+
+definition
+  INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool"  where
+  "INJF == %f n. f (NUMFST n) (NUMSND n)"
+
+lemma INJF: "INJF f = (%n. f (NUMFST n) (NUMSND n))"
+  sorry
+
+lemma INJF_INJ: "(INJF f1 = INJF f2) = (f1 = f2)"
+  sorry
+
+definition
+  INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool"  where
   "INJP ==
-%(f1::nat => 'a::type => bool) (f2::nat => 'a::type => bool) (n::nat)
-   a::'a::type. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a"
-
-lemma INJP: "ALL (f1::nat => 'a::type => bool) f2::nat => 'a::type => bool.
-   INJP f1 f2 =
-   (%(n::nat) a::'a::type.
-       if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)"
-  by (import ind_type INJP)
-
-lemma INJP_INJ: "ALL (f1::nat => 'a::type => bool) (f1'::nat => 'a::type => bool)
-   (f2::nat => 'a::type => bool) f2'::nat => 'a::type => bool.
-   (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
-  by (import ind_type INJP_INJ)
-
-definition ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" where 
-  "ZCONSTR ==
-%(c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-   INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
-
-lemma ZCONSTR: "ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-   ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
-  by (import ind_type ZCONSTR)
-
-definition ZBOT :: "nat => 'a => bool" where 
-  "ZBOT == INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
-
-lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
-  by (import ind_type ZBOT)
-
-lemma ZCONSTR_ZBOT: "ALL (x::nat) (xa::'a::type) xb::nat => nat => 'a::type => bool.
-   ZCONSTR x xa xb ~= ZBOT"
-  by (import ind_type ZCONSTR_ZBOT)
-
-definition ZRECSPACE :: "(nat => 'a => bool) => bool" where 
+%f1 f2 n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a"
+
+lemma INJP: "INJP f1 f2 =
+(%n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)"
+  sorry
+
+lemma INJP_INJ: "(INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
+  sorry
+
+definition
+  ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool"  where
+  "ZCONSTR == %c i r. INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
+
+lemma ZCONSTR: "ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
+  sorry
+
+definition
+  ZBOT :: "nat => 'a => bool"  where
+  "ZBOT == INJP (INJN 0) (SOME z. True)"
+
+lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z. True)"
+  sorry
+
+lemma ZCONSTR_ZBOT: "ZCONSTR x xa xb ~= ZBOT"
+  sorry
+
+definition
+  ZRECSPACE :: "(nat => 'a => bool) => bool"  where
   "ZRECSPACE ==
-%a0::nat => 'a::type => bool.
-   ALL ZRECSPACE'::(nat => 'a::type => bool) => bool.
-      (ALL a0::nat => 'a::type => bool.
-          a0 = ZBOT |
-          (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-              a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
-          ZRECSPACE' a0) -->
-      ZRECSPACE' a0"
+%a0. ALL ZRECSPACE'.
+        (ALL a0.
+            a0 = ZBOT |
+            (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE' (r n))) -->
+            ZRECSPACE' a0) -->
+        ZRECSPACE' a0"
 
 lemma ZRECSPACE: "ZRECSPACE =
-(%a0::nat => 'a::type => bool.
-    ALL ZRECSPACE'::(nat => 'a::type => bool) => bool.
-       (ALL a0::nat => 'a::type => bool.
-           a0 = ZBOT |
-           (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-               a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
-           ZRECSPACE' a0) -->
-       ZRECSPACE' a0)"
-  by (import ind_type ZRECSPACE)
+(%a0. ALL ZRECSPACE'.
+         (ALL a0.
+             a0 = ZBOT |
+             (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE' (r n))) -->
+             ZRECSPACE' a0) -->
+         ZRECSPACE' a0)"
+  sorry
 
 lemma ZRECSPACE_rules: "(op &::bool => bool => bool)
  ((ZRECSPACE::(nat => 'a::type => bool) => bool)
@@ -2904,26 +2212,19 @@
                                   => (nat => nat => 'a::type => bool)
                                      => nat => 'a::type => bool)
                       c i r))))))"
-  by (import ind_type ZRECSPACE_rules)
-
-lemma ZRECSPACE_ind: "ALL x::(nat => 'a::type => bool) => bool.
-   x ZBOT &
-   (ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-       (ALL n::nat. x (r n)) --> x (ZCONSTR c i r)) -->
-   (ALL a0::nat => 'a::type => bool. ZRECSPACE a0 --> x a0)"
-  by (import ind_type ZRECSPACE_ind)
-
-lemma ZRECSPACE_cases: "ALL a0::nat => 'a::type => bool.
-   ZRECSPACE a0 =
-   (a0 = ZBOT |
-    (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
-        a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE (r n))))"
-  by (import ind_type ZRECSPACE_cases)
-
-typedef (open) ('a) recspace = "(Collect::((nat => 'a::type => bool) => bool)
-          => (nat => 'a::type => bool) set)
- (ZRECSPACE::(nat => 'a::type => bool) => bool)" 
-  by (rule typedef_helper,import ind_type recspace_TY_DEF)
+  sorry
+
+lemma ZRECSPACE_ind: "[| x ZBOT & (ALL c i r. (ALL n. x (r n)) --> x (ZCONSTR c i r));
+   ZRECSPACE a0 |]
+==> x a0"
+  sorry
+
+lemma ZRECSPACE_cases: "ZRECSPACE a0 =
+(a0 = ZBOT | (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE (r n))))"
+  sorry
+
+typedef (open) ('a) recspace = "Collect ZRECSPACE :: (nat \<Rightarrow> 'a\<Colon>type \<Rightarrow> bool) set"
+  sorry
 
 lemmas recspace_TY_DEF = typedef_hol2hol4 [OF type_definition_recspace]
 
@@ -2931,110 +2232,85 @@
   mk_rec :: "(nat => 'a => bool) => 'a recspace" 
   dest_rec :: "'a recspace => nat => 'a => bool" 
 
-specification (dest_rec mk_rec) recspace_repfns: "(ALL a::'a::type recspace. mk_rec (dest_rec a) = a) &
-(ALL r::nat => 'a::type => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))"
-  by (import ind_type recspace_repfns)
-
-definition BOTTOM :: "'a recspace" where 
+specification (dest_rec mk_rec) recspace_repfns: "(ALL a::'a recspace. mk_rec (dest_rec a) = a) &
+(ALL r::nat => 'a => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))"
+  sorry
+
+definition
+  BOTTOM :: "'a recspace"  where
   "BOTTOM == mk_rec ZBOT"
 
 lemma BOTTOM: "BOTTOM = mk_rec ZBOT"
-  by (import ind_type BOTTOM)
-
-definition CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" where 
-  "CONSTR ==
-%(c::nat) (i::'a::type) r::nat => 'a::type recspace.
-   mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))"
-
-lemma CONSTR: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
-   CONSTR c i r = mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))"
-  by (import ind_type CONSTR)
-
-lemma MK_REC_INJ: "ALL (x::nat => 'a::type => bool) y::nat => 'a::type => bool.
-   mk_rec x = mk_rec y --> ZRECSPACE x & ZRECSPACE y --> x = y"
-  by (import ind_type MK_REC_INJ)
-
-lemma DEST_REC_INJ: "ALL (x::'a::type recspace) y::'a::type recspace.
-   (dest_rec x = dest_rec y) = (x = y)"
-  by (import ind_type DEST_REC_INJ)
-
-lemma CONSTR_BOT: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
-   CONSTR c i r ~= BOTTOM"
-  by (import ind_type CONSTR_BOT)
-
-lemma CONSTR_INJ: "ALL (c1::nat) (i1::'a::type) (r1::nat => 'a::type recspace) (c2::nat)
-   (i2::'a::type) r2::nat => 'a::type recspace.
-   (CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)"
-  by (import ind_type CONSTR_INJ)
-
-lemma CONSTR_IND: "ALL P::'a::type recspace => bool.
-   P BOTTOM &
-   (ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
-       (ALL n::nat. P (r n)) --> P (CONSTR c i r)) -->
-   All P"
-  by (import ind_type CONSTR_IND)
-
-lemma CONSTR_REC: "ALL Fn::nat
-        => 'a::type
-           => (nat => 'a::type recspace) => (nat => 'b::type) => 'b::type.
-   EX f::'a::type recspace => 'b::type.
-      ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
-         f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))"
-  by (import ind_type CONSTR_REC)
+  sorry
+
+definition
+  CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace"  where
+  "CONSTR == %c i r. mk_rec (ZCONSTR c i (%n. dest_rec (r n)))"
+
+lemma CONSTR: "CONSTR c i r = mk_rec (ZCONSTR c i (%n. dest_rec (r n)))"
+  sorry
+
+lemma MK_REC_INJ: "[| mk_rec x = mk_rec y; ZRECSPACE x & ZRECSPACE y |] ==> x = y"
+  sorry
+
+lemma DEST_REC_INJ: "(dest_rec x = dest_rec y) = (x = y)"
+  sorry
+
+lemma CONSTR_BOT: "CONSTR c i r ~= BOTTOM"
+  sorry
+
+lemma CONSTR_INJ: "(CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)"
+  sorry
+
+lemma CONSTR_IND: "P BOTTOM & (ALL c i r. (ALL n. P (r n)) --> P (CONSTR c i r)) ==> P x"
+  sorry
+
+lemma CONSTR_REC: "EX f. ALL c i r. f (CONSTR c i r) = Fn c i r (%n. f (r n))"
+  sorry
 
 consts
   FCONS :: "'a => (nat => 'a) => nat => 'a" 
 
-specification (FCONS) FCONS: "(ALL (a::'a::type) f::nat => 'a::type. FCONS a f 0 = a) &
-(ALL (a::'a::type) (f::nat => 'a::type) n::nat. FCONS a f (Suc n) = f n)"
-  by (import ind_type FCONS)
-
-definition FNIL :: "nat => 'a" where 
-  "FNIL == %n::nat. SOME x::'a::type. True"
-
-lemma FNIL: "ALL n::nat. FNIL n = (SOME x::'a::type. True)"
-  by (import ind_type FNIL)
-
-definition ISO :: "('a => 'b) => ('b => 'a) => bool" where 
-  "ISO ==
-%(f::'a::type => 'b::type) g::'b::type => 'a::type.
-   (ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y)"
-
-lemma ISO: "ALL (f::'a::type => 'b::type) g::'b::type => 'a::type.
-   ISO f g =
-   ((ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y))"
-  by (import ind_type ISO)
-
-lemma ISO_REFL: "ISO (%x::'a::type. x) (%x::'a::type. x)"
-  by (import ind_type ISO_REFL)
-
-lemma ISO_FUN: "ISO (f::'a::type => 'c::type) (f'::'c::type => 'a::type) &
-ISO (g::'b::type => 'd::type) (g'::'d::type => 'b::type) -->
-ISO (%(h::'a::type => 'b::type) a'::'c::type. g (h (f' a')))
- (%(h::'c::type => 'd::type) a::'a::type. g' (h (f a)))"
-  by (import ind_type ISO_FUN)
-
-lemma ISO_USAGE: "ISO (f::'a::type => 'b::type) (g::'b::type => 'a::type) -->
-(ALL P::'a::type => bool. All P = (ALL x::'b::type. P (g x))) &
-(ALL P::'a::type => bool. Ex P = (EX x::'b::type. P (g x))) &
-(ALL (a::'a::type) b::'b::type. (a = g b) = (f a = b))"
-  by (import ind_type ISO_USAGE)
+specification (FCONS) FCONS: "(ALL (a::'a) f::nat => 'a. FCONS a f (0::nat) = a) &
+(ALL (a::'a) (f::nat => 'a) n::nat. FCONS a f (Suc n) = f n)"
+  sorry
+
+definition
+  FNIL :: "nat => 'a"  where
+  "FNIL == %n. SOME x. True"
+
+lemma FNIL: "FNIL n = (SOME x. True)"
+  sorry
+
+definition
+  ISO :: "('a => 'b) => ('b => 'a) => bool"  where
+  "ISO == %f g. (ALL x. f (g x) = x) & (ALL y. g (f y) = y)"
+
+lemma ISO: "ISO f g = ((ALL x. f (g x) = x) & (ALL y. g (f y) = y))"
+  sorry
+
+lemma ISO_REFL: "ISO (%x. x) (%x. x)"
+  sorry
+
+lemma ISO_FUN: "ISO (f::'a => 'c) (f'::'c => 'a) & ISO (g::'b => 'd) (g'::'d => 'b)
+==> ISO (%(h::'a => 'b) a'::'c. g (h (f' a')))
+     (%(h::'c => 'd) a::'a. g' (h (f a)))"
+  sorry
+
+lemma ISO_USAGE: "ISO f g
+==> (ALL P. All P = (ALL x. P (g x))) &
+    (ALL P. Ex P = (EX x. P (g x))) & (ALL a b. (a = g b) = (f a = b))"
+  sorry
 
 ;end_setup
 
 ;setup_theory divides
 
-lemma ONE_DIVIDES_ALL: "(All::(nat => bool) => bool) ((op dvd::nat => nat => bool) (1::nat))"
-  by (import divides ONE_DIVIDES_ALL)
-
-lemma DIVIDES_ADD_2: "ALL (a::nat) (b::nat) c::nat. a dvd b & a dvd b + c --> a dvd c"
-  by (import divides DIVIDES_ADD_2)
-
-lemma DIVIDES_FACT: "ALL b>0. b dvd FACT b"
-  by (import divides DIVIDES_FACT)
-
-lemma DIVIDES_MULT_LEFT: "ALL (x::nat) xa::nat. (x * xa dvd xa) = (xa = 0 | x = 1)"
-  by (import divides DIVIDES_MULT_LEFT)
+lemma DIVIDES_FACT: "0 < b ==> b dvd FACT b"
+  sorry
+
+lemma DIVIDES_MULT_LEFT: "((x::nat) * (xa::nat) dvd xa) = (xa = (0::nat) | x = (1::nat))"
+  sorry
 
 ;end_setup
 
@@ -3044,17 +2320,16 @@
   prime :: "nat => bool" 
 
 defs
-  prime_primdef: "prime.prime == %a::nat. a ~= 1 & (ALL b::nat. b dvd a --> b = a | b = 1)"
-
-lemma prime_def: "ALL a::nat.
-   prime.prime a = (a ~= 1 & (ALL b::nat. b dvd a --> b = a | b = 1))"
-  by (import prime prime_def)
+  prime_primdef: "prime.prime == %a. a ~= 1 & (ALL b. b dvd a --> b = a | b = 1)"
+
+lemma prime_def: "prime.prime a = (a ~= 1 & (ALL b. b dvd a --> b = a | b = 1))"
+  sorry
 
 lemma NOT_PRIME_0: "~ prime.prime 0"
-  by (import prime NOT_PRIME_0)
+  sorry
 
 lemma NOT_PRIME_1: "~ prime.prime 1"
-  by (import prime NOT_PRIME_1)
+  sorry
 
 ;end_setup
 
@@ -3063,997 +2338,758 @@
 consts
   EL :: "nat => 'a list => 'a" 
 
-specification (EL) EL: "(ALL l::'a::type list. EL 0 l = hd l) &
-(ALL (l::'a::type list) n::nat. EL (Suc n) l = EL n (tl l))"
-  by (import list EL)
+specification (EL) EL: "(ALL l::'a list. EL (0::nat) l = hd l) &
+(ALL (l::'a list) n::nat. EL (Suc n) l = EL n (tl l))"
+  sorry
 
 lemma NULL: "(op &::bool => bool => bool)
- ((null::'a::type list => bool) ([]::'a::type list))
+ ((List.null::'a::type list => bool) ([]::'a::type list))
  ((All::('a::type => bool) => bool)
    (%x::'a::type.
        (All::('a::type list => bool) => bool)
         (%xa::'a::type list.
             (Not::bool => bool)
-             ((null::'a::type list => bool)
+             ((List.null::'a::type list => bool)
                ((op #::'a::type => 'a::type list => 'a::type list) x xa)))))"
-  by (import list NULL)
-
-lemma list_case_compute: "ALL l::'a::type list.
-   list_case (b::'b::type) (f::'a::type => 'a::type list => 'b::type) l =
-   (if null l then b else f (hd l) (tl l))"
-  by (import list list_case_compute)
-
-lemma LIST_NOT_EQ: "ALL (l1::'a::type list) l2::'a::type list.
-   l1 ~= l2 --> (ALL (x::'a::type) xa::'a::type. x # l1 ~= xa # l2)"
-  by (import list LIST_NOT_EQ)
-
-lemma NOT_EQ_LIST: "ALL (h1::'a::type) h2::'a::type.
-   h1 ~= h2 -->
-   (ALL (x::'a::type list) xa::'a::type list. h1 # x ~= h2 # xa)"
-  by (import list NOT_EQ_LIST)
-
-lemma EQ_LIST: "ALL (h1::'a::type) h2::'a::type.
-   h1 = h2 -->
-   (ALL (l1::'a::type list) l2::'a::type list.
-       l1 = l2 --> h1 # l1 = h2 # l2)"
-  by (import list EQ_LIST)
-
-lemma CONS: "ALL l::'a::type list. ~ null l --> hd l # tl l = l"
-  by (import list CONS)
-
-lemma MAP_EQ_NIL: "ALL (l::'a::type list) f::'a::type => 'b::type.
-   (map f l = []) = (l = []) & ([] = map f l) = (l = [])"
-  by (import list MAP_EQ_NIL)
-
-lemma EVERY_EL: "(All::('a::type list => bool) => bool)
- (%l::'a::type list.
-     (All::(('a::type => bool) => bool) => bool)
-      (%P::'a::type => bool.
-          (op =::bool => bool => bool)
-           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
-           ((All::(nat => bool) => bool)
-             (%n::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <::nat => nat => bool) n
-                    ((size::'a::type list => nat) l))
-                  (P ((EL::nat => 'a::type list => 'a::type) n l))))))"
-  by (import list EVERY_EL)
-
-lemma EVERY_CONJ: "ALL l::'a::type list.
-   list_all
-    (%x::'a::type. (P::'a::type => bool) x & (Q::'a::type => bool) x) l =
-   (list_all P l & list_all Q l)"
-  by (import list EVERY_CONJ)
-
-lemma EVERY_MEM: "ALL (P::'a::type => bool) l::'a::type list.
-   list_all P l = (ALL e::'a::type. e mem l --> P e)"
-  by (import list EVERY_MEM)
-
-lemma EXISTS_MEM: "ALL (P::'a::type => bool) l::'a::type list.
-   list_ex P l = (EX e::'a::type. e mem l & P e)"
-  by (import list EXISTS_MEM)
-
-lemma MEM_APPEND: "ALL (e::'a::type) (l1::'a::type list) l2::'a::type list.
-   e mem l1 @ l2 = (e mem l1 | e mem l2)"
-  by (import list MEM_APPEND)
-
-lemma EXISTS_APPEND: "ALL (P::'a::type => bool) (l1::'a::type list) l2::'a::type list.
-   list_ex P (l1 @ l2) = (list_ex P l1 | list_ex P l2)"
-  by (import list EXISTS_APPEND)
-
-lemma NOT_EVERY: "ALL (P::'a::type => bool) l::'a::type list.
-   (~ list_all P l) = list_ex (Not o P) l"
-  by (import list NOT_EVERY)
-
-lemma NOT_EXISTS: "ALL (P::'a::type => bool) l::'a::type list.
-   (~ list_ex P l) = list_all (Not o P) l"
-  by (import list NOT_EXISTS)
-
-lemma MEM_MAP: "ALL (l::'a::type list) (f::'a::type => 'b::type) x::'b::type.
-   x mem map f l = (EX y::'a::type. x = f y & y mem l)"
-  by (import list MEM_MAP)
-
-lemma LENGTH_CONS: "ALL (l::'a::type list) n::nat.
-   (length l = Suc n) =
-   (EX (h::'a::type) l'::'a::type list. length l' = n & l = h # l')"
-  by (import list LENGTH_CONS)
-
-lemma LENGTH_EQ_CONS: "ALL (P::'a::type list => bool) n::nat.
-   (ALL l::'a::type list. length l = Suc n --> P l) =
-   (ALL l::'a::type list. length l = n --> (ALL x::'a::type. P (x # l)))"
-  by (import list LENGTH_EQ_CONS)
-
-lemma LENGTH_EQ_NIL: "ALL P::'a::type list => bool.
-   (ALL l::'a::type list. length l = 0 --> P l) = P []"
-  by (import list LENGTH_EQ_NIL)
-
-lemma CONS_ACYCLIC: "ALL (l::'a::type list) x::'a::type. l ~= x # l & x # l ~= l"
-  by (import list CONS_ACYCLIC)
-
-lemma APPEND_eq_NIL: "(ALL (l1::'a::type list) l2::'a::type list.
-    ([] = l1 @ l2) = (l1 = [] & l2 = [])) &
-(ALL (l1::'a::type list) l2::'a::type list.
-    (l1 @ l2 = []) = (l1 = [] & l2 = []))"
-  by (import list APPEND_eq_NIL)
-
-lemma APPEND_11: "(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list.
+  sorry
+
+lemma list_case_compute: "list_case (b::'b) (f::'a => 'a list => 'b) (l::'a list) =
+(if List.null l then b else f (hd l) (tl l))"
+  sorry
+
+lemma LIST_NOT_EQ: "l1 ~= l2 ==> x # l1 ~= xa # l2"
+  sorry
+
+lemma NOT_EQ_LIST: "h1 ~= h2 ==> h1 # x ~= h2 # xa"
+  sorry
+
+lemma EQ_LIST: "[| h1 = h2; l1 = l2 |] ==> h1 # l1 = h2 # l2"
+  sorry
+
+lemma CONS: "~ List.null l ==> hd l # tl l = l"
+  sorry
+
+lemma MAP_EQ_NIL: "(map (f::'a => 'b) (l::'a list) = []) = (l = []) & ([] = map f l) = (l = [])"
+  sorry
+
+lemma EVERY_EL: "list_all P l = (ALL n<length l. P (EL n l))"
+  sorry
+
+lemma EVERY_CONJ: "list_all (%x. P x & Q x) l = (list_all P l & list_all Q l)"
+  sorry
+
+lemma EVERY_MEM: "list_all P l = (ALL e. List.member l e --> P e)"
+  sorry
+
+lemma EXISTS_MEM: "list_ex P l = (EX e. List.member l e & P e)"
+  sorry
+
+lemma MEM_APPEND: "List.member (l1 @ l2) e = (List.member l1 e | List.member l2 e)"
+  sorry
+
+lemma NOT_EVERY: "(~ list_all P l) = list_ex (Not o P) l"
+  sorry
+
+lemma NOT_EXISTS: "(~ list_ex P l) = list_all (Not o P) l"
+  sorry
+
+lemma MEM_MAP: "List.member (map (f::'a => 'b) (l::'a list)) (x::'b) =
+(EX y::'a. x = f y & List.member l y)"
+  sorry
+
+lemma LENGTH_CONS: "(length l = Suc n) = (EX h l'. length l' = n & l = h # l')"
+  sorry
+
+lemma LENGTH_EQ_CONS: "(ALL l. length l = Suc n --> P l) =
+(ALL l. length l = n --> (ALL x. P (x # l)))"
+  sorry
+
+lemma LENGTH_EQ_NIL: "(ALL l. length l = 0 --> P l) = P []"
+  sorry
+
+lemma CONS_ACYCLIC: "l ~= x # l & x # l ~= l"
+  sorry
+
+lemma APPEND_eq_NIL: "(ALL (l1::'a list) l2::'a list. ([] = l1 @ l2) = (l1 = [] & l2 = [])) &
+(ALL (l1::'a list) l2::'a list. (l1 @ l2 = []) = (l1 = [] & l2 = []))"
+  sorry
+
+lemma APPEND_11: "(ALL (l1::'a list) (l2::'a list) l3::'a list.
     (l1 @ l2 = l1 @ l3) = (l2 = l3)) &
-(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list.
+(ALL (l1::'a list) (l2::'a list) l3::'a list.
     (l2 @ l1 = l3 @ l1) = (l2 = l3))"
-  by (import list APPEND_11)
-
-lemma EL_compute: "ALL n::nat.
-   EL n (l::'a::type list) = (if n = 0 then hd l else EL (PRE n) (tl l))"
-  by (import list EL_compute)
-
-lemma WF_LIST_PRED: "WF (%(L1::'a::type list) L2::'a::type list. EX h::'a::type. L2 = h # L1)"
-  by (import list WF_LIST_PRED)
-
-lemma list_size_cong: "ALL (M::'a::type list) (N::'a::type list) (f::'a::type => nat)
-   f'::'a::type => nat.
-   M = N & (ALL x::'a::type. x mem N --> f x = f' x) -->
-   list_size f M = list_size f' N"
-  by (import list list_size_cong)
-
-lemma FOLDR_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type)
-   (f::'a::type => 'b::type => 'b::type)
-   f'::'a::type => 'b::type => 'b::type.
-   l = l' &
-   b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f x a = f' x a) -->
-   foldr f l b = foldr f' l' b'"
-  by (import list FOLDR_CONG)
-
-lemma FOLDL_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type)
-   (f::'b::type => 'a::type => 'b::type)
-   f'::'b::type => 'a::type => 'b::type.
-   l = l' &
-   b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f a x = f' a x) -->
-   foldl f b l = foldl f' b' l'"
-  by (import list FOLDL_CONG)
-
-lemma MAP_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (f::'a::type => 'b::type)
-   f'::'a::type => 'b::type.
-   l1 = l2 & (ALL x::'a::type. x mem l2 --> f x = f' x) -->
-   map f l1 = map f' l2"
-  by (import list MAP_CONG)
-
-lemma EXISTS_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool)
-   P'::'a::type => bool.
-   l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) -->
-   list_ex P l1 = list_ex P' l2"
-  by (import list EXISTS_CONG)
-
-lemma EVERY_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool)
-   P'::'a::type => bool.
-   l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) -->
-   list_all P l1 = list_all P' l2"
-  by (import list EVERY_CONG)
-
-lemma EVERY_MONOTONIC: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   (ALL x::'a::type. P x --> Q x) -->
-   (ALL l::'a::type list. list_all P l --> list_all Q l)"
-  by (import list EVERY_MONOTONIC)
-
-lemma LENGTH_ZIP: "ALL (l1::'a::type list) l2::'b::type list.
-   length l1 = length l2 -->
-   length (zip l1 l2) = length l1 & length (zip l1 l2) = length l2"
-  by (import list LENGTH_ZIP)
-
-lemma LENGTH_UNZIP: "ALL pl::('a::type * 'b::type) list.
-   length (fst (unzip pl)) = length pl & length (snd (unzip pl)) = length pl"
-  by (import list LENGTH_UNZIP)
-
-lemma ZIP_UNZIP: "ALL l::('a::type * 'b::type) list. ZIP (unzip l) = l"
-  by (import list ZIP_UNZIP)
-
-lemma UNZIP_ZIP: "ALL (l1::'a::type list) l2::'b::type list.
-   length l1 = length l2 --> unzip (zip l1 l2) = (l1, l2)"
-  by (import list UNZIP_ZIP)
-
-lemma ZIP_MAP: "ALL (l1::'a::type list) (l2::'b::type list) (f1::'a::type => 'c::type)
-   f2::'b::type => 'd::type.
-   length l1 = length l2 -->
-   zip (map f1 l1) l2 =
-   map (%p::'a::type * 'b::type. (f1 (fst p), snd p)) (zip l1 l2) &
-   zip l1 (map f2 l2) =
-   map (%p::'a::type * 'b::type. (fst p, f2 (snd p))) (zip l1 l2)"
-  by (import list ZIP_MAP)
-
-lemma MEM_ZIP: "(All::('a::type list => bool) => bool)
- (%l1::'a::type list.
-     (All::('b::type list => bool) => bool)
-      (%l2::'b::type list.
-          (All::('a::type * 'b::type => bool) => bool)
-           (%p::'a::type * 'b::type.
-               (op -->::bool => bool => bool)
-                ((op =::nat => nat => bool)
-                  ((size::'a::type list => nat) l1)
-                  ((size::'b::type list => nat) l2))
-                ((op =::bool => bool => bool)
-                  ((op mem::'a::type * 'b::type
-                            => ('a::type * 'b::type) list => bool)
-                    p ((zip::'a::type list
-                             => 'b::type list => ('a::type * 'b::type) list)
-                        l1 l2))
-                  ((Ex::(nat => bool) => bool)
-                    (%n::nat.
-                        (op &::bool => bool => bool)
-                         ((op <::nat => nat => bool) n
-                           ((size::'a::type list => nat) l1))
-                         ((op =::'a::type * 'b::type
-                                 => 'a::type * 'b::type => bool)
-                           p ((Pair::'a::type
-                                     => 'b::type => 'a::type * 'b::type)
-                               ((EL::nat => 'a::type list => 'a::type) n l1)
-                               ((EL::nat => 'b::type list => 'b::type) n
-                                 l2)))))))))"
-  by (import list MEM_ZIP)
-
-lemma EL_ZIP: "ALL (l1::'a::type list) (l2::'b::type list) n::nat.
-   length l1 = length l2 & n < length l1 -->
-   EL n (zip l1 l2) = (EL n l1, EL n l2)"
-  by (import list EL_ZIP)
-
-lemma MAP2_ZIP: "(All::('a::type list => bool) => bool)
- (%l1::'a::type list.
-     (All::('b::type list => bool) => bool)
-      (%l2::'b::type list.
-          (op -->::bool => bool => bool)
-           ((op =::nat => nat => bool) ((size::'a::type list => nat) l1)
-             ((size::'b::type list => nat) l2))
-           ((All::(('a::type => 'b::type => 'c::type) => bool) => bool)
-             (%f::'a::type => 'b::type => 'c::type.
-                 (op =::'c::type list => 'c::type list => bool)
-                  ((map2::('a::type => 'b::type => 'c::type)
-                          => 'a::type list
-                             => 'b::type list => 'c::type list)
-                    f l1 l2)
-                  ((map::('a::type * 'b::type => 'c::type)
-                         => ('a::type * 'b::type) list => 'c::type list)
-                    ((split::('a::type => 'b::type => 'c::type)
-                             => 'a::type * 'b::type => 'c::type)
-                      f)
-                    ((zip::'a::type list
-                           => 'b::type list => ('a::type * 'b::type) list)
-                      l1 l2))))))"
-  by (import list MAP2_ZIP)
-
-lemma MEM_EL: "(All::('a::type list => bool) => bool)
- (%l::'a::type list.
-     (All::('a::type => bool) => bool)
-      (%x::'a::type.
-          (op =::bool => bool => bool)
-           ((op mem::'a::type => 'a::type list => bool) x l)
-           ((Ex::(nat => bool) => bool)
-             (%n::nat.
-                 (op &::bool => bool => bool)
-                  ((op <::nat => nat => bool) n
-                    ((size::'a::type list => nat) l))
-                  ((op =::'a::type => 'a::type => bool) x
-                    ((EL::nat => 'a::type list => 'a::type) n l))))))"
-  by (import list MEM_EL)
-
-lemma LAST_CONS: "(ALL x::'a::type. last [x] = x) &
-(ALL (x::'a::type) (xa::'a::type) xb::'a::type list.
-    last (x # xa # xb) = last (xa # xb))"
-  by (import list LAST_CONS)
-
-lemma FRONT_CONS: "(ALL x::'a::type. butlast [x] = []) &
-(ALL (x::'a::type) (xa::'a::type) xb::'a::type list.
+  sorry
+
+lemma EL_compute: "EL n l = (if n = 0 then hd l else EL (PRE n) (tl l))"
+  sorry
+
+lemma WF_LIST_PRED: "WF (%L1 L2. EX h. L2 = h # L1)"
+  sorry
+
+lemma list_size_cong: "M = N & (ALL x. List.member N x --> f x = f' x)
+==> HOL4Compat.list_size f M = HOL4Compat.list_size f' N"
+  sorry
+
+lemma FOLDR_CONG: "l = l' & b = b' & (ALL x a. List.member l' x --> f x a = f' x a)
+==> foldr f l b = foldr f' l' b'"
+  sorry
+
+lemma FOLDL_CONG: "l = l' & b = b' & (ALL x a. List.member l' x --> f a x = f' a x)
+==> foldl f b l = foldl f' b' l'"
+  sorry
+
+lemma MAP_CONG: "l1 = l2 & (ALL x. List.member l2 x --> f x = f' x) ==> map f l1 = map f' l2"
+  sorry
+
+lemma EXISTS_CONG: "l1 = l2 & (ALL x. List.member l2 x --> P x = P' x)
+==> list_ex P l1 = list_ex P' l2"
+  sorry
+
+lemma EVERY_CONG: "l1 = l2 & (ALL x. List.member l2 x --> P x = P' x)
+==> list_all P l1 = list_all P' l2"
+  sorry
+
+lemma EVERY_MONOTONIC: "[| !!x. P x ==> Q x; list_all P l |] ==> list_all Q l"
+  sorry
+
+lemma LENGTH_ZIP: "length l1 = length l2
+==> length (zip l1 l2) = length l1 & length (zip l1 l2) = length l2"
+  sorry
+
+lemma LENGTH_UNZIP: "length (fst (unzip pl)) = length pl & length (snd (unzip pl)) = length pl"
+  sorry
+
+lemma ZIP_UNZIP: "ZIP (unzip l) = l"
+  sorry
+
+lemma UNZIP_ZIP: "length l1 = length l2 ==> unzip (zip l1 l2) = (l1, l2)"
+  sorry
+
+lemma ZIP_MAP: "length l1 = length l2
+==> zip (map f1 l1) l2 = map (%p. (f1 (fst p), snd p)) (zip l1 l2) &
+    zip l1 (map f2 l2) = map (%p. (fst p, f2 (snd p))) (zip l1 l2)"
+  sorry
+
+lemma MEM_ZIP: "length l1 = length l2
+==> List.member (zip l1 l2) p = (EX n<length l1. p = (EL n l1, EL n l2))"
+  sorry
+
+lemma EL_ZIP: "length l1 = length l2 & n < length l1
+==> EL n (zip l1 l2) = (EL n l1, EL n l2)"
+  sorry
+
+lemma MAP2_ZIP: "length l1 = length l2 ==> map2 f l1 l2 = map (%(x, y). f x y) (zip l1 l2)"
+  sorry
+
+lemma MEM_EL: "List.member l x = (EX n<length l. x = EL n l)"
+  sorry
+
+lemma LAST_CONS: "(ALL x::'a. last [x] = x) &
+(ALL (x::'a) (xa::'a) xb::'a list. last (x # xa # xb) = last (xa # xb))"
+  sorry
+
+lemma FRONT_CONS: "(ALL x::'a. butlast [x] = []) &
+(ALL (x::'a) (xa::'a) xb::'a list.
     butlast (x # xa # xb) = x # butlast (xa # xb))"
-  by (import list FRONT_CONS)
+  sorry
 
 ;end_setup
 
 ;setup_theory pred_set
 
-lemma EXTENSION: "ALL (s::'a::type => bool) t::'a::type => bool.
-   (s = t) = (ALL x::'a::type. IN x s = IN x t)"
-  by (import pred_set EXTENSION)
-
-lemma NOT_EQUAL_SETS: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   (x ~= xa) = (EX xb::'a::type. IN xb xa = (~ IN xb x))"
-  by (import pred_set NOT_EQUAL_SETS)
-
-lemma NUM_SET_WOP: "ALL s::nat => bool.
-   (EX n::nat. IN n s) =
-   (EX n::nat. IN n s & (ALL m::nat. IN m s --> n <= m))"
-  by (import pred_set NUM_SET_WOP)
+lemma EXTENSION: "(s = t) = (ALL x. IN x s = IN x t)"
+  sorry
+
+lemma NOT_EQUAL_SETS: "(x ~= xa) = (EX xb. IN xb xa = (~ IN xb x))"
+  sorry
+
+lemma NUM_SET_WOP: "(EX n::nat. IN n (s::nat => bool)) =
+(EX n::nat. IN n s & (ALL m::nat. IN m s --> n <= m))"
+  sorry
 
 consts
   GSPEC :: "('b => 'a * bool) => 'a => bool" 
 
-specification (GSPEC) GSPECIFICATION: "ALL (f::'b::type => 'a::type * bool) v::'a::type.
-   IN v (GSPEC f) = (EX x::'b::type. (v, True) = f x)"
-  by (import pred_set GSPECIFICATION)
-
-lemma SET_MINIMUM: "ALL (s::'a::type => bool) M::'a::type => nat.
-   (EX x::'a::type. IN x s) =
-   (EX x::'a::type. IN x s & (ALL y::'a::type. IN y s --> M x <= M y))"
-  by (import pred_set SET_MINIMUM)
-
-definition EMPTY :: "'a => bool" where 
-  "EMPTY == %x::'a::type. False"
-
-lemma EMPTY_DEF: "EMPTY = (%x::'a::type. False)"
-  by (import pred_set EMPTY_DEF)
-
-lemma NOT_IN_EMPTY: "ALL x::'a::type. ~ IN x EMPTY"
-  by (import pred_set NOT_IN_EMPTY)
-
-lemma MEMBER_NOT_EMPTY: "ALL x::'a::type => bool. (EX xa::'a::type. IN xa x) = (x ~= EMPTY)"
-  by (import pred_set MEMBER_NOT_EMPTY)
-
-consts
-  UNIV :: "'a => bool" 
-
-defs
-  UNIV_def: "pred_set.UNIV == %x::'a::type. True"
-
-lemma UNIV_DEF: "pred_set.UNIV = (%x::'a::type. True)"
-  by (import pred_set UNIV_DEF)
-
-lemma IN_UNIV: "ALL x::'a::type. IN x pred_set.UNIV"
-  by (import pred_set IN_UNIV)
+specification (GSPEC) GSPECIFICATION: "ALL (f::'b => 'a * bool) v::'a. IN v (GSPEC f) = (EX x::'b. (v, True) = f x)"
+  sorry
+
+lemma SET_MINIMUM: "(EX x::'a. IN x (s::'a => bool)) =
+(EX x::'a. IN x s & (ALL y::'a. IN y s --> (M::'a => nat) x <= M y))"
+  sorry
+
+definition
+  EMPTY :: "'a => bool"  where
+  "EMPTY == %x. False"
+
+lemma EMPTY_DEF: "EMPTY = (%x. False)"
+  sorry
+
+lemma NOT_IN_EMPTY: "~ IN x EMPTY"
+  sorry
+
+lemma MEMBER_NOT_EMPTY: "(EX xa. IN xa x) = (x ~= EMPTY)"
+  sorry
+
+definition
+  UNIV :: "'a => bool"  where
+  "UNIV == %x. True"
+
+lemma UNIV_DEF: "pred_set.UNIV = (%x. True)"
+  sorry
+
+lemma IN_UNIV: "IN x pred_set.UNIV"
+  sorry
 
 lemma UNIV_NOT_EMPTY: "pred_set.UNIV ~= EMPTY"
-  by (import pred_set UNIV_NOT_EMPTY)
+  sorry
 
 lemma EMPTY_NOT_UNIV: "EMPTY ~= pred_set.UNIV"
-  by (import pred_set EMPTY_NOT_UNIV)
-
-lemma EQ_UNIV: "(ALL x::'a::type. IN x (s::'a::type => bool)) = (s = pred_set.UNIV)"
-  by (import pred_set EQ_UNIV)
-
-definition SUBSET :: "('a => bool) => ('a => bool) => bool" where 
-  "SUBSET ==
-%(s::'a::type => bool) t::'a::type => bool.
-   ALL x::'a::type. IN x s --> IN x t"
-
-lemma SUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   SUBSET s t = (ALL x::'a::type. IN x s --> IN x t)"
-  by (import pred_set SUBSET_DEF)
-
-lemma SUBSET_TRANS: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   SUBSET x xa & SUBSET xa xb --> SUBSET x xb"
-  by (import pred_set SUBSET_TRANS)
-
-lemma SUBSET_REFL: "ALL x::'a::type => bool. SUBSET x x"
-  by (import pred_set SUBSET_REFL)
-
-lemma SUBSET_ANTISYM: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   SUBSET x xa & SUBSET xa x --> x = xa"
-  by (import pred_set SUBSET_ANTISYM)
-
-lemma EMPTY_SUBSET: "All (SUBSET EMPTY)"
-  by (import pred_set EMPTY_SUBSET)
-
-lemma SUBSET_EMPTY: "ALL x::'a::type => bool. SUBSET x EMPTY = (x = EMPTY)"
-  by (import pred_set SUBSET_EMPTY)
-
-lemma SUBSET_UNIV: "ALL x::'a::type => bool. SUBSET x pred_set.UNIV"
-  by (import pred_set SUBSET_UNIV)
-
-lemma UNIV_SUBSET: "ALL x::'a::type => bool. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)"
-  by (import pred_set UNIV_SUBSET)
-
-definition PSUBSET :: "('a => bool) => ('a => bool) => bool" where 
-  "PSUBSET == %(s::'a::type => bool) t::'a::type => bool. SUBSET s t & s ~= t"
-
-lemma PSUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   PSUBSET s t = (SUBSET s t & s ~= t)"
-  by (import pred_set PSUBSET_DEF)
-
-lemma PSUBSET_TRANS: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb"
-  by (import pred_set PSUBSET_TRANS)
-
-lemma PSUBSET_IRREFL: "ALL x::'a::type => bool. ~ PSUBSET x x"
-  by (import pred_set PSUBSET_IRREFL)
-
-lemma NOT_PSUBSET_EMPTY: "ALL x::'a::type => bool. ~ PSUBSET x EMPTY"
-  by (import pred_set NOT_PSUBSET_EMPTY)
-
-lemma NOT_UNIV_PSUBSET: "ALL x::'a::type => bool. ~ PSUBSET pred_set.UNIV x"
-  by (import pred_set NOT_UNIV_PSUBSET)
-
-lemma PSUBSET_UNIV: "ALL x::'a::type => bool.
-   PSUBSET x pred_set.UNIV = (EX xa::'a::type. ~ IN xa x)"
-  by (import pred_set PSUBSET_UNIV)
-
-consts
-  UNION :: "('a => bool) => ('a => bool) => 'a => bool" 
-
-defs
-  UNION_def: "pred_set.UNION ==
-%(s::'a::type => bool) t::'a::type => bool.
-   GSPEC (%x::'a::type. (x, IN x s | IN x t))"
-
-lemma UNION_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   pred_set.UNION s t = GSPEC (%x::'a::type. (x, IN x s | IN x t))"
-  by (import pred_set UNION_DEF)
-
-lemma IN_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
-   IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)"
-  by (import pred_set IN_UNION)
-
-lemma UNION_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   pred_set.UNION x (pred_set.UNION xa xb) =
-   pred_set.UNION (pred_set.UNION x xa) xb"
-  by (import pred_set UNION_ASSOC)
-
-lemma UNION_IDEMPOT: "ALL x::'a::type => bool. pred_set.UNION x x = x"
-  by (import pred_set UNION_IDEMPOT)
-
-lemma UNION_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   pred_set.UNION x xa = pred_set.UNION xa x"
-  by (import pred_set UNION_COMM)
-
-lemma SUBSET_UNION: "(ALL (x::'a::type => bool) xa::'a::type => bool.
-    SUBSET x (pred_set.UNION x xa)) &
-(ALL (x::'a::type => bool) xa::'a::type => bool.
-    SUBSET x (pred_set.UNION xa x))"
-  by (import pred_set SUBSET_UNION)
-
-lemma UNION_SUBSET: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
-   SUBSET (pred_set.UNION s t) u = (SUBSET s u & SUBSET t u)"
-  by (import pred_set UNION_SUBSET)
-
-lemma SUBSET_UNION_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   SUBSET x xa = (pred_set.UNION x xa = xa)"
-  by (import pred_set SUBSET_UNION_ABSORPTION)
-
-lemma UNION_EMPTY: "(ALL x::'a::type => bool. pred_set.UNION EMPTY x = x) &
-(ALL x::'a::type => bool. pred_set.UNION x EMPTY = x)"
-  by (import pred_set UNION_EMPTY)
-
-lemma UNION_UNIV: "(ALL x::'a::type => bool. pred_set.UNION pred_set.UNIV x = pred_set.UNIV) &
-(ALL x::'a::type => bool. pred_set.UNION x pred_set.UNIV = pred_set.UNIV)"
-  by (import pred_set UNION_UNIV)
-
-lemma EMPTY_UNION: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   (pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)"
-  by (import pred_set EMPTY_UNION)
-
-consts
-  INTER :: "('a => bool) => ('a => bool) => 'a => bool" 
-
-defs
-  INTER_def: "pred_set.INTER ==
-%(s::'a::type => bool) t::'a::type => bool.
-   GSPEC (%x::'a::type. (x, IN x s & IN x t))"
-
-lemma INTER_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   pred_set.INTER s t = GSPEC (%x::'a::type. (x, IN x s & IN x t))"
-  by (import pred_set INTER_DEF)
-
-lemma IN_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
-   IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)"
-  by (import pred_set IN_INTER)
-
-lemma INTER_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   pred_set.INTER x (pred_set.INTER xa xb) =
-   pred_set.INTER (pred_set.INTER x xa) xb"
-  by (import pred_set INTER_ASSOC)
-
-lemma INTER_IDEMPOT: "ALL x::'a::type => bool. pred_set.INTER x x = x"
-  by (import pred_set INTER_IDEMPOT)
-
-lemma INTER_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   pred_set.INTER x xa = pred_set.INTER xa x"
-  by (import pred_set INTER_COMM)
-
-lemma INTER_SUBSET: "(ALL (x::'a::type => bool) xa::'a::type => bool.
-    SUBSET (pred_set.INTER x xa) x) &
-(ALL (x::'a::type => bool) xa::'a::type => bool.
-    SUBSET (pred_set.INTER xa x) x)"
-  by (import pred_set INTER_SUBSET)
-
-lemma SUBSET_INTER: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
-   SUBSET s (pred_set.INTER t u) = (SUBSET s t & SUBSET s u)"
-  by (import pred_set SUBSET_INTER)
-
-lemma SUBSET_INTER_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   SUBSET x xa = (pred_set.INTER x xa = x)"
-  by (import pred_set SUBSET_INTER_ABSORPTION)
-
-lemma INTER_EMPTY: "(ALL x::'a::type => bool. pred_set.INTER EMPTY x = EMPTY) &
-(ALL x::'a::type => bool. pred_set.INTER x EMPTY = EMPTY)"
-  by (import pred_set INTER_EMPTY)
-
-lemma INTER_UNIV: "(ALL x::'a::type => bool. pred_set.INTER pred_set.UNIV x = x) &
-(ALL x::'a::type => bool. pred_set.INTER x pred_set.UNIV = x)"
-  by (import pred_set INTER_UNIV)
-
-lemma UNION_OVER_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   pred_set.INTER x (pred_set.UNION xa xb) =
-   pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER x xb)"
-  by (import pred_set UNION_OVER_INTER)
-
-lemma INTER_OVER_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   pred_set.UNION x (pred_set.INTER xa xb) =
-   pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)"
-  by (import pred_set INTER_OVER_UNION)
-
-definition DISJOINT :: "('a => bool) => ('a => bool) => bool" where 
-  "DISJOINT ==
-%(s::'a::type => bool) t::'a::type => bool. pred_set.INTER s t = EMPTY"
-
-lemma DISJOINT_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   DISJOINT s t = (pred_set.INTER s t = EMPTY)"
-  by (import pred_set DISJOINT_DEF)
-
-lemma IN_DISJOINT: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   DISJOINT x xa = (~ (EX xb::'a::type. IN xb x & IN xb xa))"
-  by (import pred_set IN_DISJOINT)
-
-lemma DISJOINT_SYM: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   DISJOINT x xa = DISJOINT xa x"
-  by (import pred_set DISJOINT_SYM)
-
-lemma DISJOINT_EMPTY: "ALL x::'a::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY"
-  by (import pred_set DISJOINT_EMPTY)
-
-lemma DISJOINT_EMPTY_REFL: "ALL x::'a::type => bool. (x = EMPTY) = DISJOINT x x"
-  by (import pred_set DISJOINT_EMPTY_REFL)
-
-lemma DISJOINT_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
-   DISJOINT (pred_set.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)"
-  by (import pred_set DISJOINT_UNION)
-
-lemma DISJOINT_UNION_BOTH: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
-   DISJOINT (pred_set.UNION s t) u = (DISJOINT s u & DISJOINT t u) &
-   DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)"
-  by (import pred_set DISJOINT_UNION_BOTH)
-
-definition DIFF :: "('a => bool) => ('a => bool) => 'a => bool" where 
-  "DIFF ==
-%(s::'a::type => bool) t::'a::type => bool.
-   GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))"
-
-lemma DIFF_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
-   DIFF s t = GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))"
-  by (import pred_set DIFF_DEF)
-
-lemma IN_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type.
-   IN x (DIFF s t) = (IN x s & ~ IN x t)"
-  by (import pred_set IN_DIFF)
-
-lemma DIFF_EMPTY: "ALL s::'a::type => bool. DIFF s EMPTY = s"
-  by (import pred_set DIFF_EMPTY)
-
-lemma EMPTY_DIFF: "ALL s::'a::type => bool. DIFF EMPTY s = EMPTY"
-  by (import pred_set EMPTY_DIFF)
-
-lemma DIFF_UNIV: "ALL s::'a::type => bool. DIFF s pred_set.UNIV = EMPTY"
-  by (import pred_set DIFF_UNIV)
-
-lemma DIFF_DIFF: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   DIFF (DIFF x xa) xa = DIFF x xa"
-  by (import pred_set DIFF_DIFF)
-
-lemma DIFF_EQ_EMPTY: "ALL x::'a::type => bool. DIFF x x = EMPTY"
-  by (import pred_set DIFF_EQ_EMPTY)
-
-definition INSERT :: "'a => ('a => bool) => 'a => bool" where 
-  "INSERT ==
-%(x::'a::type) s::'a::type => bool.
-   GSPEC (%y::'a::type. (y, y = x | IN y s))"
-
-lemma INSERT_DEF: "ALL (x::'a::type) s::'a::type => bool.
-   INSERT x s = GSPEC (%y::'a::type. (y, y = x | IN y s))"
-  by (import pred_set INSERT_DEF)
-
-lemma IN_INSERT: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
-   IN x (INSERT xa xb) = (x = xa | IN x xb)"
-  by (import pred_set IN_INSERT)
-
-lemma COMPONENT: "ALL (x::'a::type) xa::'a::type => bool. IN x (INSERT x xa)"
-  by (import pred_set COMPONENT)
-
-lemma SET_CASES: "ALL x::'a::type => bool.
-   x = EMPTY |
-   (EX (xa::'a::type) xb::'a::type => bool. x = INSERT xa xb & ~ IN xa xb)"
-  by (import pred_set SET_CASES)
-
-lemma DECOMPOSITION: "ALL (s::'a::type => bool) x::'a::type.
-   IN x s = (EX t::'a::type => bool. s = INSERT x t & ~ IN x t)"
-  by (import pred_set DECOMPOSITION)
-
-lemma ABSORPTION: "ALL (x::'a::type) xa::'a::type => bool. IN x xa = (INSERT x xa = xa)"
-  by (import pred_set ABSORPTION)
-
-lemma INSERT_INSERT: "ALL (x::'a::type) xa::'a::type => bool. INSERT x (INSERT x xa) = INSERT x xa"
-  by (import pred_set INSERT_INSERT)
-
-lemma INSERT_COMM: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
-   INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)"
-  by (import pred_set INSERT_COMM)
-
-lemma INSERT_UNIV: "ALL x::'a::type. INSERT x pred_set.UNIV = pred_set.UNIV"
-  by (import pred_set INSERT_UNIV)
-
-lemma NOT_INSERT_EMPTY: "ALL (x::'a::type) xa::'a::type => bool. INSERT x xa ~= EMPTY"
-  by (import pred_set NOT_INSERT_EMPTY)
-
-lemma NOT_EMPTY_INSERT: "ALL (x::'a::type) xa::'a::type => bool. EMPTY ~= INSERT x xa"
-  by (import pred_set NOT_EMPTY_INSERT)
-
-lemma INSERT_UNION: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
-   pred_set.UNION (INSERT x s) t =
-   (if IN x t then pred_set.UNION s t else INSERT x (pred_set.UNION s t))"
-  by (import pred_set INSERT_UNION)
-
-lemma INSERT_UNION_EQ: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
-   pred_set.UNION (INSERT x s) t = INSERT x (pred_set.UNION s t)"
-  by (import pred_set INSERT_UNION_EQ)
-
-lemma INSERT_INTER: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
-   pred_set.INTER (INSERT x s) t =
-   (if IN x t then INSERT x (pred_set.INTER s t) else pred_set.INTER s t)"
-  by (import pred_set INSERT_INTER)
-
-lemma DISJOINT_INSERT: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
-   DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)"
-  by (import pred_set DISJOINT_INSERT)
-
-lemma INSERT_SUBSET: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
-   SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)"
-  by (import pred_set INSERT_SUBSET)
-
-lemma SUBSET_INSERT: "ALL (x::'a::type) xa::'a::type => bool.
-   ~ IN x xa -->
-   (ALL xb::'a::type => bool. SUBSET xa (INSERT x xb) = SUBSET xa xb)"
-  by (import pred_set SUBSET_INSERT)
-
-lemma INSERT_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type.
-   DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))"
-  by (import pred_set INSERT_DIFF)
-
-definition DELETE :: "('a => bool) => 'a => 'a => bool" where 
-  "DELETE == %(s::'a::type => bool) x::'a::type. DIFF s (INSERT x EMPTY)"
-
-lemma DELETE_DEF: "ALL (s::'a::type => bool) x::'a::type. DELETE s x = DIFF s (INSERT x EMPTY)"
-  by (import pred_set DELETE_DEF)
-
-lemma IN_DELETE: "ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type.
-   IN xa (DELETE x xb) = (IN xa x & xa ~= xb)"
-  by (import pred_set IN_DELETE)
-
-lemma DELETE_NON_ELEMENT: "ALL (x::'a::type) xa::'a::type => bool. (~ IN x xa) = (DELETE xa x = xa)"
-  by (import pred_set DELETE_NON_ELEMENT)
-
-lemma IN_DELETE_EQ: "ALL (s::'a::type => bool) (x::'a::type) x'::'a::type.
-   (IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))"
-  by (import pred_set IN_DELETE_EQ)
-
-lemma EMPTY_DELETE: "ALL x::'a::type. DELETE EMPTY x = EMPTY"
-  by (import pred_set EMPTY_DELETE)
-
-lemma DELETE_DELETE: "ALL (x::'a::type) xa::'a::type => bool. DELETE (DELETE xa x) x = DELETE xa x"
-  by (import pred_set DELETE_DELETE)
-
-lemma DELETE_COMM: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
-   DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x"
-  by (import pred_set DELETE_COMM)
-
-lemma DELETE_SUBSET: "ALL (x::'a::type) xa::'a::type => bool. SUBSET (DELETE xa x) xa"
-  by (import pred_set DELETE_SUBSET)
-
-lemma SUBSET_DELETE: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
-   SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)"
-  by (import pred_set SUBSET_DELETE)
-
-lemma SUBSET_INSERT_DELETE: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
-   SUBSET s (INSERT x t) = SUBSET (DELETE s x) t"
-  by (import pred_set SUBSET_INSERT_DELETE)
-
-lemma DIFF_INSERT: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
-   DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa"
-  by (import pred_set DIFF_INSERT)
-
-lemma PSUBSET_INSERT_SUBSET: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   PSUBSET x xa = (EX xb::'a::type. ~ IN xb x & SUBSET (INSERT xb x) xa)"
-  by (import pred_set PSUBSET_INSERT_SUBSET)
-
-lemma PSUBSET_MEMBER: "ALL (s::'a::type => bool) t::'a::type => bool.
-   PSUBSET s t = (SUBSET s t & (EX y::'a::type. IN y t & ~ IN y s))"
-  by (import pred_set PSUBSET_MEMBER)
-
-lemma DELETE_INSERT: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
-   DELETE (INSERT x xb) xa =
-   (if x = xa then DELETE xb xa else INSERT x (DELETE xb xa))"
-  by (import pred_set DELETE_INSERT)
-
-lemma INSERT_DELETE: "ALL (x::'a::type) xa::'a::type => bool.
-   IN x xa --> INSERT x (DELETE xa x) = xa"
-  by (import pred_set INSERT_DELETE)
-
-lemma DELETE_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
-   pred_set.INTER (DELETE x xb) xa = DELETE (pred_set.INTER x xa) xb"
-  by (import pred_set DELETE_INTER)
-
-lemma DISJOINT_DELETE_SYM: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
-   DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x"
-  by (import pred_set DISJOINT_DELETE_SYM)
+  sorry
+
+lemma EQ_UNIV: "(ALL x. IN x s) = (s = pred_set.UNIV)"
+  sorry
+
+definition
+  SUBSET :: "('a => bool) => ('a => bool) => bool"  where
+  "SUBSET == %s t. ALL x. IN x s --> IN x t"
+
+lemma SUBSET_DEF: "SUBSET s t = (ALL x. IN x s --> IN x t)"
+  sorry
+
+lemma SUBSET_TRANS: "SUBSET x xa & SUBSET xa xb ==> SUBSET x xb"
+  sorry
+
+lemma SUBSET_REFL: "SUBSET x x"
+  sorry
+
+lemma SUBSET_ANTISYM: "SUBSET x xa & SUBSET xa x ==> x = xa"
+  sorry
+
+lemma EMPTY_SUBSET: "SUBSET EMPTY x"
+  sorry
+
+lemma SUBSET_EMPTY: "SUBSET x EMPTY = (x = EMPTY)"
+  sorry
+
+lemma SUBSET_UNIV: "SUBSET x pred_set.UNIV"
+  sorry
+
+lemma UNIV_SUBSET: "SUBSET pred_set.UNIV x = (x = pred_set.UNIV)"
+  sorry
+
+definition
+  PSUBSET :: "('a => bool) => ('a => bool) => bool"  where
+  "PSUBSET == %s t. SUBSET s t & s ~= t"
+
+lemma PSUBSET_DEF: "PSUBSET s t = (SUBSET s t & s ~= t)"
+  sorry
+
+lemma PSUBSET_TRANS: "PSUBSET x xa & PSUBSET xa xb ==> PSUBSET x xb"
+  sorry
+
+lemma PSUBSET_IRREFL: "~ PSUBSET x x"
+  sorry
+
+lemma NOT_PSUBSET_EMPTY: "~ PSUBSET x EMPTY"
+  sorry
+
+lemma NOT_UNIV_PSUBSET: "~ PSUBSET pred_set.UNIV x"
+  sorry
+
+lemma PSUBSET_UNIV: "PSUBSET x pred_set.UNIV = (EX xa. ~ IN xa x)"
+  sorry
+
+definition
+  UNION :: "('a => bool) => ('a => bool) => 'a => bool"  where
+  "UNION == %s t. GSPEC (%x. (x, IN x s | IN x t))"
+
+lemma UNION_DEF: "pred_set.UNION s t = GSPEC (%x. (x, IN x s | IN x t))"
+  sorry
+
+lemma IN_UNION: "IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)"
+  sorry
+
+lemma UNION_ASSOC: "pred_set.UNION x (pred_set.UNION xa xb) =
+pred_set.UNION (pred_set.UNION x xa) xb"
+  sorry
+
+lemma UNION_IDEMPOT: "pred_set.UNION x x = x"
+  sorry
+
+lemma UNION_COMM: "pred_set.UNION x xa = pred_set.UNION xa x"
+  sorry
+
+lemma SUBSET_UNION: "(ALL (x::'a => bool) xa::'a => bool. SUBSET x (pred_set.UNION x xa)) &
+(ALL (x::'a => bool) xa::'a => bool. SUBSET x (pred_set.UNION xa x))"
+  sorry
+
+lemma UNION_SUBSET: "SUBSET (pred_set.UNION s t) u = (SUBSET s u & SUBSET t u)"
+  sorry
+
+lemma SUBSET_UNION_ABSORPTION: "SUBSET x xa = (pred_set.UNION x xa = xa)"
+  sorry
+
+lemma UNION_EMPTY: "(ALL x::'a => bool. pred_set.UNION EMPTY x = x) &
+(ALL x::'a => bool. pred_set.UNION x EMPTY = x)"
+  sorry
+
+lemma UNION_UNIV: "(ALL x::'a => bool. pred_set.UNION pred_set.UNIV x = pred_set.UNIV) &
+(ALL x::'a => bool. pred_set.UNION x pred_set.UNIV = pred_set.UNIV)"
+  sorry
+
+lemma EMPTY_UNION: "(pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)"
+  sorry
+
+definition
+  INTER :: "('a => bool) => ('a => bool) => 'a => bool"  where
+  "INTER == %s t. GSPEC (%x. (x, IN x s & IN x t))"
+
+lemma INTER_DEF: "pred_set.INTER s t = GSPEC (%x. (x, IN x s & IN x t))"
+  sorry
+
+lemma IN_INTER: "IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)"
+  sorry
+
+lemma INTER_ASSOC: "pred_set.INTER x (pred_set.INTER xa xb) =
+pred_set.INTER (pred_set.INTER x xa) xb"
+  sorry
+
+lemma INTER_IDEMPOT: "pred_set.INTER x x = x"
+  sorry
+
+lemma INTER_COMM: "pred_set.INTER x xa = pred_set.INTER xa x"
+  sorry
+
+lemma INTER_SUBSET: "(ALL (x::'a => bool) xa::'a => bool. SUBSET (pred_set.INTER x xa) x) &
+(ALL (x::'a => bool) xa::'a => bool. SUBSET (pred_set.INTER xa x) x)"
+  sorry
+
+lemma SUBSET_INTER: "SUBSET s (pred_set.INTER t u) = (SUBSET s t & SUBSET s u)"
+  sorry
+
+lemma SUBSET_INTER_ABSORPTION: "SUBSET x xa = (pred_set.INTER x xa = x)"
+  sorry
+
+lemma INTER_EMPTY: "(ALL x::'a => bool. pred_set.INTER EMPTY x = EMPTY) &
+(ALL x::'a => bool. pred_set.INTER x EMPTY = EMPTY)"
+  sorry
+
+lemma INTER_UNIV: "(ALL x::'a => bool. pred_set.INTER pred_set.UNIV x = x) &
+(ALL x::'a => bool. pred_set.INTER x pred_set.UNIV = x)"
+  sorry
+
+lemma UNION_OVER_INTER: "pred_set.INTER x (pred_set.UNION xa xb) =
+pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER x xb)"
+  sorry
+
+lemma INTER_OVER_UNION: "pred_set.UNION x (pred_set.INTER xa xb) =
+pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)"
+  sorry
+
+definition
+  DISJOINT :: "('a => bool) => ('a => bool) => bool"  where
+  "DISJOINT == %s t. pred_set.INTER s t = EMPTY"
+
+lemma DISJOINT_DEF: "DISJOINT s t = (pred_set.INTER s t = EMPTY)"
+  sorry
+
+lemma IN_DISJOINT: "DISJOINT x xa = (~ (EX xb. IN xb x & IN xb xa))"
+  sorry
+
+lemma DISJOINT_SYM: "DISJOINT x xa = DISJOINT xa x"
+  sorry
+
+lemma DISJOINT_EMPTY: "DISJOINT EMPTY x & DISJOINT x EMPTY"
+  sorry
+
+lemma DISJOINT_EMPTY_REFL: "(x = EMPTY) = DISJOINT x x"
+  sorry
+
+lemma DISJOINT_UNION: "DISJOINT (pred_set.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)"
+  sorry
+
+lemma DISJOINT_UNION_BOTH: "DISJOINT (pred_set.UNION s t) u = (DISJOINT s u & DISJOINT t u) &
+DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)"
+  sorry
+
+definition
+  DIFF :: "('a => bool) => ('a => bool) => 'a => bool"  where
+  "DIFF == %s t. GSPEC (%x. (x, IN x s & ~ IN x t))"
+
+lemma DIFF_DEF: "DIFF s t = GSPEC (%x. (x, IN x s & ~ IN x t))"
+  sorry
+
+lemma IN_DIFF: "IN x (DIFF s t) = (IN x s & ~ IN x t)"
+  sorry
+
+lemma DIFF_EMPTY: "DIFF s EMPTY = s"
+  sorry
+
+lemma EMPTY_DIFF: "DIFF EMPTY s = EMPTY"
+  sorry
+
+lemma DIFF_UNIV: "DIFF s pred_set.UNIV = EMPTY"
+  sorry
+
+lemma DIFF_DIFF: "DIFF (DIFF x xa) xa = DIFF x xa"
+  sorry
+
+lemma DIFF_EQ_EMPTY: "DIFF x x = EMPTY"
+  sorry
+
+definition
+  INSERT :: "'a => ('a => bool) => 'a => bool"  where
+  "INSERT == %x s. GSPEC (%y. (y, y = x | IN y s))"
+
+lemma INSERT_DEF: "INSERT x s = GSPEC (%y. (y, y = x | IN y s))"
+  sorry
+
+lemma IN_INSERT: "IN x (INSERT xa xb) = (x = xa | IN x xb)"
+  sorry
+
+lemma COMPONENT: "IN x (INSERT x xa)"
+  sorry
+
+lemma SET_CASES: "x = EMPTY | (EX xa xb. x = INSERT xa xb & ~ IN xa xb)"
+  sorry
+
+lemma DECOMPOSITION: "IN x s = (EX t. s = INSERT x t & ~ IN x t)"
+  sorry
+
+lemma ABSORPTION: "IN x xa = (INSERT x xa = xa)"
+  sorry
+
+lemma INSERT_INSERT: "INSERT x (INSERT x xa) = INSERT x xa"
+  sorry
+
+lemma INSERT_COMM: "INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)"
+  sorry
+
+lemma INSERT_UNIV: "INSERT x pred_set.UNIV = pred_set.UNIV"
+  sorry
+
+lemma NOT_INSERT_EMPTY: "INSERT x xa ~= EMPTY"
+  sorry
+
+lemma NOT_EMPTY_INSERT: "EMPTY ~= INSERT x xa"
+  sorry
+
+lemma INSERT_UNION: "pred_set.UNION (INSERT x s) t =
+(if IN x t then pred_set.UNION s t else INSERT x (pred_set.UNION s t))"
+  sorry
+
+lemma INSERT_UNION_EQ: "pred_set.UNION (INSERT x s) t = INSERT x (pred_set.UNION s t)"
+  sorry
+
+lemma INSERT_INTER: "pred_set.INTER (INSERT x s) t =
+(if IN x t then INSERT x (pred_set.INTER s t) else pred_set.INTER s t)"
+  sorry
+
+lemma DISJOINT_INSERT: "DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)"
+  sorry
+
+lemma INSERT_SUBSET: "SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)"
+  sorry
+
+lemma SUBSET_INSERT: "~ IN x xa ==> SUBSET xa (INSERT x xb) = SUBSET xa xb"
+  sorry
+
+lemma INSERT_DIFF: "DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))"
+  sorry
+
+definition
+  DELETE :: "('a => bool) => 'a => 'a => bool"  where
+  "DELETE == %s x. DIFF s (INSERT x EMPTY)"
+
+lemma DELETE_DEF: "DELETE s x = DIFF s (INSERT x EMPTY)"
+  sorry
+
+lemma IN_DELETE: "IN xa (DELETE x xb) = (IN xa x & xa ~= xb)"
+  sorry
+
+lemma DELETE_NON_ELEMENT: "(~ IN x xa) = (DELETE xa x = xa)"
+  sorry
+
+lemma IN_DELETE_EQ: "(IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))"
+  sorry
+
+lemma EMPTY_DELETE: "DELETE EMPTY x = EMPTY"
+  sorry
+
+lemma DELETE_DELETE: "DELETE (DELETE xa x) x = DELETE xa x"
+  sorry
+
+lemma DELETE_COMM: "DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x"
+  sorry
+
+lemma DELETE_SUBSET: "SUBSET (DELETE xa x) xa"
+  sorry
+
+lemma SUBSET_DELETE: "SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)"
+  sorry
+
+lemma SUBSET_INSERT_DELETE: "SUBSET s (INSERT x t) = SUBSET (DELETE s x) t"
+  sorry
+
+lemma DIFF_INSERT: "DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa"
+  sorry
+
+lemma PSUBSET_INSERT_SUBSET: "PSUBSET x xa = (EX xb. ~ IN xb x & SUBSET (INSERT xb x) xa)"
+  sorry
+
+lemma PSUBSET_MEMBER: "PSUBSET s t = (SUBSET s t & (EX y. IN y t & ~ IN y s))"
+  sorry
+
+lemma DELETE_INSERT: "DELETE (INSERT x xb) xa =
+(if x = xa then DELETE xb xa else INSERT x (DELETE xb xa))"
+  sorry
+
+lemma INSERT_DELETE: "IN x xa ==> INSERT x (DELETE xa x) = xa"
+  sorry
+
+lemma DELETE_INTER: "pred_set.INTER (DELETE x xb) xa = DELETE (pred_set.INTER x xa) xb"
+  sorry
+
+lemma DISJOINT_DELETE_SYM: "DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x"
+  sorry
 
 consts
   CHOICE :: "('a => bool) => 'a" 
 
-specification (CHOICE) CHOICE_DEF: "ALL x::'a::type => bool. x ~= EMPTY --> IN (CHOICE x) x"
-  by (import pred_set CHOICE_DEF)
-
-definition REST :: "('a => bool) => 'a => bool" where 
-  "REST == %s::'a::type => bool. DELETE s (CHOICE s)"
-
-lemma REST_DEF: "ALL s::'a::type => bool. REST s = DELETE s (CHOICE s)"
-  by (import pred_set REST_DEF)
-
-lemma CHOICE_NOT_IN_REST: "ALL x::'a::type => bool. ~ IN (CHOICE x) (REST x)"
-  by (import pred_set CHOICE_NOT_IN_REST)
-
-lemma CHOICE_INSERT_REST: "ALL s::'a::type => bool. s ~= EMPTY --> INSERT (CHOICE s) (REST s) = s"
-  by (import pred_set CHOICE_INSERT_REST)
-
-lemma REST_SUBSET: "ALL x::'a::type => bool. SUBSET (REST x) x"
-  by (import pred_set REST_SUBSET)
-
-lemma REST_PSUBSET: "ALL x::'a::type => bool. x ~= EMPTY --> PSUBSET (REST x) x"
-  by (import pred_set REST_PSUBSET)
-
-definition SING :: "('a => bool) => bool" where 
-  "SING == %s::'a::type => bool. EX x::'a::type. s = INSERT x EMPTY"
-
-lemma SING_DEF: "ALL s::'a::type => bool. SING s = (EX x::'a::type. s = INSERT x EMPTY)"
-  by (import pred_set SING_DEF)
-
-lemma SING: "ALL x::'a::type. SING (INSERT x EMPTY)"
-  by (import pred_set SING)
-
-lemma IN_SING: "ALL (x::'a::type) xa::'a::type. IN x (INSERT xa EMPTY) = (x = xa)"
-  by (import pred_set IN_SING)
-
-lemma NOT_SING_EMPTY: "ALL x::'a::type. INSERT x EMPTY ~= EMPTY"
-  by (import pred_set NOT_SING_EMPTY)
-
-lemma NOT_EMPTY_SING: "ALL x::'a::type. EMPTY ~= INSERT x EMPTY"
-  by (import pred_set NOT_EMPTY_SING)
-
-lemma EQUAL_SING: "ALL (x::'a::type) xa::'a::type.
-   (INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)"
-  by (import pred_set EQUAL_SING)
-
-lemma DISJOINT_SING_EMPTY: "ALL x::'a::type. DISJOINT (INSERT x EMPTY) EMPTY"
-  by (import pred_set DISJOINT_SING_EMPTY)
-
-lemma INSERT_SING_UNION: "ALL (x::'a::type => bool) xa::'a::type.
-   INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x"
-  by (import pred_set INSERT_SING_UNION)
-
-lemma SING_DELETE: "ALL x::'a::type. DELETE (INSERT x EMPTY) x = EMPTY"
-  by (import pred_set SING_DELETE)
-
-lemma DELETE_EQ_SING: "ALL (x::'a::type => bool) xa::'a::type.
-   IN xa x --> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)"
-  by (import pred_set DELETE_EQ_SING)
-
-lemma CHOICE_SING: "ALL x::'a::type. CHOICE (INSERT x EMPTY) = x"
-  by (import pred_set CHOICE_SING)
-
-lemma REST_SING: "ALL x::'a::type. REST (INSERT x EMPTY) = EMPTY"
-  by (import pred_set REST_SING)
-
-lemma SING_IFF_EMPTY_REST: "ALL x::'a::type => bool. SING x = (x ~= EMPTY & REST x = EMPTY)"
-  by (import pred_set SING_IFF_EMPTY_REST)
-
-definition IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" where 
-  "IMAGE ==
-%(f::'a::type => 'b::type) s::'a::type => bool.
-   GSPEC (%x::'a::type. (f x, IN x s))"
-
-lemma IMAGE_DEF: "ALL (f::'a::type => 'b::type) s::'a::type => bool.
-   IMAGE f s = GSPEC (%x::'a::type. (f x, IN x s))"
-  by (import pred_set IMAGE_DEF)
-
-lemma IN_IMAGE: "ALL (x::'b::type) (xa::'a::type => bool) xb::'a::type => 'b::type.
-   IN x (IMAGE xb xa) = (EX xc::'a::type. x = xb xc & IN xc xa)"
-  by (import pred_set IN_IMAGE)
-
-lemma IMAGE_IN: "ALL (x::'a::type) xa::'a::type => bool.
-   IN x xa --> (ALL xb::'a::type => 'b::type. IN (xb x) (IMAGE xb xa))"
-  by (import pred_set IMAGE_IN)
-
-lemma IMAGE_EMPTY: "ALL x::'a::type => 'b::type. IMAGE x EMPTY = EMPTY"
-  by (import pred_set IMAGE_EMPTY)
-
-lemma IMAGE_ID: "ALL x::'a::type => bool. IMAGE (%x::'a::type. x) x = x"
-  by (import pred_set IMAGE_ID)
-
-lemma IMAGE_COMPOSE: "ALL (x::'b::type => 'c::type) (xa::'a::type => 'b::type)
-   xb::'a::type => bool. IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)"
-  by (import pred_set IMAGE_COMPOSE)
-
-lemma IMAGE_INSERT: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type => bool.
-   IMAGE x (INSERT xa xb) = INSERT (x xa) (IMAGE x xb)"
-  by (import pred_set IMAGE_INSERT)
-
-lemma IMAGE_EQ_EMPTY: "ALL (s::'a::type => bool) x::'a::type => 'b::type.
-   (IMAGE x s = EMPTY) = (s = EMPTY)"
-  by (import pred_set IMAGE_EQ_EMPTY)
-
-lemma IMAGE_DELETE: "ALL (f::'a::type => 'b::type) (x::'a::type) s::'a::type => bool.
-   ~ IN x s --> IMAGE f (DELETE s x) = IMAGE f s"
-  by (import pred_set IMAGE_DELETE)
-
-lemma IMAGE_UNION: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'a::type => bool.
-   IMAGE x (pred_set.UNION xa xb) = pred_set.UNION (IMAGE x xa) (IMAGE x xb)"
-  by (import pred_set IMAGE_UNION)
-
-lemma IMAGE_SUBSET: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   SUBSET x xa -->
-   (ALL xb::'a::type => 'b::type. SUBSET (IMAGE xb x) (IMAGE xb xa))"
-  by (import pred_set IMAGE_SUBSET)
-
-lemma IMAGE_INTER: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'a::type => bool.
-   SUBSET (IMAGE f (pred_set.INTER s t))
-    (pred_set.INTER (IMAGE f s) (IMAGE f t))"
-  by (import pred_set IMAGE_INTER)
-
-definition INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
+specification (CHOICE) CHOICE_DEF: "ALL x. x ~= EMPTY --> IN (CHOICE x) x"
+  sorry
+
+definition
+  REST :: "('a => bool) => 'a => bool"  where
+  "REST == %s. DELETE s (CHOICE s)"
+
+lemma REST_DEF: "REST s = DELETE s (CHOICE s)"
+  sorry
+
+lemma CHOICE_NOT_IN_REST: "~ IN (CHOICE x) (REST x)"
+  sorry
+
+lemma CHOICE_INSERT_REST: "s ~= EMPTY ==> INSERT (CHOICE s) (REST s) = s"
+  sorry
+
+lemma REST_SUBSET: "SUBSET (REST x) x"
+  sorry
+
+lemma REST_PSUBSET: "x ~= EMPTY ==> PSUBSET (REST x) x"
+  sorry
+
+definition
+  SING :: "('a => bool) => bool"  where
+  "SING == %s. EX x. s = INSERT x EMPTY"
+
+lemma SING_DEF: "SING s = (EX x. s = INSERT x EMPTY)"
+  sorry
+
+lemma SING: "SING (INSERT x EMPTY)"
+  sorry
+
+lemma IN_SING: "IN x (INSERT xa EMPTY) = (x = xa)"
+  sorry
+
+lemma NOT_SING_EMPTY: "INSERT x EMPTY ~= EMPTY"
+  sorry
+
+lemma NOT_EMPTY_SING: "EMPTY ~= INSERT x EMPTY"
+  sorry
+
+lemma EQUAL_SING: "(INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)"
+  sorry
+
+lemma DISJOINT_SING_EMPTY: "DISJOINT (INSERT x EMPTY) EMPTY"
+  sorry
+
+lemma INSERT_SING_UNION: "INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x"
+  sorry
+
+lemma SING_DELETE: "DELETE (INSERT x EMPTY) x = EMPTY"
+  sorry
+
+lemma DELETE_EQ_SING: "IN xa x ==> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)"
+  sorry
+
+lemma CHOICE_SING: "CHOICE (INSERT x EMPTY) = x"
+  sorry
+
+lemma REST_SING: "REST (INSERT x EMPTY) = EMPTY"
+  sorry
+
+lemma SING_IFF_EMPTY_REST: "SING x = (x ~= EMPTY & REST x = EMPTY)"
+  sorry
+
+definition
+  IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool"  where
+  "IMAGE == %f s. GSPEC (%x. (f x, IN x s))"
+
+lemma IMAGE_DEF: "IMAGE (f::'a => 'b) (s::'a => bool) = GSPEC (%x::'a. (f x, IN x s))"
+  sorry
+
+lemma IN_IMAGE: "IN (x::'b) (IMAGE (xb::'a => 'b) (xa::'a => bool)) =
+(EX xc::'a. x = xb xc & IN xc xa)"
+  sorry
+
+lemma IMAGE_IN: "IN x xa ==> IN (xb x) (IMAGE xb xa)"
+  sorry
+
+lemma IMAGE_EMPTY: "IMAGE (x::'a => 'b) EMPTY = EMPTY"
+  sorry
+
+lemma IMAGE_ID: "IMAGE (%x. x) x = x"
+  sorry
+
+lemma IMAGE_COMPOSE: "IMAGE ((x::'b => 'c) o (xa::'a => 'b)) (xb::'a => bool) =
+IMAGE x (IMAGE xa xb)"
+  sorry
+
+lemma IMAGE_INSERT: "IMAGE (x::'a => 'b) (INSERT (xa::'a) (xb::'a => bool)) =
+INSERT (x xa) (IMAGE x xb)"
+  sorry
+
+lemma IMAGE_EQ_EMPTY: "(IMAGE (x::'a => 'b) (s::'a => bool) = EMPTY) = (s = EMPTY)"
+  sorry
+
+lemma IMAGE_DELETE: "~ IN x s ==> IMAGE f (DELETE s x) = IMAGE f s"
+  sorry
+
+lemma IMAGE_UNION: "IMAGE (x::'a => 'b) (pred_set.UNION (xa::'a => bool) (xb::'a => bool)) =
+pred_set.UNION (IMAGE x xa) (IMAGE x xb)"
+  sorry
+
+lemma IMAGE_SUBSET: "SUBSET x xa ==> SUBSET (IMAGE xb x) (IMAGE xb xa)"
+  sorry
+
+lemma IMAGE_INTER: "SUBSET
+ (IMAGE (f::'a => 'b) (pred_set.INTER (s::'a => bool) (t::'a => bool)))
+ (pred_set.INTER (IMAGE f s) (IMAGE f t))"
+  sorry
+
+definition
+  INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool"  where
   "INJ ==
-%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   (ALL x::'a::type. IN x s --> IN (f x) t) &
-   (ALL (x::'a::type) y::'a::type. IN x s & IN y s --> f x = f y --> x = y)"
-
-lemma INJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   INJ f s t =
-   ((ALL x::'a::type. IN x s --> IN (f x) t) &
-    (ALL (x::'a::type) y::'a::type.
-        IN x s & IN y s --> f x = f y --> x = y))"
-  by (import pred_set INJ_DEF)
-
-lemma INJ_ID: "ALL x::'a::type => bool. INJ (%x::'a::type. x) x x"
-  by (import pred_set INJ_ID)
-
-lemma INJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
-   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
-   INJ x xb xc & INJ xa xc xd --> INJ (xa o x) xb xd"
-  by (import pred_set INJ_COMPOSE)
-
-lemma INJ_EMPTY: "ALL x::'a::type => 'b::type.
-   All (INJ x EMPTY) &
-   (ALL xa::'a::type => bool. INJ x xa EMPTY = (xa = EMPTY))"
-  by (import pred_set INJ_EMPTY)
-
-definition SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
+%f s t.
+   (ALL x. IN x s --> IN (f x) t) &
+   (ALL x y. IN x s & IN y s --> f x = f y --> x = y)"
+
+lemma INJ_DEF: "INJ f s t =
+((ALL x. IN x s --> IN (f x) t) &
+ (ALL x y. IN x s & IN y s --> f x = f y --> x = y))"
+  sorry
+
+lemma INJ_ID: "INJ (%x. x) x x"
+  sorry
+
+lemma INJ_COMPOSE: "INJ x xb xc & INJ xa xc xd ==> INJ (xa o x) xb xd"
+  sorry
+
+lemma INJ_EMPTY: "All (INJ (x::'a => 'b) EMPTY) &
+(ALL xa::'a => bool. INJ x xa EMPTY = (xa = EMPTY))"
+  sorry
+
+definition
+  SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool"  where
   "SURJ ==
-%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   (ALL x::'a::type. IN x s --> IN (f x) t) &
-   (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x))"
-
-lemma SURJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   SURJ f s t =
-   ((ALL x::'a::type. IN x s --> IN (f x) t) &
-    (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x)))"
-  by (import pred_set SURJ_DEF)
-
-lemma SURJ_ID: "ALL x::'a::type => bool. SURJ (%x::'a::type. x) x x"
-  by (import pred_set SURJ_ID)
-
-lemma SURJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
-   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
-   SURJ x xb xc & SURJ xa xc xd --> SURJ (xa o x) xb xd"
-  by (import pred_set SURJ_COMPOSE)
-
-lemma SURJ_EMPTY: "ALL x::'a::type => 'b::type.
-   (ALL xa::'b::type => bool. SURJ x EMPTY xa = (xa = EMPTY)) &
-   (ALL xa::'a::type => bool. SURJ x xa EMPTY = (xa = EMPTY))"
-  by (import pred_set SURJ_EMPTY)
-
-lemma IMAGE_SURJ: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'b::type => bool.
-   SURJ x xa xb = (IMAGE x xa = xb)"
-  by (import pred_set IMAGE_SURJ)
-
-definition BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where 
-  "BIJ ==
-%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   INJ f s t & SURJ f s t"
-
-lemma BIJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   BIJ f s t = (INJ f s t & SURJ f s t)"
-  by (import pred_set BIJ_DEF)
-
-lemma BIJ_ID: "ALL x::'a::type => bool. BIJ (%x::'a::type. x) x x"
-  by (import pred_set BIJ_ID)
-
-lemma BIJ_EMPTY: "ALL x::'a::type => 'b::type.
-   (ALL xa::'b::type => bool. BIJ x EMPTY xa = (xa = EMPTY)) &
-   (ALL xa::'a::type => bool. BIJ x xa EMPTY = (xa = EMPTY))"
-  by (import pred_set BIJ_EMPTY)
-
-lemma BIJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
-   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
-   BIJ x xb xc & BIJ xa xc xd --> BIJ (xa o x) xb xd"
-  by (import pred_set BIJ_COMPOSE)
+%f s t.
+   (ALL x. IN x s --> IN (f x) t) &
+   (ALL x. IN x t --> (EX y. IN y s & f y = x))"
+
+lemma SURJ_DEF: "SURJ f s t =
+((ALL x. IN x s --> IN (f x) t) &
+ (ALL x. IN x t --> (EX y. IN y s & f y = x)))"
+  sorry
+
+lemma SURJ_ID: "SURJ (%x. x) x x"
+  sorry
+
+lemma SURJ_COMPOSE: "SURJ x xb xc & SURJ xa xc xd ==> SURJ (xa o x) xb xd"
+  sorry
+
+lemma SURJ_EMPTY: "(ALL xa::'b => bool. SURJ (x::'a => 'b) EMPTY xa = (xa = EMPTY)) &
+(ALL xa::'a => bool. SURJ x xa EMPTY = (xa = EMPTY))"
+  sorry
+
+lemma IMAGE_SURJ: "SURJ x xa xb = (IMAGE x xa = xb)"
+  sorry
+
+definition
+  BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool"  where
+  "BIJ == %f s t. INJ f s t & SURJ f s t"
+
+lemma BIJ_DEF: "BIJ f s t = (INJ f s t & SURJ f s t)"
+  sorry
+
+lemma BIJ_ID: "BIJ (%x. x) x x"
+  sorry
+
+lemma BIJ_EMPTY: "(ALL xa::'b => bool. BIJ (x::'a => 'b) EMPTY xa = (xa = EMPTY)) &
+(ALL xa::'a => bool. BIJ x xa EMPTY = (xa = EMPTY))"
+  sorry
+
+lemma BIJ_COMPOSE: "BIJ x xb xc & BIJ xa xc xd ==> BIJ (xa o x) xb xd"
+  sorry
 
 consts
   LINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
 
-specification (LINV) LINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   INJ f s t --> (ALL x::'a::type. IN x s --> LINV f s (f x) = x)"
-  by (import pred_set LINV_DEF)
+specification (LINV) LINV_DEF: "ALL f s t. INJ f s t --> (ALL x. IN x s --> LINV f s (f x) = x)"
+  sorry
 
 consts
   RINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
 
-specification (RINV) RINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
-   SURJ f s t --> (ALL x::'b::type. IN x t --> f (RINV f s x) = x)"
-  by (import pred_set RINV_DEF)
-
-definition FINITE :: "('a => bool) => bool" where 
+specification (RINV) RINV_DEF: "ALL f s t. SURJ f s t --> (ALL x. IN x t --> f (RINV f s x) = x)"
+  sorry
+
+definition
+  FINITE :: "('a => bool) => bool"  where
   "FINITE ==
-%s::'a::type => bool.
-   ALL P::('a::type => bool) => bool.
-      P EMPTY &
-      (ALL s::'a::type => bool.
-          P s --> (ALL e::'a::type. P (INSERT e s))) -->
-      P s"
-
-lemma FINITE_DEF: "ALL s::'a::type => bool.
-   FINITE s =
-   (ALL P::('a::type => bool) => bool.
-       P EMPTY &
-       (ALL s::'a::type => bool.
-           P s --> (ALL e::'a::type. P (INSERT e s))) -->
-       P s)"
-  by (import pred_set FINITE_DEF)
+%s. ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s"
+
+lemma FINITE_DEF: "FINITE s =
+(ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s)"
+  sorry
 
 lemma FINITE_EMPTY: "FINITE EMPTY"
-  by (import pred_set FINITE_EMPTY)
-
-lemma FINITE_INDUCT: "ALL P::('a::type => bool) => bool.
-   P EMPTY &
-   (ALL s::'a::type => bool.
-       FINITE s & P s -->
-       (ALL e::'a::type. ~ IN e s --> P (INSERT e s))) -->
-   (ALL s::'a::type => bool. FINITE s --> P s)"
-  by (import pred_set FINITE_INDUCT)
-
-lemma FINITE_INSERT: "ALL (x::'a::type) s::'a::type => bool. FINITE (INSERT x s) = FINITE s"
-  by (import pred_set FINITE_INSERT)
-
-lemma FINITE_DELETE: "ALL (x::'a::type) s::'a::type => bool. FINITE (DELETE s x) = FINITE s"
-  by (import pred_set FINITE_DELETE)
-
-lemma FINITE_UNION: "ALL (s::'a::type => bool) t::'a::type => bool.
-   FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)"
-  by (import pred_set FINITE_UNION)
-
-lemma INTER_FINITE: "ALL s::'a::type => bool.
-   FINITE s --> (ALL t::'a::type => bool. FINITE (pred_set.INTER s t))"
-  by (import pred_set INTER_FINITE)
-
-lemma SUBSET_FINITE: "ALL s::'a::type => bool.
-   FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> FINITE t)"
-  by (import pred_set SUBSET_FINITE)
-
-lemma PSUBSET_FINITE: "ALL x::'a::type => bool.
-   FINITE x --> (ALL xa::'a::type => bool. PSUBSET xa x --> FINITE xa)"
-  by (import pred_set PSUBSET_FINITE)
-
-lemma FINITE_DIFF: "ALL s::'a::type => bool.
-   FINITE s --> (ALL t::'a::type => bool. FINITE (DIFF s t))"
-  by (import pred_set FINITE_DIFF)
-
-lemma FINITE_SING: "ALL x::'a::type. FINITE (INSERT x EMPTY)"
-  by (import pred_set FINITE_SING)
-
-lemma SING_FINITE: "ALL x::'a::type => bool. SING x --> FINITE x"
-  by (import pred_set SING_FINITE)
-
-lemma IMAGE_FINITE: "ALL s::'a::type => bool.
-   FINITE s --> (ALL f::'a::type => 'b::type. FINITE (IMAGE f s))"
-  by (import pred_set IMAGE_FINITE)
+  sorry
+
+lemma FINITE_INDUCT: "[| P EMPTY &
+   (ALL s. FINITE s & P s --> (ALL e. ~ IN e s --> P (INSERT e s)));
+   FINITE s |]
+==> P s"
+  sorry
+
+lemma FINITE_INSERT: "FINITE (INSERT x s) = FINITE s"
+  sorry
+
+lemma FINITE_DELETE: "FINITE (DELETE s x) = FINITE s"
+  sorry
+
+lemma FINITE_UNION: "FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)"
+  sorry
+
+lemma INTER_FINITE: "FINITE s ==> FINITE (pred_set.INTER s t)"
+  sorry
+
+lemma SUBSET_FINITE: "[| FINITE s; SUBSET t s |] ==> FINITE t"
+  sorry
+
+lemma PSUBSET_FINITE: "[| FINITE x; PSUBSET xa x |] ==> FINITE xa"
+  sorry
+
+lemma FINITE_DIFF: "FINITE s ==> FINITE (DIFF s t)"
+  sorry
+
+lemma FINITE_SING: "FINITE (INSERT x EMPTY)"
+  sorry
+
+lemma SING_FINITE: "SING x ==> FINITE x"
+  sorry
+
+lemma IMAGE_FINITE: "FINITE s ==> FINITE (IMAGE f s)"
+  sorry
 
 consts
   CARD :: "('a => bool) => nat" 
@@ -4077,77 +3113,56 @@
                  ((CARD::('a::type => bool) => nat) s)
                  ((Suc::nat => nat)
                    ((CARD::('a::type => bool) => nat) s)))))))"
-  by (import pred_set CARD_DEF)
+  sorry
 
 lemma CARD_EMPTY: "CARD EMPTY = 0"
-  by (import pred_set CARD_EMPTY)
-
-lemma CARD_INSERT: "ALL s::'a::type => bool.
-   FINITE s -->
-   (ALL x::'a::type.
-       CARD (INSERT x s) = (if IN x s then CARD s else Suc (CARD s)))"
-  by (import pred_set CARD_INSERT)
-
-lemma CARD_EQ_0: "ALL s::'a::type => bool. FINITE s --> (CARD s = 0) = (s = EMPTY)"
-  by (import pred_set CARD_EQ_0)
-
-lemma CARD_DELETE: "ALL s::'a::type => bool.
-   FINITE s -->
-   (ALL x::'a::type.
-       CARD (DELETE s x) = (if IN x s then CARD s - 1 else CARD s))"
-  by (import pred_set CARD_DELETE)
-
-lemma CARD_INTER_LESS_EQ: "ALL s::'a::type => bool.
-   FINITE s -->
-   (ALL t::'a::type => bool. CARD (pred_set.INTER s t) <= CARD s)"
-  by (import pred_set CARD_INTER_LESS_EQ)
-
-lemma CARD_UNION: "ALL s::'a::type => bool.
-   FINITE s -->
-   (ALL t::'a::type => bool.
-       FINITE t -->
-       CARD (pred_set.UNION s t) + CARD (pred_set.INTER s t) =
-       CARD s + CARD t)"
-  by (import pred_set CARD_UNION)
-
-lemma CARD_SUBSET: "ALL s::'a::type => bool.
-   FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> CARD t <= CARD s)"
-  by (import pred_set CARD_SUBSET)
-
-lemma CARD_PSUBSET: "ALL s::'a::type => bool.
-   FINITE s --> (ALL t::'a::type => bool. PSUBSET t s --> CARD t < CARD s)"
-  by (import pred_set CARD_PSUBSET)
-
-lemma CARD_SING: "ALL x::'a::type. CARD (INSERT x EMPTY) = 1"
-  by (import pred_set CARD_SING)
-
-lemma SING_IFF_CARD1: "ALL x::'a::type => bool. SING x = (CARD x = 1 & FINITE x)"
-  by (import pred_set SING_IFF_CARD1)
-
-lemma CARD_DIFF: "ALL t::'a::type => bool.
-   FINITE t -->
-   (ALL s::'a::type => bool.
-       FINITE s --> CARD (DIFF s t) = CARD s - CARD (pred_set.INTER s t))"
-  by (import pred_set CARD_DIFF)
-
-lemma LESS_CARD_DIFF: "ALL t::'a::type => bool.
-   FINITE t -->
-   (ALL s::'a::type => bool.
-       FINITE s --> CARD t < CARD s --> 0 < CARD (DIFF s t))"
-  by (import pred_set LESS_CARD_DIFF)
-
-lemma FINITE_COMPLETE_INDUCTION: "ALL P::('a::type => bool) => bool.
-   (ALL x::'a::type => bool.
-       (ALL y::'a::type => bool. PSUBSET y x --> P y) -->
-       FINITE x --> P x) -->
-   (ALL x::'a::type => bool. FINITE x --> P x)"
-  by (import pred_set FINITE_COMPLETE_INDUCTION)
-
-definition INFINITE :: "('a => bool) => bool" where 
-  "INFINITE == %s::'a::type => bool. ~ FINITE s"
-
-lemma INFINITE_DEF: "ALL s::'a::type => bool. INFINITE s = (~ FINITE s)"
-  by (import pred_set INFINITE_DEF)
+  sorry
+
+lemma CARD_INSERT: "FINITE s ==> CARD (INSERT x s) = (if IN x s then CARD s else Suc (CARD s))"
+  sorry
+
+lemma CARD_EQ_0: "FINITE s ==> (CARD s = 0) = (s = EMPTY)"
+  sorry
+
+lemma CARD_DELETE: "FINITE s ==> CARD (DELETE s x) = (if IN x s then CARD s - 1 else CARD s)"
+  sorry
+
+lemma CARD_INTER_LESS_EQ: "FINITE s ==> CARD (pred_set.INTER s t) <= CARD s"
+  sorry
+
+lemma CARD_UNION: "[| FINITE s; FINITE t |]
+==> CARD (pred_set.UNION s t) + CARD (pred_set.INTER s t) = CARD s + CARD t"
+  sorry
+
+lemma CARD_SUBSET: "[| FINITE s; SUBSET t s |] ==> CARD t <= CARD s"
+  sorry
+
+lemma CARD_PSUBSET: "[| FINITE s; PSUBSET t s |] ==> CARD t < CARD s"
+  sorry
+
+lemma CARD_SING: "CARD (INSERT x EMPTY) = 1"
+  sorry
+
+lemma SING_IFF_CARD1: "SING x = (CARD x = 1 & FINITE x)"
+  sorry
+
+lemma CARD_DIFF: "[| FINITE t; FINITE s |]
+==> CARD (DIFF s t) = CARD s - CARD (pred_set.INTER s t)"
+  sorry
+
+lemma LESS_CARD_DIFF: "[| FINITE t; FINITE s; CARD t < CARD s |] ==> 0 < CARD (DIFF s t)"
+  sorry
+
+lemma FINITE_COMPLETE_INDUCTION: "[| !!x. [| !!y. PSUBSET y x ==> P y; FINITE x |] ==> P x; FINITE x |]
+==> P x"
+  sorry
+
+definition
+  INFINITE :: "('a => bool) => bool"  where
+  "INFINITE == %s. ~ FINITE s"
+
+lemma INFINITE_DEF: "INFINITE s = (~ FINITE s)"
+  sorry
 
 lemma NOT_IN_FINITE: "(op =::bool => bool => bool)
  ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
@@ -4159,23 +3174,19 @@
           (%x::'a::type.
               (Not::bool => bool)
                ((IN::'a::type => ('a::type => bool) => bool) x s)))))"
-  by (import pred_set NOT_IN_FINITE)
-
-lemma INFINITE_INHAB: "ALL x::'a::type => bool. INFINITE x --> (EX xa::'a::type. IN xa x)"
-  by (import pred_set INFINITE_INHAB)
-
-lemma IMAGE_11_INFINITE: "ALL f::'a::type => 'b::type.
-   (ALL (x::'a::type) y::'a::type. f x = f y --> x = y) -->
-   (ALL s::'a::type => bool. INFINITE s --> INFINITE (IMAGE f s))"
-  by (import pred_set IMAGE_11_INFINITE)
-
-lemma INFINITE_SUBSET: "ALL x::'a::type => bool.
-   INFINITE x --> (ALL xa::'a::type => bool. SUBSET x xa --> INFINITE xa)"
-  by (import pred_set INFINITE_SUBSET)
-
-lemma IN_INFINITE_NOT_FINITE: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   INFINITE x & FINITE xa --> (EX xb::'a::type. IN xb x & ~ IN xb xa)"
-  by (import pred_set IN_INFINITE_NOT_FINITE)
+  sorry
+
+lemma INFINITE_INHAB: "INFINITE x ==> EX xa. IN xa x"
+  sorry
+
+lemma IMAGE_11_INFINITE: "[| !!x y. f x = f y ==> x = y; INFINITE s |] ==> INFINITE (IMAGE f s)"
+  sorry
+
+lemma INFINITE_SUBSET: "[| INFINITE x; SUBSET x xa |] ==> INFINITE xa"
+  sorry
+
+lemma IN_INFINITE_NOT_FINITE: "INFINITE x & FINITE xa ==> EX xb. IN xb x & ~ IN xb xa"
+  sorry
 
 lemma INFINITE_UNIV: "(op =::bool => bool => bool)
  ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
@@ -4193,14 +3204,11 @@
           (%y::'a::type.
               (All::('a::type => bool) => bool)
                (%x::'a::type.
-                   (Not::bool => bool)
-                    ((op =::'a::type => 'a::type => bool) (f x) y))))))"
-  by (import pred_set INFINITE_UNIV)
-
-lemma FINITE_PSUBSET_INFINITE: "ALL x::'a::type => bool.
-   INFINITE x =
-   (ALL xa::'a::type => bool. FINITE xa --> SUBSET xa x --> PSUBSET xa x)"
-  by (import pred_set FINITE_PSUBSET_INFINITE)
+                   (op ~=::'a::type => 'a::type => bool) (f x) y)))))"
+  sorry
+
+lemma FINITE_PSUBSET_INFINITE: "INFINITE x = (ALL xa. FINITE xa --> SUBSET xa x --> PSUBSET xa x)"
+  sorry
 
 lemma FINITE_PSUBSET_UNIV: "(op =::bool => bool => bool)
  ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
@@ -4210,362 +3218,283 @@
         ((FINITE::('a::type => bool) => bool) s)
         ((PSUBSET::('a::type => bool) => ('a::type => bool) => bool) s
           (pred_set.UNIV::'a::type => bool))))"
-  by (import pred_set FINITE_PSUBSET_UNIV)
-
-lemma INFINITE_DIFF_FINITE: "ALL (s::'a::type => bool) t::'a::type => bool.
-   INFINITE s & FINITE t --> DIFF s t ~= EMPTY"
-  by (import pred_set INFINITE_DIFF_FINITE)
-
-lemma FINITE_ISO_NUM: "ALL s::'a::type => bool.
-   FINITE s -->
-   (EX f::nat => 'a::type.
-       (ALL (n::nat) m::nat.
-           n < CARD s & m < CARD s --> f n = f m --> n = m) &
-       s = GSPEC (%n::nat. (f n, n < CARD s)))"
-  by (import pred_set FINITE_ISO_NUM)
-
-lemma FINITE_WEAK_ENUMERATE: "(All::(('a::type => bool) => bool) => bool)
- (%x::'a::type => bool.
-     (op =::bool => bool => bool) ((FINITE::('a::type => bool) => bool) x)
-      ((Ex::((nat => 'a::type) => bool) => bool)
-        (%f::nat => 'a::type.
-            (Ex::(nat => bool) => bool)
-             (%b::nat.
-                 (All::('a::type => bool) => bool)
-                  (%e::'a::type.
-                      (op =::bool => bool => bool)
-                       ((IN::'a::type => ('a::type => bool) => bool) e x)
-                       ((Ex::(nat => bool) => bool)
-                         (%n::nat.
-                             (op &::bool => bool => bool)
-                              ((op <::nat => nat => bool) n b)
-                              ((op =::'a::type => 'a::type => bool) e
-                                (f n)))))))))"
-  by (import pred_set FINITE_WEAK_ENUMERATE)
-
-definition BIGUNION :: "(('a => bool) => bool) => 'a => bool" where 
-  "BIGUNION ==
-%P::('a::type => bool) => bool.
-   GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))"
-
-lemma BIGUNION: "ALL P::('a::type => bool) => bool.
-   BIGUNION P =
-   GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))"
-  by (import pred_set BIGUNION)
-
-lemma IN_BIGUNION: "ALL (x::'a::type) xa::('a::type => bool) => bool.
-   IN x (BIGUNION xa) = (EX s::'a::type => bool. IN x s & IN s xa)"
-  by (import pred_set IN_BIGUNION)
+  sorry
+
+lemma INFINITE_DIFF_FINITE: "INFINITE s & FINITE t ==> DIFF s t ~= EMPTY"
+  sorry
+
+lemma FINITE_ISO_NUM: "FINITE s
+==> EX f. (ALL n m. n < CARD s & m < CARD s --> f n = f m --> n = m) &
+          s = GSPEC (%n. (f n, n < CARD s))"
+  sorry
+
+lemma FINITE_WEAK_ENUMERATE: "FINITE (x::'a => bool) =
+(EX (f::nat => 'a) b::nat. ALL e::'a. IN e x = (EX n<b. e = f n))"
+  sorry
+
+definition
+  BIGUNION :: "(('a => bool) => bool) => 'a => bool"  where
+  "BIGUNION == %P. GSPEC (%x. (x, EX p. IN p P & IN x p))"
+
+lemma BIGUNION: "BIGUNION P = GSPEC (%x. (x, EX p. IN p P & IN x p))"
+  sorry
+
+lemma IN_BIGUNION: "IN x (BIGUNION xa) = (EX s. IN x s & IN s xa)"
+  sorry
 
 lemma BIGUNION_EMPTY: "BIGUNION EMPTY = EMPTY"
-  by (import pred_set BIGUNION_EMPTY)
-
-lemma BIGUNION_SING: "ALL x::'a::type => bool. BIGUNION (INSERT x EMPTY) = x"
-  by (import pred_set BIGUNION_SING)
-
-lemma BIGUNION_UNION: "ALL (x::('a::type => bool) => bool) xa::('a::type => bool) => bool.
-   BIGUNION (pred_set.UNION x xa) =
-   pred_set.UNION (BIGUNION x) (BIGUNION xa)"
-  by (import pred_set BIGUNION_UNION)
-
-lemma DISJOINT_BIGUNION: "(ALL (s::('a::type => bool) => bool) t::'a::type => bool.
+  sorry
+
+lemma BIGUNION_SING: "BIGUNION (INSERT x EMPTY) = x"
+  sorry
+
+lemma BIGUNION_UNION: "BIGUNION (pred_set.UNION x xa) = pred_set.UNION (BIGUNION x) (BIGUNION xa)"
+  sorry
+
+lemma DISJOINT_BIGUNION: "(ALL (s::('a => bool) => bool) t::'a => bool.
     DISJOINT (BIGUNION s) t =
-    (ALL s'::'a::type => bool. IN s' s --> DISJOINT s' t)) &
-(ALL (x::('a::type => bool) => bool) xa::'a::type => bool.
+    (ALL s'::'a => bool. IN s' s --> DISJOINT s' t)) &
+(ALL (x::('a => bool) => bool) xa::'a => bool.
     DISJOINT xa (BIGUNION x) =
-    (ALL xb::'a::type => bool. IN xb x --> DISJOINT xa xb))"
-  by (import pred_set DISJOINT_BIGUNION)
-
-lemma BIGUNION_INSERT: "ALL (x::'a::type => bool) xa::('a::type => bool) => bool.
-   BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)"
-  by (import pred_set BIGUNION_INSERT)
-
-lemma BIGUNION_SUBSET: "ALL (X::'a::type => bool) P::('a::type => bool) => bool.
-   SUBSET (BIGUNION P) X = (ALL Y::'a::type => bool. IN Y P --> SUBSET Y X)"
-  by (import pred_set BIGUNION_SUBSET)
-
-lemma FINITE_BIGUNION: "ALL x::('a::type => bool) => bool.
-   FINITE x & (ALL s::'a::type => bool. IN s x --> FINITE s) -->
-   FINITE (BIGUNION x)"
-  by (import pred_set FINITE_BIGUNION)
-
-definition BIGINTER :: "(('a => bool) => bool) => 'a => bool" where 
-  "BIGINTER ==
-%B::('a::type => bool) => bool.
-   GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))"
-
-lemma BIGINTER: "ALL B::('a::type => bool) => bool.
-   BIGINTER B =
-   GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))"
-  by (import pred_set BIGINTER)
-
-lemma IN_BIGINTER: "IN (x::'a::type) (BIGINTER (B::('a::type => bool) => bool)) =
-(ALL P::'a::type => bool. IN P B --> IN x P)"
-  by (import pred_set IN_BIGINTER)
-
-lemma BIGINTER_INSERT: "ALL (P::'a::type => bool) B::('a::type => bool) => bool.
-   BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)"
-  by (import pred_set BIGINTER_INSERT)
+    (ALL xb::'a => bool. IN xb x --> DISJOINT xa xb))"
+  sorry
+
+lemma BIGUNION_INSERT: "BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)"
+  sorry
+
+lemma BIGUNION_SUBSET: "SUBSET (BIGUNION P) X = (ALL Y. IN Y P --> SUBSET Y X)"
+  sorry
+
+lemma FINITE_BIGUNION: "FINITE x & (ALL s. IN s x --> FINITE s) ==> FINITE (BIGUNION x)"
+  sorry
+
+definition
+  BIGINTER :: "(('a => bool) => bool) => 'a => bool"  where
+  "BIGINTER == %B. GSPEC (%x. (x, ALL P. IN P B --> IN x P))"
+
+lemma BIGINTER: "BIGINTER B = GSPEC (%x. (x, ALL P. IN P B --> IN x P))"
+  sorry
+
+lemma IN_BIGINTER: "IN x (BIGINTER B) = (ALL P. IN P B --> IN x P)"
+  sorry
+
+lemma BIGINTER_INSERT: "BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)"
+  sorry
 
 lemma BIGINTER_EMPTY: "BIGINTER EMPTY = pred_set.UNIV"
-  by (import pred_set BIGINTER_EMPTY)
-
-lemma BIGINTER_INTER: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa"
-  by (import pred_set BIGINTER_INTER)
-
-lemma BIGINTER_SING: "ALL x::'a::type => bool. BIGINTER (INSERT x EMPTY) = x"
-  by (import pred_set BIGINTER_SING)
-
-lemma SUBSET_BIGINTER: "ALL (X::'a::type => bool) P::('a::type => bool) => bool.
-   SUBSET X (BIGINTER P) = (ALL x::'a::type => bool. IN x P --> SUBSET X x)"
-  by (import pred_set SUBSET_BIGINTER)
-
-lemma DISJOINT_BIGINTER: "ALL (x::'a::type => bool) (xa::'a::type => bool)
-   xb::('a::type => bool) => bool.
-   IN xa xb & DISJOINT xa x -->
-   DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x"
-  by (import pred_set DISJOINT_BIGINTER)
-
-definition CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" where 
-  "CROSS ==
-%(P::'a::type => bool) Q::'b::type => bool.
-   GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))"
-
-lemma CROSS_DEF: "ALL (P::'a::type => bool) Q::'b::type => bool.
-   CROSS P Q =
-   GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))"
-  by (import pred_set CROSS_DEF)
-
-lemma IN_CROSS: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type * 'b::type.
-   IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)"
-  by (import pred_set IN_CROSS)
-
-lemma CROSS_EMPTY: "ALL x::'a::type => bool. CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY"
-  by (import pred_set CROSS_EMPTY)
-
-lemma CROSS_INSERT_LEFT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type.
-   CROSS (INSERT xb x) xa =
-   pred_set.UNION (CROSS (INSERT xb EMPTY) xa) (CROSS x xa)"
-  by (import pred_set CROSS_INSERT_LEFT)
-
-lemma CROSS_INSERT_RIGHT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'b::type.
-   CROSS x (INSERT xb xa) =
-   pred_set.UNION (CROSS x (INSERT xb EMPTY)) (CROSS x xa)"
-  by (import pred_set CROSS_INSERT_RIGHT)
-
-lemma FINITE_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool.
-   FINITE x & FINITE xa --> FINITE (CROSS x xa)"
-  by (import pred_set FINITE_CROSS)
-
-lemma CROSS_SINGS: "ALL (x::'a::type) xa::'b::type.
-   CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY"
-  by (import pred_set CROSS_SINGS)
-
-lemma CARD_SING_CROSS: "ALL (x::'a::type) s::'b::type => bool.
-   FINITE s --> CARD (CROSS (INSERT x EMPTY) s) = CARD s"
-  by (import pred_set CARD_SING_CROSS)
-
-lemma CARD_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool.
-   FINITE x & FINITE xa --> CARD (CROSS x xa) = CARD x * CARD xa"
-  by (import pred_set CARD_CROSS)
-
-lemma CROSS_SUBSET: "ALL (x::'a::type => bool) (xa::'b::type => bool) (xb::'a::type => bool)
-   xc::'b::type => bool.
-   SUBSET (CROSS xb xc) (CROSS x xa) =
-   (xb = EMPTY | xc = EMPTY | SUBSET xb x & SUBSET xc xa)"
-  by (import pred_set CROSS_SUBSET)
-
-lemma FINITE_CROSS_EQ: "ALL (P::'a::type => bool) Q::'b::type => bool.
-   FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)"
-  by (import pred_set FINITE_CROSS_EQ)
-
-definition COMPL :: "('a => bool) => 'a => bool" where 
+  sorry
+
+lemma BIGINTER_INTER: "BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa"
+  sorry
+
+lemma BIGINTER_SING: "BIGINTER (INSERT x EMPTY) = x"
+  sorry
+
+lemma SUBSET_BIGINTER: "SUBSET X (BIGINTER P) = (ALL x. IN x P --> SUBSET X x)"
+  sorry
+
+lemma DISJOINT_BIGINTER: "IN xa xb & DISJOINT xa x
+==> DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x"
+  sorry
+
+definition
+  CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool"  where
+  "CROSS == %P Q. GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))"
+
+lemma CROSS_DEF: "CROSS P Q = GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))"
+  sorry
+
+lemma IN_CROSS: "IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)"
+  sorry
+
+lemma CROSS_EMPTY: "CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY"
+  sorry
+
+lemma CROSS_INSERT_LEFT: "CROSS (INSERT xb x) xa =
+pred_set.UNION (CROSS (INSERT xb EMPTY) xa) (CROSS x xa)"
+  sorry
+
+lemma CROSS_INSERT_RIGHT: "CROSS x (INSERT xb xa) =
+pred_set.UNION (CROSS x (INSERT xb EMPTY)) (CROSS x xa)"
+  sorry
+
+lemma FINITE_CROSS: "FINITE x & FINITE xa ==> FINITE (CROSS x xa)"
+  sorry
+
+lemma CROSS_SINGS: "CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY"
+  sorry
+
+lemma CARD_SING_CROSS: "FINITE (s::'b => bool) ==> CARD (CROSS (INSERT (x::'a) EMPTY) s) = CARD s"
+  sorry
+
+lemma CARD_CROSS: "FINITE x & FINITE xa ==> CARD (CROSS x xa) = CARD x * CARD xa"
+  sorry
+
+lemma CROSS_SUBSET: "SUBSET (CROSS xb xc) (CROSS x xa) =
+(xb = EMPTY | xc = EMPTY | SUBSET xb x & SUBSET xc xa)"
+  sorry
+
+lemma FINITE_CROSS_EQ: "FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)"
+  sorry
+
+definition
+  COMPL :: "('a => bool) => 'a => bool"  where
   "COMPL == DIFF pred_set.UNIV"
 
-lemma COMPL_DEF: "ALL P::'a::type => bool. COMPL P = DIFF pred_set.UNIV P"
-  by (import pred_set COMPL_DEF)
-
-lemma IN_COMPL: "ALL (x::'a::type) xa::'a::type => bool. IN x (COMPL xa) = (~ IN x xa)"
-  by (import pred_set IN_COMPL)
-
-lemma COMPL_COMPL: "ALL x::'a::type => bool. COMPL (COMPL x) = x"
-  by (import pred_set COMPL_COMPL)
-
-lemma COMPL_CLAUSES: "ALL x::'a::type => bool.
-   pred_set.INTER (COMPL x) x = EMPTY &
-   pred_set.UNION (COMPL x) x = pred_set.UNIV"
-  by (import pred_set COMPL_CLAUSES)
-
-lemma COMPL_SPLITS: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER (COMPL x) xa) = xa"
-  by (import pred_set COMPL_SPLITS)
-
-lemma INTER_UNION_COMPL: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))"
-  by (import pred_set INTER_UNION_COMPL)
+lemma COMPL_DEF: "COMPL P = DIFF pred_set.UNIV P"
+  sorry
+
+lemma IN_COMPL: "IN x (COMPL xa) = (~ IN x xa)"
+  sorry
+
+lemma COMPL_COMPL: "COMPL (COMPL x) = x"
+  sorry
+
+lemma COMPL_CLAUSES: "pred_set.INTER (COMPL x) x = EMPTY &
+pred_set.UNION (COMPL x) x = pred_set.UNIV"
+  sorry
+
+lemma COMPL_SPLITS: "pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER (COMPL x) xa) = xa"
+  sorry
+
+lemma INTER_UNION_COMPL: "pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))"
+  sorry
 
 lemma COMPL_EMPTY: "COMPL EMPTY = pred_set.UNIV"
-  by (import pred_set COMPL_EMPTY)
+  sorry
 
 consts
   count :: "nat => nat => bool" 
 
 defs
-  count_primdef: "count == %n::nat. GSPEC (%m::nat. (m, m < n))"
-
-lemma count_def: "ALL n::nat. count n = GSPEC (%m::nat. (m, m < n))"
-  by (import pred_set count_def)
-
-lemma IN_COUNT: "ALL (m::nat) n::nat. IN m (count n) = (m < n)"
-  by (import pred_set IN_COUNT)
+  count_primdef: "count == %n. GSPEC (%m. (m, m < n))"
+
+lemma count_def: "count n = GSPEC (%m. (m, m < n))"
+  sorry
+
+lemma IN_COUNT: "IN m (count n) = (m < n)"
+  sorry
 
 lemma COUNT_ZERO: "count 0 = EMPTY"
-  by (import pred_set COUNT_ZERO)
-
-lemma COUNT_SUC: "ALL n::nat. count (Suc n) = INSERT n (count n)"
-  by (import pred_set COUNT_SUC)
-
-lemma FINITE_COUNT: "ALL n::nat. FINITE (count n)"
-  by (import pred_set FINITE_COUNT)
-
-lemma CARD_COUNT: "ALL n::nat. CARD (count n) = n"
-  by (import pred_set CARD_COUNT)
-
-definition ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" where 
+  sorry
+
+lemma COUNT_SUC: "count (Suc n) = INSERT n (count n)"
+  sorry
+
+lemma FINITE_COUNT: "FINITE (count n)"
+  sorry
+
+lemma CARD_COUNT: "CARD (count n) = n"
+  sorry
+
+definition
+  ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b"  where
   "ITSET_tupled ==
-%f::'a::type => 'b::type => 'b::type.
-   WFREC
-    (SOME R::('a::type => bool) * 'b::type
-             => ('a::type => bool) * 'b::type => bool.
-        WF R &
-        (ALL (b::'b::type) s::'a::type => bool.
-            FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
-    (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type)
-        (v::'a::type => bool, v1::'b::type).
-        if FINITE v
-        then if v = EMPTY then v1
-             else ITSET_tupled (REST v, f (CHOICE v) v1)
-        else ARB)"
-
-lemma ITSET_tupled_primitive_def: "ALL f::'a::type => 'b::type => 'b::type.
-   ITSET_tupled f =
-   WFREC
-    (SOME R::('a::type => bool) * 'b::type
-             => ('a::type => bool) * 'b::type => bool.
-        WF R &
-        (ALL (b::'b::type) s::'a::type => bool.
-            FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
-    (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type)
-        (v::'a::type => bool, v1::'b::type).
-        if FINITE v
-        then if v = EMPTY then v1
-             else ITSET_tupled (REST v, f (CHOICE v) v1)
-        else ARB)"
-  by (import pred_set ITSET_tupled_primitive_def)
-
-definition ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" where 
-  "ITSET ==
-%(f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) x1::'b::type.
-   ITSET_tupled f (x, x1)"
-
-lemma ITSET_curried_def: "ALL (f::'a::type => 'b::type => 'b::type) (x::'a::type => bool)
-   x1::'b::type. ITSET f x x1 = ITSET_tupled f (x, x1)"
-  by (import pred_set ITSET_curried_def)
-
-lemma ITSET_IND: "ALL P::('a::type => bool) => 'b::type => bool.
-   (ALL (s::'a::type => bool) b::'b::type.
-       (FINITE s & s ~= EMPTY -->
-        P (REST s) ((f::'a::type => 'b::type => 'b::type) (CHOICE s) b)) -->
-       P s b) -->
-   (ALL v::'a::type => bool. All (P v))"
-  by (import pred_set ITSET_IND)
-
-lemma ITSET_THM: "ALL (s::'a::type => bool) (f::'a::type => 'b::type => 'b::type) b::'b::type.
-   FINITE s -->
-   ITSET f s b =
-   (if s = EMPTY then b else ITSET f (REST s) (f (CHOICE s) b))"
-  by (import pred_set ITSET_THM)
-
-lemma ITSET_EMPTY: "ALL (x::'a::type => 'b::type => 'b::type) xa::'b::type.
-   ITSET x EMPTY xa = xa"
-  by (import pred_set ITSET_EMPTY)
+%f. WFREC
+     (SOME R.
+         WF R &
+         (ALL b s.
+             FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
+     (%ITSET_tupled (v, v1).
+         if FINITE v
+         then if v = EMPTY then v1
+              else ITSET_tupled (REST v, f (CHOICE v) v1)
+         else ARB)"
+
+lemma ITSET_tupled_primitive_def: "ITSET_tupled f =
+WFREC
+ (SOME R.
+     WF R &
+     (ALL b s. FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
+ (%ITSET_tupled (v, v1).
+     if FINITE v
+     then if v = EMPTY then v1 else ITSET_tupled (REST v, f (CHOICE v) v1)
+     else ARB)"
+  sorry
+
+definition
+  ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b"  where
+  "ITSET == %f x x1. ITSET_tupled f (x, x1)"
+
+lemma ITSET_curried_def: "ITSET (f::'a => 'b => 'b) (x::'a => bool) (x1::'b) = ITSET_tupled f (x, x1)"
+  sorry
+
+lemma ITSET_IND: "(!!(s::'a => bool) b::'b.
+    (FINITE s & s ~= EMPTY
+     ==> (P::('a => bool) => 'b => bool) (REST s)
+          ((f::'a => 'b => 'b) (CHOICE s) b))
+    ==> P s b)
+==> P (v::'a => bool) (x::'b)"
+  sorry
+
+lemma ITSET_THM: "FINITE s
+==> ITSET f s b =
+    (if s = EMPTY then b else ITSET f (REST s) (f (CHOICE s) b))"
+  sorry
+
+lemma ITSET_EMPTY: "ITSET (x::'a => 'b => 'b) EMPTY (xa::'b) = xa"
+  sorry
 
 ;end_setup
 
 ;setup_theory operator
 
-definition ASSOC :: "('a => 'a => 'a) => bool" where 
-  "ASSOC ==
-%f::'a::type => 'a::type => 'a::type.
-   ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z"
-
-lemma ASSOC_DEF: "ALL f::'a::type => 'a::type => 'a::type.
-   ASSOC f =
-   (ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z)"
-  by (import operator ASSOC_DEF)
-
-definition COMM :: "('a => 'a => 'b) => bool" where 
-  "COMM ==
-%f::'a::type => 'a::type => 'b::type.
-   ALL (x::'a::type) y::'a::type. f x y = f y x"
-
-lemma COMM_DEF: "ALL f::'a::type => 'a::type => 'b::type.
-   COMM f = (ALL (x::'a::type) y::'a::type. f x y = f y x)"
-  by (import operator COMM_DEF)
-
-definition FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" where 
-  "FCOMM ==
-%(f::'a::type => 'b::type => 'a::type) g::'c::type => 'a::type => 'a::type.
-   ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z"
-
-lemma FCOMM_DEF: "ALL (f::'a::type => 'b::type => 'a::type)
-   g::'c::type => 'a::type => 'a::type.
-   FCOMM f g =
-   (ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z)"
-  by (import operator FCOMM_DEF)
-
-definition RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" where 
-  "RIGHT_ID ==
-%(f::'a::type => 'b::type => 'a::type) e::'b::type.
-   ALL x::'a::type. f x e = x"
-
-lemma RIGHT_ID_DEF: "ALL (f::'a::type => 'b::type => 'a::type) e::'b::type.
-   RIGHT_ID f e = (ALL x::'a::type. f x e = x)"
-  by (import operator RIGHT_ID_DEF)
-
-definition LEFT_ID :: "('a => 'b => 'b) => 'a => bool" where 
-  "LEFT_ID ==
-%(f::'a::type => 'b::type => 'b::type) e::'a::type.
-   ALL x::'b::type. f e x = x"
-
-lemma LEFT_ID_DEF: "ALL (f::'a::type => 'b::type => 'b::type) e::'a::type.
-   LEFT_ID f e = (ALL x::'b::type. f e x = x)"
-  by (import operator LEFT_ID_DEF)
-
-definition MONOID :: "('a => 'a => 'a) => 'a => bool" where 
-  "MONOID ==
-%(f::'a::type => 'a::type => 'a::type) e::'a::type.
-   ASSOC f & RIGHT_ID f e & LEFT_ID f e"
-
-lemma MONOID_DEF: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type.
-   MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)"
-  by (import operator MONOID_DEF)
+definition
+  ASSOC :: "('a => 'a => 'a) => bool"  where
+  "ASSOC == %f. ALL x y z. f x (f y z) = f (f x y) z"
+
+lemma ASSOC_DEF: "ASSOC f = (ALL x y z. f x (f y z) = f (f x y) z)"
+  sorry
+
+definition
+  COMM :: "('a => 'a => 'b) => bool"  where
+  "COMM == %f. ALL x y. f x y = f y x"
+
+lemma COMM_DEF: "COMM f = (ALL x y. f x y = f y x)"
+  sorry
+
+definition
+  FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool"  where
+  "FCOMM == %f g. ALL x y z. g x (f y z) = f (g x y) z"
+
+lemma FCOMM_DEF: "FCOMM f g = (ALL x y z. g x (f y z) = f (g x y) z)"
+  sorry
+
+definition
+  RIGHT_ID :: "('a => 'b => 'a) => 'b => bool"  where
+  "RIGHT_ID == %f e. ALL x. f x e = x"
+
+lemma RIGHT_ID_DEF: "RIGHT_ID f e = (ALL x. f x e = x)"
+  sorry
+
+definition
+  LEFT_ID :: "('a => 'b => 'b) => 'a => bool"  where
+  "LEFT_ID == %f e. ALL x. f e x = x"
+
+lemma LEFT_ID_DEF: "LEFT_ID f e = (ALL x. f e x = x)"
+  sorry
+
+definition
+  MONOID :: "('a => 'a => 'a) => 'a => bool"  where
+  "MONOID == %f e. ASSOC f & RIGHT_ID f e & LEFT_ID f e"
+
+lemma MONOID_DEF: "MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)"
+  sorry
 
 lemma ASSOC_CONJ: "ASSOC op &"
-  by (import operator ASSOC_CONJ)
+  sorry
 
 lemma ASSOC_DISJ: "ASSOC op |"
-  by (import operator ASSOC_DISJ)
-
-lemma FCOMM_ASSOC: "ALL x::'a::type => 'a::type => 'a::type. FCOMM x x = ASSOC x"
-  by (import operator FCOMM_ASSOC)
+  sorry
+
+lemma FCOMM_ASSOC: "FCOMM x x = ASSOC x"
+  sorry
 
 lemma MONOID_CONJ_T: "MONOID op & True"
-  by (import operator MONOID_CONJ_T)
+  sorry
 
 lemma MONOID_DISJ_F: "MONOID op | False"
-  by (import operator MONOID_DISJ_F)
+  sorry
 
 ;end_setup
 
@@ -4574,1371 +3503,995 @@
 consts
   SNOC :: "'a => 'a list => 'a list" 
 
-specification (SNOC) SNOC: "(ALL x::'a::type. SNOC x [] = [x]) &
-(ALL (x::'a::type) (x'::'a::type) l::'a::type list.
-    SNOC x (x' # l) = x' # SNOC x l)"
-  by (import rich_list SNOC)
+specification (SNOC) SNOC: "(ALL x::'a. SNOC x [] = [x]) &
+(ALL (x::'a) (x'::'a) l::'a list. SNOC x (x' # l) = x' # SNOC x l)"
+  sorry
 
 consts
   SCANL :: "('b => 'a => 'b) => 'b => 'a list => 'b list" 
 
-specification (SCANL) SCANL: "(ALL (f::'b::type => 'a::type => 'b::type) e::'b::type.
-    SCANL f e [] = [e]) &
-(ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type)
-    l::'a::type list. SCANL f e (x # l) = e # SCANL f (f e x) l)"
-  by (import rich_list SCANL)
+specification (SCANL) SCANL: "(ALL (f::'b => 'a => 'b) e::'b. SCANL f e [] = [e]) &
+(ALL (f::'b => 'a => 'b) (e::'b) (x::'a) l::'a list.
+    SCANL f e (x # l) = e # SCANL f (f e x) l)"
+  sorry
 
 consts
   SCANR :: "('a => 'b => 'b) => 'b => 'a list => 'b list" 
 
-specification (SCANR) SCANR: "(ALL (f::'a::type => 'b::type => 'b::type) e::'b::type.
-    SCANR f e [] = [e]) &
-(ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type)
-    l::'a::type list.
+specification (SCANR) SCANR: "(ALL (f::'a => 'b => 'b) e::'b. SCANR f e [] = [e]) &
+(ALL (f::'a => 'b => 'b) (e::'b) (x::'a) l::'a list.
     SCANR f e (x # l) = f x (hd (SCANR f e l)) # SCANR f e l)"
-  by (import rich_list SCANR)
-
-lemma IS_EL_DEF: "ALL (x::'a::type) l::'a::type list. x mem l = list_ex (op = x) l"
-  by (import rich_list IS_EL_DEF)
-
-definition AND_EL :: "bool list => bool" where 
+  sorry
+
+lemma IS_EL_DEF: "List.member l x = list_ex (op = x) l"
+  sorry
+
+definition
+  AND_EL :: "bool list => bool"  where
   "AND_EL == list_all I"
 
 lemma AND_EL_DEF: "AND_EL = list_all I"
-  by (import rich_list AND_EL_DEF)
-
-definition OR_EL :: "bool list => bool" where 
+  sorry
+
+definition
+  OR_EL :: "bool list => bool"  where
   "OR_EL == list_ex I"
 
 lemma OR_EL_DEF: "OR_EL = list_ex I"
-  by (import rich_list OR_EL_DEF)
+  sorry
 
 consts
   FIRSTN :: "nat => 'a list => 'a list" 
 
-specification (FIRSTN) FIRSTN: "(ALL l::'a::type list. FIRSTN 0 l = []) &
-(ALL (n::nat) (x::'a::type) l::'a::type list.
-    FIRSTN (Suc n) (x # l) = x # FIRSTN n l)"
-  by (import rich_list FIRSTN)
+specification (FIRSTN) FIRSTN: "(ALL l::'a list. FIRSTN (0::nat) l = []) &
+(ALL (n::nat) (x::'a) l::'a list. FIRSTN (Suc n) (x # l) = x # FIRSTN n l)"
+  sorry
 
 consts
   BUTFIRSTN :: "nat => 'a list => 'a list" 
 
-specification (BUTFIRSTN) BUTFIRSTN: "(ALL l::'a::type list. BUTFIRSTN 0 l = l) &
-(ALL (n::nat) (x::'a::type) l::'a::type list.
-    BUTFIRSTN (Suc n) (x # l) = BUTFIRSTN n l)"
-  by (import rich_list BUTFIRSTN)
+specification (BUTFIRSTN) BUTFIRSTN: "(ALL l::'a list. BUTFIRSTN (0::nat) l = l) &
+(ALL (n::nat) (x::'a) l::'a list. BUTFIRSTN (Suc n) (x # l) = BUTFIRSTN n l)"
+  sorry
 
 consts
   SEG :: "nat => nat => 'a list => 'a list" 
 
-specification (SEG) SEG: "(ALL (k::nat) l::'a::type list. SEG 0 k l = []) &
-(ALL (m::nat) (x::'a::type) l::'a::type list.
-    SEG (Suc m) 0 (x # l) = x # SEG m 0 l) &
-(ALL (m::nat) (k::nat) (x::'a::type) l::'a::type list.
+specification (SEG) SEG: "(ALL (k::nat) l::'a list. SEG (0::nat) k l = []) &
+(ALL (m::nat) (x::'a) l::'a list.
+    SEG (Suc m) (0::nat) (x # l) = x # SEG m (0::nat) l) &
+(ALL (m::nat) (k::nat) (x::'a) l::'a list.
     SEG (Suc m) (Suc k) (x # l) = SEG (Suc m) k l)"
-  by (import rich_list SEG)
-
-lemma LAST: "ALL (x::'a::type) l::'a::type list. last (SNOC x l) = x"
-  by (import rich_list LAST)
-
-lemma BUTLAST: "ALL (x::'a::type) l::'a::type list. butlast (SNOC x l) = l"
-  by (import rich_list BUTLAST)
+  sorry
+
+lemma LAST: "last (SNOC x l) = x"
+  sorry
+
+lemma BUTLAST: "butlast (SNOC x l) = l"
+  sorry
 
 consts
   LASTN :: "nat => 'a list => 'a list" 
 
-specification (LASTN) LASTN: "(ALL l::'a::type list. LASTN 0 l = []) &
-(ALL (n::nat) (x::'a::type) l::'a::type list.
+specification (LASTN) LASTN: "(ALL l::'a list. LASTN (0::nat) l = []) &
+(ALL (n::nat) (x::'a) l::'a list.
     LASTN (Suc n) (SNOC x l) = SNOC x (LASTN n l))"
-  by (import rich_list LASTN)
+  sorry
 
 consts
   BUTLASTN :: "nat => 'a list => 'a list" 
 
-specification (BUTLASTN) BUTLASTN: "(ALL l::'a::type list. BUTLASTN 0 l = l) &
-(ALL (n::nat) (x::'a::type) l::'a::type list.
+specification (BUTLASTN) BUTLASTN: "(ALL l::'a list. BUTLASTN (0::nat) l = l) &
+(ALL (n::nat) (x::'a) l::'a list.
     BUTLASTN (Suc n) (SNOC x l) = BUTLASTN n l)"
-  by (import rich_list BUTLASTN)
-
-lemma EL: "(ALL x::'a::type list. EL 0 x = hd x) &
-(ALL (x::nat) xa::'a::type list. EL (Suc x) xa = EL x (tl xa))"
-  by (import rich_list EL)
+  sorry
+
+lemma EL: "(ALL x::'a list. EL (0::nat) x = hd x) &
+(ALL (x::nat) xa::'a list. EL (Suc x) xa = EL x (tl xa))"
+  sorry
 
 consts
   ELL :: "nat => 'a list => 'a" 
 
-specification (ELL) ELL: "(ALL l::'a::type list. ELL 0 l = last l) &
-(ALL (n::nat) l::'a::type list. ELL (Suc n) l = ELL n (butlast l))"
-  by (import rich_list ELL)
+specification (ELL) ELL: "(ALL l::'a list. ELL (0::nat) l = last l) &
+(ALL (n::nat) l::'a list. ELL (Suc n) l = ELL n (butlast l))"
+  sorry
 
 consts
   IS_PREFIX :: "'a list => 'a list => bool" 
 
-specification (IS_PREFIX) IS_PREFIX: "(ALL l::'a::type list. IS_PREFIX l [] = True) &
-(ALL (x::'a::type) l::'a::type list. IS_PREFIX [] (x # l) = False) &
-(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
+specification (IS_PREFIX) IS_PREFIX: "(ALL l::'a list. IS_PREFIX l [] = True) &
+(ALL (x::'a) l::'a list. IS_PREFIX [] (x # l) = False) &
+(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list.
     IS_PREFIX (x1 # l1) (x2 # l2) = (x1 = x2 & IS_PREFIX l1 l2))"
-  by (import rich_list IS_PREFIX)
-
-lemma SNOC_APPEND: "ALL (x::'a::type) l::'a::type list. SNOC x l = l @ [x]"
-  by (import rich_list SNOC_APPEND)
-
-lemma REVERSE: "rev [] = [] &
-(ALL (x::'a::type) xa::'a::type list. rev (x # xa) = SNOC x (rev xa))"
-  by (import rich_list REVERSE)
-
-lemma REVERSE_SNOC: "ALL (x::'a::type) l::'a::type list. rev (SNOC x l) = x # rev l"
-  by (import rich_list REVERSE_SNOC)
-
-lemma SNOC_Axiom: "ALL (e::'b::type) f::'a::type => 'a::type list => 'b::type => 'b::type.
-   EX x::'a::type list => 'b::type.
-      x [] = e &
-      (ALL (xa::'a::type) l::'a::type list. x (SNOC xa l) = f xa l (x l))"
-  by (import rich_list SNOC_Axiom)
+  sorry
+
+lemma SNOC_APPEND: "SNOC x l = l @ [x]"
+  sorry
+
+lemma REVERSE: "rev [] = [] & (ALL (x::'a) xa::'a list. rev (x # xa) = SNOC x (rev xa))"
+  sorry
+
+lemma REVERSE_SNOC: "rev (SNOC x l) = x # rev l"
+  sorry
+
+lemma SNOC_Axiom: "EX x. x [] = e & (ALL xa l. x (SNOC xa l) = f xa l (x l))"
+  sorry
 
 consts
   IS_SUFFIX :: "'a list => 'a list => bool" 
 
-specification (IS_SUFFIX) IS_SUFFIX: "(ALL l::'a::type list. IS_SUFFIX l [] = True) &
-(ALL (x::'a::type) l::'a::type list. IS_SUFFIX [] (SNOC x l) = False) &
-(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
+specification (IS_SUFFIX) IS_SUFFIX: "(ALL l::'a list. IS_SUFFIX l [] = True) &
+(ALL (x::'a) l::'a list. IS_SUFFIX [] (SNOC x l) = False) &
+(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list.
     IS_SUFFIX (SNOC x1 l1) (SNOC x2 l2) = (x1 = x2 & IS_SUFFIX l1 l2))"
-  by (import rich_list IS_SUFFIX)
+  sorry
 
 consts
   IS_SUBLIST :: "'a list => 'a list => bool" 
 
-specification (IS_SUBLIST) IS_SUBLIST: "(ALL l::'a::type list. IS_SUBLIST l [] = True) &
-(ALL (x::'a::type) l::'a::type list. IS_SUBLIST [] (x # l) = False) &
-(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
+specification (IS_SUBLIST) IS_SUBLIST: "(ALL l::'a list. IS_SUBLIST l [] = True) &
+(ALL (x::'a) l::'a list. IS_SUBLIST [] (x # l) = False) &
+(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list.
     IS_SUBLIST (x1 # l1) (x2 # l2) =
     (x1 = x2 & IS_PREFIX l1 l2 | IS_SUBLIST l1 (x2 # l2)))"
-  by (import rich_list IS_SUBLIST)
+  sorry
 
 consts
   SPLITP :: "('a => bool) => 'a list => 'a list * 'a list" 
 
-specification (SPLITP) SPLITP: "(ALL P::'a::type => bool. SPLITP P [] = ([], [])) &
-(ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
+specification (SPLITP) SPLITP: "(ALL P::'a => bool. SPLITP P [] = ([], [])) &
+(ALL (P::'a => bool) (x::'a) l::'a list.
     SPLITP P (x # l) =
     (if P x then ([], x # l) else (x # fst (SPLITP P l), snd (SPLITP P l))))"
-  by (import rich_list SPLITP)
-
-definition PREFIX :: "('a => bool) => 'a list => 'a list" where 
-  "PREFIX == %(P::'a::type => bool) l::'a::type list. fst (SPLITP (Not o P) l)"
-
-lemma PREFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list.
-   PREFIX P l = fst (SPLITP (Not o P) l)"
-  by (import rich_list PREFIX_DEF)
-
-definition SUFFIX :: "('a => bool) => 'a list => 'a list" where 
-  "SUFFIX ==
-%P::'a::type => bool.
-   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else [])
-    []"
-
-lemma SUFFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list.
-   SUFFIX P l =
-   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else [])
-    [] l"
-  by (import rich_list SUFFIX_DEF)
-
-definition UNZIP_FST :: "('a * 'b) list => 'a list" where 
-  "UNZIP_FST == %l::('a::type * 'b::type) list. fst (unzip l)"
-
-lemma UNZIP_FST_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_FST l = fst (unzip l)"
-  by (import rich_list UNZIP_FST_DEF)
-
-definition UNZIP_SND :: "('a * 'b) list => 'b list" where 
-  "UNZIP_SND == %l::('a::type * 'b::type) list. snd (unzip l)"
-
-lemma UNZIP_SND_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_SND l = snd (unzip l)"
-  by (import rich_list UNZIP_SND_DEF)
+  sorry
+
+definition
+  PREFIX :: "('a => bool) => 'a list => 'a list"  where
+  "PREFIX == %P l. fst (SPLITP (Not o P) l)"
+
+lemma PREFIX_DEF: "PREFIX P l = fst (SPLITP (Not o P) l)"
+  sorry
+
+definition
+  SUFFIX :: "('a => bool) => 'a list => 'a list"  where
+  "SUFFIX == %P. foldl (%l' x. if P x then SNOC x l' else []) []"
+
+lemma SUFFIX_DEF: "SUFFIX P l = foldl (%l' x. if P x then SNOC x l' else []) [] l"
+  sorry
+
+definition
+  UNZIP_FST :: "('a * 'b) list => 'a list"  where
+  "UNZIP_FST == %l. fst (unzip l)"
+
+lemma UNZIP_FST_DEF: "UNZIP_FST l = fst (unzip l)"
+  sorry
+
+definition
+  UNZIP_SND :: "('a * 'b) list => 'b list"  where
+  "UNZIP_SND == %l. snd (unzip l)"
+
+lemma UNZIP_SND_DEF: "UNZIP_SND (l::('a * 'b) list) = snd (unzip l)"
+  sorry
 
 consts
   GENLIST :: "(nat => 'a) => nat => 'a list" 
 
-specification (GENLIST) GENLIST: "(ALL f::nat => 'a::type. GENLIST f 0 = []) &
-(ALL (f::nat => 'a::type) n::nat.
-    GENLIST f (Suc n) = SNOC (f n) (GENLIST f n))"
-  by (import rich_list GENLIST)
+specification (GENLIST) GENLIST: "(ALL f::nat => 'a. GENLIST f (0::nat) = []) &
+(ALL (f::nat => 'a) n::nat. GENLIST f (Suc n) = SNOC (f n) (GENLIST f n))"
+  sorry
 
 consts
   REPLICATE :: "nat => 'a => 'a list" 
 
-specification (REPLICATE) REPLICATE: "(ALL x::'a::type. REPLICATE 0 x = []) &
-(ALL (n::nat) x::'a::type. REPLICATE (Suc n) x = x # REPLICATE n x)"
-  by (import rich_list REPLICATE)
-
-lemma LENGTH_MAP2: "ALL (l1::'a::type list) l2::'b::type list.
-   length l1 = length l2 -->
-   (ALL f::'a::type => 'b::type => 'c::type.
-       length (map2 f l1 l2) = length l1 &
-       length (map2 f l1 l2) = length l2)"
-  by (import rich_list LENGTH_MAP2)
-
-lemma NULL_EQ_NIL: "ALL l::'a::type list. null l = (l = [])"
-  by (import rich_list NULL_EQ_NIL)
-
-lemma LENGTH_EQ: "ALL (x::'a::type list) y::'a::type list. x = y --> length x = length y"
-  by (import rich_list LENGTH_EQ)
-
-lemma LENGTH_NOT_NULL: "ALL l::'a::type list. (0 < length l) = (~ null l)"
-  by (import rich_list LENGTH_NOT_NULL)
-
-lemma SNOC_INDUCT: "ALL P::'a::type list => bool.
-   P [] &
-   (ALL l::'a::type list. P l --> (ALL x::'a::type. P (SNOC x l))) -->
-   All P"
-  by (import rich_list SNOC_INDUCT)
-
-lemma SNOC_CASES: "ALL x'::'a::type list.
-   x' = [] | (EX (x::'a::type) l::'a::type list. x' = SNOC x l)"
-  by (import rich_list SNOC_CASES)
-
-lemma LENGTH_SNOC: "ALL (x::'a::type) l::'a::type list. length (SNOC x l) = Suc (length l)"
-  by (import rich_list LENGTH_SNOC)
-
-lemma NOT_NIL_SNOC: "ALL (x::'a::type) xa::'a::type list. [] ~= SNOC x xa"
-  by (import rich_list NOT_NIL_SNOC)
-
-lemma NOT_SNOC_NIL: "ALL (x::'a::type) xa::'a::type list. SNOC x xa ~= []"
-  by (import rich_list NOT_SNOC_NIL)
-
-lemma SNOC_11: "ALL (x::'a::type) (l::'a::type list) (x'::'a::type) l'::'a::type list.
-   (SNOC x l = SNOC x' l') = (x = x' & l = l')"
-  by (import rich_list SNOC_11)
-
-lemma SNOC_EQ_LENGTH_EQ: "ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
-   SNOC x1 l1 = SNOC x2 l2 --> length l1 = length l2"
-  by (import rich_list SNOC_EQ_LENGTH_EQ)
-
-lemma SNOC_REVERSE_CONS: "ALL (x::'a::type) xa::'a::type list. SNOC x xa = rev (x # rev xa)"
-  by (import rich_list SNOC_REVERSE_CONS)
-
-lemma MAP_SNOC: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type list.
-   map x (SNOC xa xb) = SNOC (x xa) (map x xb)"
-  by (import rich_list MAP_SNOC)
-
-lemma FOLDR_SNOC: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type)
-   l::'a::type list. foldr f (SNOC x l) e = foldr f l (f x e)"
-  by (import rich_list FOLDR_SNOC)
-
-lemma FOLDL_SNOC: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type)
-   l::'a::type list. foldl f e (SNOC x l) = f (foldl f e l) x"
-  by (import rich_list FOLDL_SNOC)
-
-lemma FOLDR_FOLDL: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type.
-   MONOID f e --> (ALL l::'a::type list. foldr f l e = foldl f e l)"
-  by (import rich_list FOLDR_FOLDL)
-
-lemma LENGTH_FOLDR: "ALL l::'a::type list. length l = foldr (%x::'a::type. Suc) l 0"
-  by (import rich_list LENGTH_FOLDR)
-
-lemma LENGTH_FOLDL: "ALL l::'a::type list. length l = foldl (%(l'::nat) x::'a::type. Suc l') 0 l"
-  by (import rich_list LENGTH_FOLDL)
-
-lemma MAP_FOLDR: "ALL (f::'a::type => 'b::type) l::'a::type list.
-   map f l = foldr (%x::'a::type. op # (f x)) l []"
-  by (import rich_list MAP_FOLDR)
-
-lemma MAP_FOLDL: "ALL (f::'a::type => 'b::type) l::'a::type list.
-   map f l = foldl (%(l'::'b::type list) x::'a::type. SNOC (f x) l') [] l"
-  by (import rich_list MAP_FOLDL)
-
-lemma MAP_o: "ALL (f::'b::type => 'c::type) g::'a::type => 'b::type.
-   map (f o g) = map f o map g"
-  by (import rich_list MAP_o)
-
-lemma FILTER_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
-   filter P l =
-   foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else l') l []"
-  by (import rich_list FILTER_FOLDR)
-
-lemma FILTER_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
-   filter P (SNOC x l) = (if P x then SNOC x (filter P l) else filter P l)"
-  by (import rich_list FILTER_SNOC)
-
-lemma FILTER_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
-   filter P l =
-   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else l')
-    [] l"
-  by (import rich_list FILTER_FOLDL)
-
-lemma FILTER_COMM: "ALL (f1::'a::type => bool) (f2::'a::type => bool) l::'a::type list.
-   filter f1 (filter f2 l) = filter f2 (filter f1 l)"
-  by (import rich_list FILTER_COMM)
-
-lemma FILTER_IDEM: "ALL (f::'a::type => bool) l::'a::type list.
-   filter f (filter f l) = filter f l"
-  by (import rich_list FILTER_IDEM)
-
-lemma LENGTH_SEG: "ALL (n::nat) (k::nat) l::'a::type list.
-   n + k <= length l --> length (SEG n k l) = n"
-  by (import rich_list LENGTH_SEG)
-
-lemma APPEND_NIL: "(ALL l::'a::type list. l @ [] = l) & (ALL x::'a::type list. [] @ x = x)"
-  by (import rich_list APPEND_NIL)
-
-lemma APPEND_SNOC: "ALL (l1::'a::type list) (x::'a::type) l2::'a::type list.
-   l1 @ SNOC x l2 = SNOC x (l1 @ l2)"
-  by (import rich_list APPEND_SNOC)
-
-lemma APPEND_FOLDR: "ALL (l1::'a::type list) l2::'a::type list. l1 @ l2 = foldr op # l1 l2"
-  by (import rich_list APPEND_FOLDR)
-
-lemma APPEND_FOLDL: "ALL (l1::'a::type list) l2::'a::type list.
-   l1 @ l2 = foldl (%(l'::'a::type list) x::'a::type. SNOC x l') l1 l2"
-  by (import rich_list APPEND_FOLDL)
-
-lemma CONS_APPEND: "ALL (x::'a::type) l::'a::type list. x # l = [x] @ l"
-  by (import rich_list CONS_APPEND)
+specification (REPLICATE) REPLICATE: "(ALL x::'a. REPLICATE (0::nat) x = []) &
+(ALL (n::nat) x::'a. REPLICATE (Suc n) x = x # REPLICATE n x)"
+  sorry
+
+lemma LENGTH_MAP2: "length l1 = length l2
+==> length (map2 f l1 l2) = length l1 & length (map2 f l1 l2) = length l2"
+  sorry
+
+lemma LENGTH_EQ: "x = y ==> length x = length y"
+  sorry
+
+lemma LENGTH_NOT_NULL: "(0 < length l) = (~ List.null l)"
+  sorry
+
+lemma SNOC_INDUCT: "P [] & (ALL l. P l --> (ALL x. P (SNOC x l))) ==> P x"
+  sorry
+
+lemma SNOC_CASES: "x' = [] | (EX x l. x' = SNOC x l)"
+  sorry
+
+lemma LENGTH_SNOC: "length (SNOC x l) = Suc (length l)"
+  sorry
+
+lemma NOT_NIL_SNOC: "[] ~= SNOC x xa"
+  sorry
+
+lemma NOT_SNOC_NIL: "SNOC x xa ~= []"
+  sorry
+
+lemma SNOC_11: "(SNOC x l = SNOC x' l') = (x = x' & l = l')"
+  sorry
+
+lemma SNOC_EQ_LENGTH_EQ: "SNOC x1 l1 = SNOC x2 l2 ==> length l1 = length l2"
+  sorry
+
+lemma SNOC_REVERSE_CONS: "SNOC x xa = rev (x # rev xa)"
+  sorry
+
+lemma MAP_SNOC: "map (x::'a => 'b) (SNOC (xa::'a) (xb::'a list)) = SNOC (x xa) (map x xb)"
+  sorry
+
+lemma FOLDR_SNOC: "foldr (f::'a => 'b => 'b) (SNOC (x::'a) (l::'a list)) (e::'b) =
+foldr f l (f x e)"
+  sorry
+
+lemma FOLDL_SNOC: "foldl (f::'b => 'a => 'b) (e::'b) (SNOC (x::'a) (l::'a list)) =
+f (foldl f e l) x"
+  sorry
+
+lemma FOLDR_FOLDL: "MONOID f e ==> foldr f l e = foldl f e l"
+  sorry
+
+lemma LENGTH_FOLDR: "length l = foldr (%x. Suc) l 0"
+  sorry
+
+lemma LENGTH_FOLDL: "length l = foldl (%l' x. Suc l') 0 l"
+  sorry
+
+lemma MAP_FOLDR: "map (f::'a => 'b) (l::'a list) = foldr (%x::'a. op # (f x)) l []"
+  sorry
+
+lemma MAP_FOLDL: "map (f::'a => 'b) (l::'a list) =
+foldl (%(l'::'b list) x::'a. SNOC (f x) l') [] l"
+  sorry
+
+lemma FILTER_FOLDR: "filter P l = foldr (%x l'. if P x then x # l' else l') l []"
+  sorry
+
+lemma FILTER_SNOC: "filter P (SNOC x l) = (if P x then SNOC x (filter P l) else filter P l)"
+  sorry
+
+lemma FILTER_FOLDL: "filter P l = foldl (%l' x. if P x then SNOC x l' else l') [] l"
+  sorry
+
+lemma FILTER_COMM: "filter f1 (filter f2 l) = filter f2 (filter f1 l)"
+  sorry
+
+lemma FILTER_IDEM: "filter f (filter f l) = filter f l"
+  sorry
+
+lemma LENGTH_SEG: "n + k <= length l ==> length (SEG n k l) = n"
+  sorry
+
+lemma APPEND_NIL: "(ALL l::'a list. l @ [] = l) & (ALL x::'a list. [] @ x = x)"
+  sorry
+
+lemma APPEND_SNOC: "l1 @ SNOC x l2 = SNOC x (l1 @ l2)"
+  sorry
+
+lemma APPEND_FOLDR: "l1 @ l2 = foldr op # l1 l2"
+  sorry
+
+lemma APPEND_FOLDL: "l1 @ l2 = foldl (%l' x. SNOC x l') l1 l2"
+  sorry
+
+lemma CONS_APPEND: "x # l = [x] @ l"
+  sorry
 
 lemma ASSOC_APPEND: "ASSOC op @"
-  by (import rich_list ASSOC_APPEND)
+  sorry
 
 lemma MONOID_APPEND_NIL: "MONOID op @ []"
-  by (import rich_list MONOID_APPEND_NIL)
-
-lemma APPEND_LENGTH_EQ: "ALL (l1::'a::type list) l1'::'a::type list.
-   length l1 = length l1' -->
-   (ALL (l2::'a::type list) l2'::'a::type list.
-       length l2 = length l2' -->
-       (l1 @ l2 = l1' @ l2') = (l1 = l1' & l2 = l2'))"
-  by (import rich_list APPEND_LENGTH_EQ)
-
-lemma FLAT_SNOC: "ALL (x::'a::type list) l::'a::type list list.
-   concat (SNOC x l) = concat l @ x"
-  by (import rich_list FLAT_SNOC)
-
-lemma FLAT_FOLDR: "ALL l::'a::type list list. concat l = foldr op @ l []"
-  by (import rich_list FLAT_FOLDR)
-
-lemma FLAT_FOLDL: "ALL l::'a::type list list. concat l = foldl op @ [] l"
-  by (import rich_list FLAT_FOLDL)
-
-lemma LENGTH_FLAT: "ALL l::'a::type list list. length (concat l) = sum (map size l)"
-  by (import rich_list LENGTH_FLAT)
-
-lemma REVERSE_FOLDR: "ALL l::'a::type list. rev l = foldr SNOC l []"
-  by (import rich_list REVERSE_FOLDR)
-
-lemma REVERSE_FOLDL: "ALL l::'a::type list.
-   rev l = foldl (%(l'::'a::type list) x::'a::type. x # l') [] l"
-  by (import rich_list REVERSE_FOLDL)
-
-lemma ALL_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
-   list_all P (SNOC x l) = (list_all P l & P x)"
-  by (import rich_list ALL_EL_SNOC)
-
-lemma ALL_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list.
-   list_all P (map f l) = list_all (P o f) l"
-  by (import rich_list ALL_EL_MAP)
-
-lemma SOME_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
-   list_ex P (SNOC x l) = (P x | list_ex P l)"
-  by (import rich_list SOME_EL_SNOC)
-
-lemma IS_EL_SNOC: "ALL (y::'a::type) (x::'a::type) l::'a::type list.
-   y mem SNOC x l = (y = x | y mem l)"
-  by (import rich_list IS_EL_SNOC)
-
-lemma SUM_SNOC: "ALL (x::nat) l::nat list. sum (SNOC x l) = sum l + x"
-  by (import rich_list SUM_SNOC)
-
-lemma SUM_FOLDL: "ALL l::nat list. sum l = foldl op + 0 l"
-  by (import rich_list SUM_FOLDL)
-
-lemma IS_PREFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_PREFIX l1 l2 = (EX l::'a::type list. l1 = l2 @ l)"
-  by (import rich_list IS_PREFIX_APPEND)
-
-lemma IS_SUFFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_SUFFIX l1 l2 = (EX l::'a::type list. l1 = l @ l2)"
-  by (import rich_list IS_SUFFIX_APPEND)
-
-lemma IS_SUBLIST_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_SUBLIST l1 l2 =
-   (EX (l::'a::type list) l'::'a::type list. l1 = l @ l2 @ l')"
-  by (import rich_list IS_SUBLIST_APPEND)
-
-lemma IS_PREFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_PREFIX l1 l2 --> IS_SUBLIST l1 l2"
-  by (import rich_list IS_PREFIX_IS_SUBLIST)
-
-lemma IS_SUFFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_SUFFIX l1 l2 --> IS_SUBLIST l1 l2"
-  by (import rich_list IS_SUFFIX_IS_SUBLIST)
-
-lemma IS_PREFIX_REVERSE: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2"
-  by (import rich_list IS_PREFIX_REVERSE)
-
-lemma IS_SUFFIX_REVERSE: "ALL (l2::'a::type list) l1::'a::type list.
-   IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2"
-  by (import rich_list IS_SUFFIX_REVERSE)
-
-lemma IS_SUBLIST_REVERSE: "ALL (l1::'a::type list) l2::'a::type list.
-   IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2"
-  by (import rich_list IS_SUBLIST_REVERSE)
-
-lemma PREFIX_FOLDR: "ALL (P::'a::type => bool) x::'a::type list.
-   PREFIX P x =
-   foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else []) x []"
-  by (import rich_list PREFIX_FOLDR)
-
-lemma PREFIX: "(ALL x::'a::type => bool. PREFIX x [] = []) &
-(ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type list.
+  sorry
+
+lemma APPEND_LENGTH_EQ: "[| length l1 = length l1'; length l2 = length l2' |]
+==> (l1 @ l2 = l1' @ l2') = (l1 = l1' & l2 = l2')"
+  sorry
+
+lemma FLAT_SNOC: "concat (SNOC x l) = concat l @ x"
+  sorry
+
+lemma FLAT_FOLDR: "concat l = foldr op @ l []"
+  sorry
+
+lemma LENGTH_FLAT: "length (concat l) = HOL4Compat.sum (map length l)"
+  sorry
+
+lemma REVERSE_FOLDR: "rev l = foldr SNOC l []"
+  sorry
+
+lemma ALL_EL_SNOC: "list_all P (SNOC x l) = (list_all P l & P x)"
+  sorry
+
+lemma ALL_EL_MAP: "list_all (P::'b => bool) (map (f::'a => 'b) (l::'a list)) =
+list_all (P o f) l"
+  sorry
+
+lemma SOME_EL_SNOC: "list_ex P (SNOC x l) = (P x | list_ex P l)"
+  sorry
+
+lemma IS_EL_SNOC: "List.member (SNOC x l) y = (y = x | List.member l y)"
+  sorry
+
+lemma SUM_SNOC: "HOL4Compat.sum (SNOC x l) = HOL4Compat.sum l + x"
+  sorry
+
+lemma SUM_FOLDL: "HOL4Compat.sum l = foldl op + 0 l"
+  sorry
+
+lemma IS_PREFIX_APPEND: "IS_PREFIX l1 l2 = (EX l. l1 = l2 @ l)"
+  sorry
+
+lemma IS_SUFFIX_APPEND: "IS_SUFFIX l1 l2 = (EX l. l1 = l @ l2)"
+  sorry
+
+lemma IS_SUBLIST_APPEND: "IS_SUBLIST l1 l2 = (EX l l'. l1 = l @ l2 @ l')"
+  sorry
+
+lemma IS_PREFIX_IS_SUBLIST: "IS_PREFIX l1 l2 ==> IS_SUBLIST l1 l2"
+  sorry
+
+lemma IS_SUFFIX_IS_SUBLIST: "IS_SUFFIX l1 l2 ==> IS_SUBLIST l1 l2"
+  sorry
+
+lemma IS_PREFIX_REVERSE: "IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2"
+  sorry
+
+lemma IS_SUFFIX_REVERSE: "IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2"
+  sorry
+
+lemma IS_SUBLIST_REVERSE: "IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2"
+  sorry
+
+lemma PREFIX_FOLDR: "PREFIX P x = foldr (%x l'. if P x then x # l' else []) x []"
+  sorry
+
+lemma PREFIX: "(ALL x::'a => bool. PREFIX x [] = []) &
+(ALL (x::'a => bool) (xa::'a) xb::'a list.
     PREFIX x (xa # xb) = (if x xa then xa # PREFIX x xb else []))"
-  by (import rich_list PREFIX)
-
-lemma IS_PREFIX_PREFIX: "ALL (P::'a::type => bool) l::'a::type list. IS_PREFIX l (PREFIX P l)"
-  by (import rich_list IS_PREFIX_PREFIX)
-
-lemma LENGTH_SCANL: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) l::'a::type list.
-   length (SCANL f e l) = Suc (length l)"
-  by (import rich_list LENGTH_SCANL)
-
-lemma LENGTH_SCANR: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) l::'a::type list.
-   length (SCANR f e l) = Suc (length l)"
-  by (import rich_list LENGTH_SCANR)
-
-lemma COMM_MONOID_FOLDL: "ALL x::'a::type => 'a::type => 'a::type.
-   COMM x -->
-   (ALL xa::'a::type.
-       MONOID x xa -->
-       (ALL (e::'a::type) l::'a::type list.
-           foldl x e l = x e (foldl x xa l)))"
-  by (import rich_list COMM_MONOID_FOLDL)
-
-lemma COMM_MONOID_FOLDR: "ALL x::'a::type => 'a::type => 'a::type.
-   COMM x -->
-   (ALL xa::'a::type.
-       MONOID x xa -->
-       (ALL (e::'a::type) l::'a::type list.
-           foldr x l e = x e (foldr x l xa)))"
-  by (import rich_list COMM_MONOID_FOLDR)
-
-lemma FCOMM_FOLDR_APPEND: "ALL (x::'a::type => 'a::type => 'a::type)
-   xa::'b::type => 'a::type => 'a::type.
-   FCOMM x xa -->
-   (ALL xb::'a::type.
-       LEFT_ID x xb -->
-       (ALL (l1::'b::type list) l2::'b::type list.
-           foldr xa (l1 @ l2) xb = x (foldr xa l1 xb) (foldr xa l2 xb)))"
-  by (import rich_list FCOMM_FOLDR_APPEND)
-
-lemma FCOMM_FOLDL_APPEND: "ALL (x::'a::type => 'b::type => 'a::type)
-   xa::'a::type => 'a::type => 'a::type.
-   FCOMM x xa -->
-   (ALL xb::'a::type.
-       RIGHT_ID xa xb -->
-       (ALL (l1::'b::type list) l2::'b::type list.
-           foldl x xb (l1 @ l2) = xa (foldl x xb l1) (foldl x xb l2)))"
-  by (import rich_list FCOMM_FOLDL_APPEND)
-
-lemma FOLDL_SINGLE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type.
-   foldl x xa [xb] = x xa xb"
-  by (import rich_list FOLDL_SINGLE)
-
-lemma FOLDR_SINGLE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type.
-   foldr x [xb] xa = x xb xa"
-  by (import rich_list FOLDR_SINGLE)
-
-lemma FOLDR_CONS_NIL: "ALL l::'a::type list. foldr op # l [] = l"
-  by (import rich_list FOLDR_CONS_NIL)
-
-lemma FOLDL_SNOC_NIL: "ALL l::'a::type list.
-   foldl (%(xs::'a::type list) x::'a::type. SNOC x xs) [] l = l"
-  by (import rich_list FOLDL_SNOC_NIL)
-
-lemma FOLDR_REVERSE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type list.
-   foldr x (rev xb) xa = foldl (%(xa::'b::type) y::'a::type. x y xa) xa xb"
-  by (import rich_list FOLDR_REVERSE)
-
-lemma FOLDL_REVERSE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type list.
-   foldl x xa (rev xb) = foldr (%(xa::'b::type) y::'a::type. x y xa) xb xa"
-  by (import rich_list FOLDL_REVERSE)
-
-lemma FOLDR_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
-   (g::'b::type => 'a::type) l::'b::type list.
-   foldr f (map g l) e = foldr (%x::'b::type. f (g x)) l e"
-  by (import rich_list FOLDR_MAP)
-
-lemma FOLDL_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
-   (g::'b::type => 'a::type) l::'b::type list.
-   foldl f e (map g l) = foldl (%(x::'a::type) y::'b::type. f x (g y)) e l"
-  by (import rich_list FOLDL_MAP)
-
-lemma ALL_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
-   list_all P l = foldr (%x::'a::type. op & (P x)) l True"
-  by (import rich_list ALL_EL_FOLDR)
-
-lemma ALL_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
-   list_all P l = foldl (%(l'::bool) x::'a::type. l' & P x) True l"
-  by (import rich_list ALL_EL_FOLDL)
-
-lemma SOME_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
-   list_ex P l = foldr (%x::'a::type. op | (P x)) l False"
-  by (import rich_list SOME_EL_FOLDR)
-
-lemma SOME_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
-   list_ex P l = foldl (%(l'::bool) x::'a::type. l' | P x) False l"
-  by (import rich_list SOME_EL_FOLDL)
-
-lemma ALL_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
-   list_all x xa = foldr op & (map x xa) True"
-  by (import rich_list ALL_EL_FOLDR_MAP)
-
-lemma ALL_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
-   list_all x xa = foldl op & True (map x xa)"
-  by (import rich_list ALL_EL_FOLDL_MAP)
-
-lemma SOME_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
-   list_ex x xa = foldr op | (map x xa) False"
-  by (import rich_list SOME_EL_FOLDR_MAP)
-
-lemma SOME_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
-   list_ex x xa = foldl op | False (map x xa)"
-  by (import rich_list SOME_EL_FOLDL_MAP)
-
-lemma FOLDR_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
-   (P::'a::type => bool) l::'a::type list.
-   foldr f (filter P l) e =
-   foldr (%(x::'a::type) y::'a::type. if P x then f x y else y) l e"
-  by (import rich_list FOLDR_FILTER)
-
-lemma FOLDL_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
-   (P::'a::type => bool) l::'a::type list.
-   foldl f e (filter P l) =
-   foldl (%(x::'a::type) y::'a::type. if P y then f x y else x) e l"
-  by (import rich_list FOLDL_FILTER)
-
-lemma ASSOC_FOLDR_FLAT: "ALL f::'a::type => 'a::type => 'a::type.
-   ASSOC f -->
-   (ALL e::'a::type.
-       LEFT_ID f e -->
-       (ALL l::'a::type list list.
-           foldr f (concat l) e = foldr f (map (FOLDR f e) l) e))"
-  by (import rich_list ASSOC_FOLDR_FLAT)
-
-lemma ASSOC_FOLDL_FLAT: "ALL f::'a::type => 'a::type => 'a::type.
-   ASSOC f -->
-   (ALL e::'a::type.
-       RIGHT_ID f e -->
-       (ALL l::'a::type list list.
-           foldl f e (concat l) = foldl f e (map (foldl f e) l)))"
-  by (import rich_list ASSOC_FOLDL_FLAT)
-
-lemma SOME_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list.
-   list_ex P (map f l) = list_ex (P o f) l"
-  by (import rich_list SOME_EL_MAP)
-
-lemma SOME_EL_DISJ: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list.
-   list_ex (%x::'a::type. P x | Q x) l =
-   (list_ex P l | list_ex Q l)"
-  by (import rich_list SOME_EL_DISJ)
-
-lemma IS_EL_FOLDR: "ALL (x::'a::type) xa::'a::type list.
-   x mem xa = foldr (%xa::'a::type. op | (x = xa)) xa False"
-  by (import rich_list IS_EL_FOLDR)
-
-lemma IS_EL_FOLDL: "ALL (x::'a::type) xa::'a::type list.
-   x mem xa = foldl (%(l'::bool) xa::'a::type. l' | x = xa) False xa"
-  by (import rich_list IS_EL_FOLDL)
-
-lemma NULL_FOLDR: "ALL l::'a::type list. null l = foldr (%(x::'a::type) l'::bool. False) l True"
-  by (import rich_list NULL_FOLDR)
-
-lemma NULL_FOLDL: "ALL l::'a::type list. null l = foldl (%(x::bool) l'::'a::type. False) True l"
-  by (import rich_list NULL_FOLDL)
-
-lemma SEG_LENGTH_ID: "ALL l::'a::type list. SEG (length l) 0 l = l"
-  by (import rich_list SEG_LENGTH_ID)
-
-lemma SEG_SUC_CONS: "ALL (m::nat) (n::nat) (l::'a::type list) x::'a::type.
-   SEG m (Suc n) (x # l) = SEG m n l"
-  by (import rich_list SEG_SUC_CONS)
-
-lemma SEG_0_SNOC: "ALL (m::nat) (l::'a::type list) x::'a::type.
-   m <= length l --> SEG m 0 (SNOC x l) = SEG m 0 l"
-  by (import rich_list SEG_0_SNOC)
-
-lemma BUTLASTN_SEG: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTLASTN n l = SEG (length l - n) 0 l"
-  by (import rich_list BUTLASTN_SEG)
-
-lemma LASTN_CONS: "ALL (n::nat) l::'a::type list.
-   n <= length l --> (ALL x::'a::type. LASTN n (x # l) = LASTN n l)"
-  by (import rich_list LASTN_CONS)
-
-lemma LENGTH_LASTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (LASTN n l) = n"
-  by (import rich_list LENGTH_LASTN)
-
-lemma LASTN_LENGTH_ID: "ALL l::'a::type list. LASTN (length l) l = l"
-  by (import rich_list LASTN_LENGTH_ID)
-
-lemma LASTN_LASTN: "ALL (l::'a::type list) (n::nat) m::nat.
-   m <= length l --> n <= m --> LASTN n (LASTN m l) = LASTN n l"
-  by (import rich_list LASTN_LASTN)
-
-lemma FIRSTN_LENGTH_ID: "ALL l::'a::type list. FIRSTN (length l) l = l"
-  by (import rich_list FIRSTN_LENGTH_ID)
-
-lemma FIRSTN_SNOC: "ALL (n::nat) l::'a::type list.
-   n <= length l --> (ALL x::'a::type. FIRSTN n (SNOC x l) = FIRSTN n l)"
-  by (import rich_list FIRSTN_SNOC)
-
-lemma BUTLASTN_LENGTH_NIL: "ALL l::'a::type list. BUTLASTN (length l) l = []"
-  by (import rich_list BUTLASTN_LENGTH_NIL)
-
-lemma BUTLASTN_SUC_BUTLAST: "ALL (n::nat) l::'a::type list.
-   n < length l --> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)"
-  by (import rich_list BUTLASTN_SUC_BUTLAST)
-
-lemma BUTLASTN_BUTLAST: "ALL (n::nat) l::'a::type list.
-   n < length l --> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)"
-  by (import rich_list BUTLASTN_BUTLAST)
-
-lemma LENGTH_BUTLASTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> length (BUTLASTN n l) = length l - n"
-  by (import rich_list LENGTH_BUTLASTN)
-
-lemma BUTLASTN_BUTLASTN: "ALL (m::nat) (n::nat) l::'a::type list.
-   n + m <= length l --> BUTLASTN n (BUTLASTN m l) = BUTLASTN (n + m) l"
-  by (import rich_list BUTLASTN_BUTLASTN)
-
-lemma APPEND_BUTLASTN_LASTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTLASTN n l @ LASTN n l = l"
-  by (import rich_list APPEND_BUTLASTN_LASTN)
-
-lemma APPEND_FIRSTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list.
-   m + n = length l --> FIRSTN n l @ LASTN m l = l"
-  by (import rich_list APPEND_FIRSTN_LASTN)
-
-lemma BUTLASTN_APPEND2: "ALL (n::nat) (l1::'a::type list) l2::'a::type list.
-   n <= length l2 --> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2"
-  by (import rich_list BUTLASTN_APPEND2)
-
-lemma BUTLASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list.
-   BUTLASTN (length l2) (l1 @ l2) = l1"
-  by (import rich_list BUTLASTN_LENGTH_APPEND)
-
-lemma LASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list. LASTN (length l2) (l1 @ l2) = l2"
-  by (import rich_list LASTN_LENGTH_APPEND)
-
-lemma BUTLASTN_CONS: "ALL (n::nat) l::'a::type list.
-   n <= length l -->
-   (ALL x::'a::type. BUTLASTN n (x # l) = x # BUTLASTN n l)"
-  by (import rich_list BUTLASTN_CONS)
-
-lemma BUTLASTN_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. BUTLASTN (length l) (x # l) = [x]"
-  by (import rich_list BUTLASTN_LENGTH_CONS)
-
-lemma LAST_LASTN_LAST: "ALL (n::nat) l::'a::type list.
-   n <= length l --> 0 < n --> last (LASTN n l) = last l"
-  by (import rich_list LAST_LASTN_LAST)
-
-lemma BUTLASTN_LASTN_NIL: "ALL (n::nat) l::'a::type list. n <= length l --> BUTLASTN n (LASTN n l) = []"
-  by (import rich_list BUTLASTN_LASTN_NIL)
-
-lemma LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l -->
-   LASTN n (BUTLASTN m l) = BUTLASTN m (LASTN (n + m) l)"
-  by (import rich_list LASTN_BUTLASTN)
-
-lemma BUTLASTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list.
-   m <= n & n <= length l -->
-   BUTLASTN m (LASTN n l) = LASTN (n - m) (BUTLASTN m l)"
-  by (import rich_list BUTLASTN_LASTN)
-
-lemma LASTN_1: "ALL l::'a::type list. l ~= [] --> LASTN 1 l = [last l]"
-  by (import rich_list LASTN_1)
-
-lemma BUTLASTN_1: "ALL l::'a::type list. l ~= [] --> BUTLASTN 1 l = butlast l"
-  by (import rich_list BUTLASTN_1)
-
-lemma BUTLASTN_APPEND1: "ALL (l2::'a::type list) n::nat.
-   length l2 <= n -->
-   (ALL l1::'a::type list.
-       BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1)"
-  by (import rich_list BUTLASTN_APPEND1)
-
-lemma LASTN_APPEND2: "ALL (n::nat) l2::'a::type list.
-   n <= length l2 -->
-   (ALL l1::'a::type list. LASTN n (l1 @ l2) = LASTN n l2)"
-  by (import rich_list LASTN_APPEND2)
-
-lemma LASTN_APPEND1: "ALL (l2::'a::type list) n::nat.
-   length l2 <= n -->
-   (ALL l1::'a::type list.
-       LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2)"
-  by (import rich_list LASTN_APPEND1)
-
-lemma LASTN_MAP: "ALL (n::nat) l::'a::type list.
-   n <= length l -->
-   (ALL f::'a::type => 'b::type. LASTN n (map f l) = map f (LASTN n l))"
-  by (import rich_list LASTN_MAP)
-
-lemma BUTLASTN_MAP: "ALL (n::nat) l::'a::type list.
-   n <= length l -->
-   (ALL f::'a::type => 'b::type.
-       BUTLASTN n (map f l) = map f (BUTLASTN n l))"
-  by (import rich_list BUTLASTN_MAP)
-
-lemma ALL_EL_LASTN: "(All::(('a::type => bool) => bool) => bool)
- (%P::'a::type => bool.
-     (All::('a::type list => bool) => bool)
-      (%l::'a::type list.
-          (op -->::bool => bool => bool)
-           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
-           ((All::(nat => bool) => bool)
-             (%m::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) m
-                    ((size::'a::type list => nat) l))
-                  ((list_all::('a::type => bool) => 'a::type list => bool) P
-                    ((LASTN::nat => 'a::type list => 'a::type list) m
-                      l))))))"
-  by (import rich_list ALL_EL_LASTN)
-
-lemma ALL_EL_BUTLASTN: "(All::(('a::type => bool) => bool) => bool)
- (%P::'a::type => bool.
-     (All::('a::type list => bool) => bool)
-      (%l::'a::type list.
-          (op -->::bool => bool => bool)
-           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
-           ((All::(nat => bool) => bool)
-             (%m::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) m
-                    ((size::'a::type list => nat) l))
-                  ((list_all::('a::type => bool) => 'a::type list => bool) P
-                    ((BUTLASTN::nat => 'a::type list => 'a::type list) m
-                      l))))))"
-  by (import rich_list ALL_EL_BUTLASTN)
-
-lemma LENGTH_FIRSTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (FIRSTN n l) = n"
-  by (import rich_list LENGTH_FIRSTN)
-
-lemma FIRSTN_FIRSTN: "(All::(nat => bool) => bool)
- (%m::nat.
-     (All::('a::type list => bool) => bool)
-      (%l::'a::type list.
-          (op -->::bool => bool => bool)
-           ((op <=::nat => nat => bool) m ((size::'a::type list => nat) l))
-           ((All::(nat => bool) => bool)
-             (%n::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) n m)
-                  ((op =::'a::type list => 'a::type list => bool)
-                    ((FIRSTN::nat => 'a::type list => 'a::type list) n
-                      ((FIRSTN::nat => 'a::type list => 'a::type list) m l))
-                    ((FIRSTN::nat => 'a::type list => 'a::type list) n
-                      l))))))"
-  by (import rich_list FIRSTN_FIRSTN)
-
-lemma LENGTH_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> length (BUTFIRSTN n l) = length l - n"
-  by (import rich_list LENGTH_BUTFIRSTN)
-
-lemma BUTFIRSTN_LENGTH_NIL: "ALL l::'a::type list. BUTFIRSTN (length l) l = []"
-  by (import rich_list BUTFIRSTN_LENGTH_NIL)
-
-lemma BUTFIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list.
-   n <= length l1 -->
-   (ALL l2::'a::type list. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2)"
-  by (import rich_list BUTFIRSTN_APPEND1)
-
-lemma BUTFIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat.
-   length l1 <= n -->
-   (ALL l2::'a::type list.
-       BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2)"
-  by (import rich_list BUTFIRSTN_APPEND2)
-
-lemma BUTFIRSTN_BUTFIRSTN: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l --> BUTFIRSTN n (BUTFIRSTN m l) = BUTFIRSTN (n + m) l"
-  by (import rich_list BUTFIRSTN_BUTFIRSTN)
-
-lemma APPEND_FIRSTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> FIRSTN n l @ BUTFIRSTN n l = l"
-  by (import rich_list APPEND_FIRSTN_BUTFIRSTN)
-
-lemma LASTN_SEG: "ALL (n::nat) l::'a::type list.
-   n <= length l --> LASTN n l = SEG n (length l - n) l"
-  by (import rich_list LASTN_SEG)
-
-lemma FIRSTN_SEG: "ALL (n::nat) l::'a::type list. n <= length l --> FIRSTN n l = SEG n 0 l"
-  by (import rich_list FIRSTN_SEG)
-
-lemma BUTFIRSTN_SEG: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTFIRSTN n l = SEG (length l - n) n l"
-  by (import rich_list BUTFIRSTN_SEG)
-
-lemma BUTFIRSTN_SNOC: "ALL (n::nat) l::'a::type list.
-   n <= length l -->
-   (ALL x::'a::type. BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l))"
-  by (import rich_list BUTFIRSTN_SNOC)
-
-lemma APPEND_BUTLASTN_BUTFIRSTN: "ALL (m::nat) (n::nat) l::'a::type list.
-   m + n = length l --> BUTLASTN m l @ BUTFIRSTN n l = l"
-  by (import rich_list APPEND_BUTLASTN_BUTFIRSTN)
-
-lemma SEG_SEG: "ALL (n1::nat) (m1::nat) (n2::nat) (m2::nat) l::'a::type list.
-   n1 + m1 <= length l & n2 + m2 <= n1 -->
-   SEG n2 m2 (SEG n1 m1 l) = SEG n2 (m1 + m2) l"
-  by (import rich_list SEG_SEG)
-
-lemma SEG_APPEND1: "ALL (n::nat) (m::nat) l1::'a::type list.
-   n + m <= length l1 -->
-   (ALL l2::'a::type list. SEG n m (l1 @ l2) = SEG n m l1)"
-  by (import rich_list SEG_APPEND1)
-
-lemma SEG_APPEND2: "ALL (l1::'a::type list) (m::nat) (n::nat) l2::'a::type list.
-   length l1 <= m & n <= length l2 -->
-   SEG n m (l1 @ l2) = SEG n (m - length l1) l2"
-  by (import rich_list SEG_APPEND2)
-
-lemma SEG_FIRSTN_BUTFISTN: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l --> SEG n m l = FIRSTN n (BUTFIRSTN m l)"
-  by (import rich_list SEG_FIRSTN_BUTFISTN)
-
-lemma SEG_APPEND: "ALL (m::nat) (l1::'a::type list) (n::nat) l2::'a::type list.
-   m < length l1 & length l1 <= n + m & n + m <= length l1 + length l2 -->
-   SEG n m (l1 @ l2) =
-   SEG (length l1 - m) m l1 @ SEG (n + m - length l1) 0 l2"
-  by (import rich_list SEG_APPEND)
-
-lemma SEG_LENGTH_SNOC: "ALL (x::'a::type list) xa::'a::type. SEG 1 (length x) (SNOC xa x) = [xa]"
-  by (import rich_list SEG_LENGTH_SNOC)
-
-lemma SEG_SNOC: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l --> (ALL x::'a::type. SEG n m (SNOC x l) = SEG n m l)"
-  by (import rich_list SEG_SNOC)
-
-lemma ELL_SEG: "ALL (n::nat) l::'a::type list.
-   n < length l --> ELL n l = hd (SEG 1 (PRE (length l - n)) l)"
-  by (import rich_list ELL_SEG)
-
-lemma SNOC_FOLDR: "ALL (x::'a::type) l::'a::type list. SNOC x l = foldr op # l [x]"
-  by (import rich_list SNOC_FOLDR)
-
-lemma IS_EL_FOLDR_MAP: "ALL (x::'a::type) xa::'a::type list.
-   x mem xa = foldr op | (map (op = x) xa) False"
-  by (import rich_list IS_EL_FOLDR_MAP)
-
-lemma IS_EL_FOLDL_MAP: "ALL (x::'a::type) xa::'a::type list.
-   x mem xa = foldl op | False (map (op = x) xa)"
-  by (import rich_list IS_EL_FOLDL_MAP)
-
-lemma FILTER_FILTER: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list.
-   filter P (filter Q l) = [x::'a::type<-l. P x & Q x]"
-  by (import rich_list FILTER_FILTER)
-
-lemma FCOMM_FOLDR_FLAT: "ALL (g::'a::type => 'a::type => 'a::type)
-   f::'b::type => 'a::type => 'a::type.
-   FCOMM g f -->
-   (ALL e::'a::type.
-       LEFT_ID g e -->
-       (ALL l::'b::type list list.
-           foldr f (concat l) e = foldr g (map (FOLDR f e) l) e))"
-  by (import rich_list FCOMM_FOLDR_FLAT)
-
-lemma FCOMM_FOLDL_FLAT: "ALL (f::'a::type => 'b::type => 'a::type)
-   g::'a::type => 'a::type => 'a::type.
-   FCOMM f g -->
-   (ALL e::'a::type.
-       RIGHT_ID g e -->
-       (ALL l::'b::type list list.
-           foldl f e (concat l) = foldl g e (map (foldl f e) l)))"
-  by (import rich_list FCOMM_FOLDL_FLAT)
-
-lemma FOLDR_MAP_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
-   (ALL (a::'a::type) (b::'a::type) c::'a::type.
-       f a (f b c) = f b (f a c)) -->
-   (ALL (e::'a::type) (g::'b::type => 'a::type) l::'b::type list.
-       foldr f (map g (rev l)) e = foldr f (map g l) e)"
-  by (import rich_list FOLDR_MAP_REVERSE)
-
-lemma FOLDR_FILTER_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
-   (ALL (a::'a::type) (b::'a::type) c::'a::type.
-       f a (f b c) = f b (f a c)) -->
-   (ALL (e::'a::type) (P::'a::type => bool) l::'a::type list.
-       foldr f (filter P (rev l)) e = foldr f (filter P l) e)"
-  by (import rich_list FOLDR_FILTER_REVERSE)
-
-lemma COMM_ASSOC_FOLDR_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
-   COMM f -->
-   ASSOC f -->
-   (ALL (e::'a::type) l::'a::type list. foldr f (rev l) e = foldr f l e)"
-  by (import rich_list COMM_ASSOC_FOLDR_REVERSE)
-
-lemma COMM_ASSOC_FOLDL_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
-   COMM f -->
-   ASSOC f -->
-   (ALL (e::'a::type) l::'a::type list. foldl f e (rev l) = foldl f e l)"
-  by (import rich_list COMM_ASSOC_FOLDL_REVERSE)
-
-lemma ELL_LAST: "ALL l::'a::type list. ~ null l --> ELL 0 l = last l"
-  by (import rich_list ELL_LAST)
-
-lemma ELL_0_SNOC: "ALL (l::'a::type list) x::'a::type. ELL 0 (SNOC x l) = x"
-  by (import rich_list ELL_0_SNOC)
-
-lemma ELL_SNOC: "ALL n>0.
-   ALL (x::'a::type) l::'a::type list. ELL n (SNOC x l) = ELL (PRE n) l"
-  by (import rich_list ELL_SNOC)
-
-lemma ELL_SUC_SNOC: "ALL (n::nat) (x::'a::type) xa::'a::type list.
-   ELL (Suc n) (SNOC x xa) = ELL n xa"
-  by (import rich_list ELL_SUC_SNOC)
-
-lemma ELL_CONS: "ALL (n::nat) l::'a::type list.
-   n < length l --> (ALL x::'a::type. ELL n (x # l) = ELL n l)"
-  by (import rich_list ELL_CONS)
-
-lemma ELL_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. ELL (length l) (x # l) = x"
-  by (import rich_list ELL_LENGTH_CONS)
-
-lemma ELL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type.
-   ELL (length l) (SNOC x l) = (if null l then x else hd l)"
-  by (import rich_list ELL_LENGTH_SNOC)
-
-lemma ELL_APPEND2: "ALL (n::nat) l2::'a::type list.
-   n < length l2 --> (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL n l2)"
-  by (import rich_list ELL_APPEND2)
-
-lemma ELL_APPEND1: "ALL (l2::'a::type list) n::nat.
-   length l2 <= n -->
-   (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL (n - length l2) l1)"
-  by (import rich_list ELL_APPEND1)
-
-lemma ELL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> ELL (PRE (length l)) l = hd l"
-  by (import rich_list ELL_PRE_LENGTH)
-
-lemma EL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type. EL (length l) (SNOC x l) = x"
-  by (import rich_list EL_LENGTH_SNOC)
-
-lemma EL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> EL (PRE (length l)) l = last l"
-  by (import rich_list EL_PRE_LENGTH)
-
-lemma EL_SNOC: "ALL (n::nat) l::'a::type list.
-   n < length l --> (ALL x::'a::type. EL n (SNOC x l) = EL n l)"
-  by (import rich_list EL_SNOC)
-
-lemma EL_ELL: "ALL (n::nat) l::'a::type list.
-   n < length l --> EL n l = ELL (PRE (length l - n)) l"
-  by (import rich_list EL_ELL)
-
-lemma EL_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list.
-   ~ null l2 --> EL (length l1) (l1 @ l2) = hd l2"
-  by (import rich_list EL_LENGTH_APPEND)
-
-lemma ELL_EL: "ALL (n::nat) l::'a::type list.
-   n < length l --> ELL n l = EL (PRE (length l - n)) l"
-  by (import rich_list ELL_EL)
-
-lemma ELL_MAP: "ALL (n::nat) (l::'a::type list) f::'a::type => 'b::type.
-   n < length l --> ELL n (map f l) = f (ELL n l)"
-  by (import rich_list ELL_MAP)
-
-lemma LENGTH_BUTLAST: "ALL l::'a::type list. l ~= [] --> length (butlast l) = PRE (length l)"
-  by (import rich_list LENGTH_BUTLAST)
-
-lemma BUTFIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
-   BUTFIRSTN (length l1) (l1 @ l2) = l2"
-  by (import rich_list BUTFIRSTN_LENGTH_APPEND)
-
-lemma FIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list.
-   n <= length l1 -->
-   (ALL l2::'a::type list. FIRSTN n (l1 @ l2) = FIRSTN n l1)"
-  by (import rich_list FIRSTN_APPEND1)
-
-lemma FIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat.
-   length l1 <= n -->
-   (ALL l2::'a::type list.
-       FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2)"
-  by (import rich_list FIRSTN_APPEND2)
-
-lemma FIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list. FIRSTN (length l1) (l1 @ l2) = l1"
-  by (import rich_list FIRSTN_LENGTH_APPEND)
-
-lemma REVERSE_FLAT: "ALL l::'a::type list list. rev (concat l) = concat (rev (map rev l))"
-  by (import rich_list REVERSE_FLAT)
-
-lemma MAP_FILTER: "ALL (f::'a::type => 'a::type) (P::'a::type => bool) l::'a::type list.
-   (ALL x::'a::type. P (f x) = P x) -->
-   map f (filter P l) = filter P (map f l)"
-  by (import rich_list MAP_FILTER)
-
-lemma FLAT_REVERSE: "ALL l::'a::type list list. concat (rev l) = rev (concat (map rev l))"
-  by (import rich_list FLAT_REVERSE)
-
-lemma FLAT_FLAT: "ALL l::'a::type list list list. concat (concat l) = concat (map concat l)"
-  by (import rich_list FLAT_FLAT)
-
-lemma SOME_EL_REVERSE: "ALL (P::'a::type => bool) l::'a::type list.
-   list_ex P (rev l) = list_ex P l"
-  by (import rich_list SOME_EL_REVERSE)
-
-lemma ALL_EL_SEG: "ALL (P::'a::type => bool) l::'a::type list.
-   list_all P l -->
-   (ALL (m::nat) k::nat. m + k <= length l --> list_all P (SEG m k l))"
-  by (import rich_list ALL_EL_SEG)
-
-lemma ALL_EL_FIRSTN: "(All::(('a::type => bool) => bool) => bool)
- (%P::'a::type => bool.
-     (All::('a::type list => bool) => bool)
-      (%l::'a::type list.
-          (op -->::bool => bool => bool)
-           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
-           ((All::(nat => bool) => bool)
-             (%m::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) m
-                    ((size::'a::type list => nat) l))
-                  ((list_all::('a::type => bool) => 'a::type list => bool) P
-                    ((FIRSTN::nat => 'a::type list => 'a::type list) m
-                      l))))))"
-  by (import rich_list ALL_EL_FIRSTN)
-
-lemma ALL_EL_BUTFIRSTN: "(All::(('a::type => bool) => bool) => bool)
- (%P::'a::type => bool.
-     (All::('a::type list => bool) => bool)
-      (%l::'a::type list.
-          (op -->::bool => bool => bool)
-           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
-           ((All::(nat => bool) => bool)
-             (%m::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) m
-                    ((size::'a::type list => nat) l))
-                  ((list_all::('a::type => bool) => 'a::type list => bool) P
-                    ((BUTFIRSTN::nat => 'a::type list => 'a::type list) m
-                      l))))))"
-  by (import rich_list ALL_EL_BUTFIRSTN)
-
-lemma SOME_EL_SEG: "ALL (m::nat) (k::nat) l::'a::type list.
-   m + k <= length l -->
-   (ALL P::'a::type => bool. list_ex P (SEG m k l) --> list_ex P l)"
-  by (import rich_list SOME_EL_SEG)
-
-lemma SOME_EL_FIRSTN: "ALL (m::nat) l::'a::type list.
-   m <= length l -->
-   (ALL P::'a::type => bool. list_ex P (FIRSTN m l) --> list_ex P l)"
-  by (import rich_list SOME_EL_FIRSTN)
-
-lemma SOME_EL_BUTFIRSTN: "ALL (m::nat) l::'a::type list.
-   m <= length l -->
-   (ALL P::'a::type => bool.
-       list_ex P (BUTFIRSTN m l) --> list_ex P l)"
-  by (import rich_list SOME_EL_BUTFIRSTN)
-
-lemma SOME_EL_LASTN: "ALL (m::nat) l::'a::type list.
-   m <= length l -->
-   (ALL P::'a::type => bool. list_ex P (LASTN m l) --> list_ex P l)"
-  by (import rich_list SOME_EL_LASTN)
-
-lemma SOME_EL_BUTLASTN: "ALL (m::nat) l::'a::type list.
-   m <= length l -->
-   (ALL P::'a::type => bool.
-       list_ex P (BUTLASTN m l) --> list_ex P l)"
-  by (import rich_list SOME_EL_BUTLASTN)
-
-lemma IS_EL_REVERSE: "ALL (x::'a::type) l::'a::type list. x mem rev l = x mem l"
-  by (import rich_list IS_EL_REVERSE)
-
-lemma IS_EL_FILTER: "ALL (P::'a::type => bool) x::'a::type.
-   P x --> (ALL l::'a::type list. x mem filter P l = x mem l)"
-  by (import rich_list IS_EL_FILTER)
-
-lemma IS_EL_SEG: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l --> (ALL x::'a::type. x mem SEG n m l --> x mem l)"
-  by (import rich_list IS_EL_SEG)
-
-lemma IS_EL_SOME_EL: "ALL (x::'a::type) l::'a::type list. x mem l = list_ex (op = x) l"
-  by (import rich_list IS_EL_SOME_EL)
-
-lemma IS_EL_FIRSTN: "ALL (x::nat) xa::'a::type list.
-   x <= length xa --> (ALL xb::'a::type. xb mem FIRSTN x xa --> xb mem xa)"
-  by (import rich_list IS_EL_FIRSTN)
-
-lemma IS_EL_BUTFIRSTN: "ALL (x::nat) xa::'a::type list.
-   x <= length xa -->
-   (ALL xb::'a::type. xb mem BUTFIRSTN x xa --> xb mem xa)"
-  by (import rich_list IS_EL_BUTFIRSTN)
-
-lemma IS_EL_BUTLASTN: "ALL (x::nat) xa::'a::type list.
-   x <= length xa --> (ALL xb::'a::type. xb mem BUTLASTN x xa --> xb mem xa)"
-  by (import rich_list IS_EL_BUTLASTN)
-
-lemma IS_EL_LASTN: "ALL (x::nat) xa::'a::type list.
-   x <= length xa --> (ALL xb::'a::type. xb mem LASTN x xa --> xb mem xa)"
-  by (import rich_list IS_EL_LASTN)
-
-lemma ZIP_SNOC: "ALL (l1::'a::type list) l2::'b::type list.
-   length l1 = length l2 -->
-   (ALL (x1::'a::type) x2::'b::type.
-       zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2))"
-  by (import rich_list ZIP_SNOC)
-
-lemma UNZIP_SNOC: "ALL (x::'a::type * 'b::type) l::('a::type * 'b::type) list.
-   unzip (SNOC x l) =
-   (SNOC (fst x) (fst (unzip l)), SNOC (snd x) (snd (unzip l)))"
-  by (import rich_list UNZIP_SNOC)
-
-lemma LENGTH_UNZIP_FST: "ALL x::('a::type * 'b::type) list. length (UNZIP_FST x) = length x"
-  by (import rich_list LENGTH_UNZIP_FST)
-
-lemma LENGTH_UNZIP_SND: "ALL x::('a::type * 'b::type) list. length (UNZIP_SND x) = length x"
-  by (import rich_list LENGTH_UNZIP_SND)
-
-lemma SUM_APPEND: "ALL (l1::nat list) l2::nat list. sum (l1 @ l2) = sum l1 + sum l2"
-  by (import rich_list SUM_APPEND)
-
-lemma SUM_REVERSE: "ALL l::nat list. sum (rev l) = sum l"
-  by (import rich_list SUM_REVERSE)
-
-lemma SUM_FLAT: "ALL l::nat list list. sum (concat l) = sum (map sum l)"
-  by (import rich_list SUM_FLAT)
-
-lemma EL_APPEND1: "ALL (n::nat) (l1::'a::type list) l2::'a::type list.
-   n < length l1 --> EL n (l1 @ l2) = EL n l1"
-  by (import rich_list EL_APPEND1)
-
-lemma EL_APPEND2: "ALL (l1::'a::type list) n::nat.
-   length l1 <= n -->
-   (ALL l2::'a::type list. EL n (l1 @ l2) = EL (n - length l1) l2)"
-  by (import rich_list EL_APPEND2)
-
-lemma EL_MAP: "ALL (n::nat) l::'a::type list.
-   n < length l -->
-   (ALL f::'a::type => 'b::type. EL n (map f l) = f (EL n l))"
-  by (import rich_list EL_MAP)
-
-lemma EL_CONS: "ALL n>0. ALL (x::'a::type) l::'a::type list. EL n (x # l) = EL (PRE n) l"
-  by (import rich_list EL_CONS)
-
-lemma EL_SEG: "ALL (n::nat) l::'a::type list. n < length l --> EL n l = hd (SEG 1 n l)"
-  by (import rich_list EL_SEG)
-
-lemma EL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> EL n l mem l"
-  by (import rich_list EL_IS_EL)
-
-lemma TL_SNOC: "ALL (x::'a::type) l::'a::type list.
-   tl (SNOC x l) = (if null l then [] else SNOC x (tl l))"
-  by (import rich_list TL_SNOC)
-
-lemma EL_REVERSE: "ALL (n::nat) l::'a::type list.
-   n < length l --> EL n (rev l) = EL (PRE (length l - n)) l"
-  by (import rich_list EL_REVERSE)
-
-lemma EL_REVERSE_ELL: "ALL (n::nat) l::'a::type list. n < length l --> EL n (rev l) = ELL n l"
-  by (import rich_list EL_REVERSE_ELL)
-
-lemma ELL_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
-   ~ null l1 --> ELL (length l2) (l1 @ l2) = last l1"
-  by (import rich_list ELL_LENGTH_APPEND)
-
-lemma ELL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n l mem l"
-  by (import rich_list ELL_IS_EL)
-
-lemma ELL_REVERSE: "ALL (n::nat) l::'a::type list.
-   n < length l --> ELL n (rev l) = ELL (PRE (length l - n)) l"
-  by (import rich_list ELL_REVERSE)
-
-lemma ELL_REVERSE_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n (rev l) = EL n l"
-  by (import rich_list ELL_REVERSE_EL)
-
-lemma FIRSTN_BUTLASTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> FIRSTN n l = BUTLASTN (length l - n) l"
-  by (import rich_list FIRSTN_BUTLASTN)
-
-lemma BUTLASTN_FIRSTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTLASTN n l = FIRSTN (length l - n) l"
-  by (import rich_list BUTLASTN_FIRSTN)
-
-lemma LASTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> LASTN n l = BUTFIRSTN (length l - n) l"
-  by (import rich_list LASTN_BUTFIRSTN)
-
-lemma BUTFIRSTN_LASTN: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTFIRSTN n l = LASTN (length l - n) l"
-  by (import rich_list BUTFIRSTN_LASTN)
-
-lemma SEG_LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l -->
-   SEG n m l = LASTN n (BUTLASTN (length l - (n + m)) l)"
-  by (import rich_list SEG_LASTN_BUTLASTN)
-
-lemma BUTFIRSTN_REVERSE: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)"
-  by (import rich_list BUTFIRSTN_REVERSE)
-
-lemma BUTLASTN_REVERSE: "ALL (n::nat) l::'a::type list.
-   n <= length l --> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)"
-  by (import rich_list BUTLASTN_REVERSE)
-
-lemma LASTN_REVERSE: "ALL (n::nat) l::'a::type list.
-   n <= length l --> LASTN n (rev l) = rev (FIRSTN n l)"
-  by (import rich_list LASTN_REVERSE)
-
-lemma FIRSTN_REVERSE: "ALL (n::nat) l::'a::type list.
-   n <= length l --> FIRSTN n (rev l) = rev (LASTN n l)"
-  by (import rich_list FIRSTN_REVERSE)
-
-lemma SEG_REVERSE: "ALL (n::nat) (m::nat) l::'a::type list.
-   n + m <= length l -->
-   SEG n m (rev l) = rev (SEG n (length l - (n + m)) l)"
-  by (import rich_list SEG_REVERSE)
-
-lemma LENGTH_GENLIST: "ALL (f::nat => 'a::type) n::nat. length (GENLIST f n) = n"
-  by (import rich_list LENGTH_GENLIST)
-
-lemma LENGTH_REPLICATE: "ALL (n::nat) x::'a::type. length (REPLICATE n x) = n"
-  by (import rich_list LENGTH_REPLICATE)
-
-lemma IS_EL_REPLICATE: "ALL n>0. ALL x::'a::type. x mem REPLICATE n x"
-  by (import rich_list IS_EL_REPLICATE)
-
-lemma ALL_EL_REPLICATE: "ALL (x::'a::type) n::nat. list_all (op = x) (REPLICATE n x)"
-  by (import rich_list ALL_EL_REPLICATE)
-
-lemma AND_EL_FOLDL: "ALL l::bool list. AND_EL l = foldl op & True l"
-  by (import rich_list AND_EL_FOLDL)
-
-lemma AND_EL_FOLDR: "ALL l::bool list. AND_EL l = foldr op & l True"
-  by (import rich_list AND_EL_FOLDR)
-
-lemma OR_EL_FOLDL: "ALL l::bool list. OR_EL l = foldl op | False l"
-  by (import rich_list OR_EL_FOLDL)
-
-lemma OR_EL_FOLDR: "ALL l::bool list. OR_EL l = foldr op | l False"
-  by (import rich_list OR_EL_FOLDR)
+  sorry
+
+lemma IS_PREFIX_PREFIX: "IS_PREFIX l (PREFIX P l)"
+  sorry
+
+lemma LENGTH_SCANL: "length (SCANL (f::'b => 'a => 'b) (e::'b) (l::'a list)) = Suc (length l)"
+  sorry
+
+lemma LENGTH_SCANR: "length (SCANR (f::'a => 'b => 'b) (e::'b) (l::'a list)) = Suc (length l)"
+  sorry
+
+lemma COMM_MONOID_FOLDL: "[| COMM x; MONOID x xa |] ==> foldl x e l = x e (foldl x xa l)"
+  sorry
+
+lemma COMM_MONOID_FOLDR: "[| COMM x; MONOID x xa |] ==> foldr x l e = x e (foldr x l xa)"
+  sorry
+
+lemma FCOMM_FOLDR_APPEND: "[| FCOMM x xa; LEFT_ID x xb |]
+==> foldr xa (l1 @ l2) xb = x (foldr xa l1 xb) (foldr xa l2 xb)"
+  sorry
+
+lemma FCOMM_FOLDL_APPEND: "[| FCOMM x xa; RIGHT_ID xa xb |]
+==> foldl x xb (l1 @ l2) = xa (foldl x xb l1) (foldl x xb l2)"
+  sorry
+
+lemma FOLDL_SINGLE: "foldl x xa [xb] = x xa xb"
+  sorry
+
+lemma FOLDR_SINGLE: "foldr (x::'a => 'b => 'b) [xb::'a] (xa::'b) = x xb xa"
+  sorry
+
+lemma FOLDR_CONS_NIL: "foldr op # l [] = l"
+  sorry
+
+lemma FOLDL_SNOC_NIL: "foldl (%xs x. SNOC x xs) [] l = l"
+  sorry
+
+lemma FOLDR_REVERSE: "foldr (x::'a => 'b => 'b) (rev (xb::'a list)) (xa::'b) =
+foldl (%(xa::'b) y::'a. x y xa) xa xb"
+  sorry
+
+lemma FOLDL_REVERSE: "foldl x xa (rev xb) = foldr (%xa y. x y xa) xb xa"
+  sorry
+
+lemma FOLDR_MAP: "foldr (f::'a => 'a => 'a) (map (g::'b => 'a) (l::'b list)) (e::'a) =
+foldr (%x::'b. f (g x)) l e"
+  sorry
+
+lemma ALL_EL_FOLDR: "list_all P l = foldr (%x. op & (P x)) l True"
+  sorry
+
+lemma ALL_EL_FOLDL: "list_all P l = foldl (%l' x. l' & P x) True l"
+  sorry
+
+lemma SOME_EL_FOLDR: "list_ex P l = foldr (%x. op | (P x)) l False"
+  sorry
+
+lemma SOME_EL_FOLDL: "list_ex P l = foldl (%l' x. l' | P x) False l"
+  sorry
+
+lemma ALL_EL_FOLDR_MAP: "list_all x xa = foldr op & (map x xa) True"
+  sorry
+
+lemma ALL_EL_FOLDL_MAP: "list_all x xa = foldl op & True (map x xa)"
+  sorry
+
+lemma SOME_EL_FOLDR_MAP: "list_ex x xa = foldr op | (map x xa) False"
+  sorry
+
+lemma SOME_EL_FOLDL_MAP: "list_ex x xa = foldl op | False (map x xa)"
+  sorry
+
+lemma FOLDR_FILTER: "foldr (f::'a => 'a => 'a) (filter (P::'a => bool) (l::'a list)) (e::'a) =
+foldr (%(x::'a) y::'a. if P x then f x y else y) l e"
+  sorry
+
+lemma FOLDL_FILTER: "foldl (f::'a => 'a => 'a) (e::'a) (filter (P::'a => bool) (l::'a list)) =
+foldl (%(x::'a) y::'a. if P y then f x y else x) e l"
+  sorry
+
+lemma ASSOC_FOLDR_FLAT: "[| ASSOC f; LEFT_ID f e |]
+==> foldr f (concat l) e = foldr f (map (FOLDR f e) l) e"
+  sorry
+
+lemma ASSOC_FOLDL_FLAT: "[| ASSOC f; RIGHT_ID f e |]
+==> foldl f e (concat l) = foldl f e (map (foldl f e) l)"
+  sorry
+
+lemma SOME_EL_MAP: "list_ex (P::'b => bool) (map (f::'a => 'b) (l::'a list)) = list_ex (P o f) l"
+  sorry
+
+lemma SOME_EL_DISJ: "list_ex (%x. P x | Q x) l = (list_ex P l | list_ex Q l)"
+  sorry
+
+lemma IS_EL_FOLDR: "List.member xa x = foldr (%xa. op | (x = xa)) xa False"
+  sorry
+
+lemma IS_EL_FOLDL: "List.member xa x = foldl (%l' xa. l' | x = xa) False xa"
+  sorry
+
+lemma NULL_FOLDR: "List.null l = foldr (%x l'. False) l True"
+  sorry
+
+lemma NULL_FOLDL: "List.null l = foldl (%x l'. False) True l"
+  sorry
+
+lemma SEG_LENGTH_ID: "SEG (length l) 0 l = l"
+  sorry
+
+lemma SEG_SUC_CONS: "SEG m (Suc n) (x # l) = SEG m n l"
+  sorry
+
+lemma SEG_0_SNOC: "m <= length l ==> SEG m 0 (SNOC x l) = SEG m 0 l"
+  sorry
+
+lemma BUTLASTN_SEG: "n <= length l ==> BUTLASTN n l = SEG (length l - n) 0 l"
+  sorry
+
+lemma LASTN_CONS: "n <= length l ==> LASTN n (x # l) = LASTN n l"
+  sorry
+
+lemma LENGTH_LASTN: "n <= length l ==> length (LASTN n l) = n"
+  sorry
+
+lemma LASTN_LENGTH_ID: "LASTN (length l) l = l"
+  sorry
+
+lemma LASTN_LASTN: "[| m <= length l; n <= m |] ==> LASTN n (LASTN m l) = LASTN n l"
+  sorry
+
+lemma FIRSTN_LENGTH_ID: "FIRSTN (length l) l = l"
+  sorry
+
+lemma FIRSTN_SNOC: "n <= length l ==> FIRSTN n (SNOC x l) = FIRSTN n l"
+  sorry
+
+lemma BUTLASTN_LENGTH_NIL: "BUTLASTN (length l) l = []"
+  sorry
+
+lemma BUTLASTN_SUC_BUTLAST: "n < length l ==> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)"
+  sorry
+
+lemma BUTLASTN_BUTLAST: "n < length l ==> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)"
+  sorry
+
+lemma LENGTH_BUTLASTN: "n <= length l ==> length (BUTLASTN n l) = length l - n"
+  sorry
+
+lemma BUTLASTN_BUTLASTN: "n + m <= length l ==> BUTLASTN n (BUTLASTN m l) = BUTLASTN (n + m) l"
+  sorry
+
+lemma APPEND_BUTLASTN_LASTN: "n <= length l ==> BUTLASTN n l @ LASTN n l = l"
+  sorry
+
+lemma APPEND_FIRSTN_LASTN: "m + n = length l ==> FIRSTN n l @ LASTN m l = l"
+  sorry
+
+lemma BUTLASTN_APPEND2: "n <= length l2 ==> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2"
+  sorry
+
+lemma BUTLASTN_LENGTH_APPEND: "BUTLASTN (length l2) (l1 @ l2) = l1"
+  sorry
+
+lemma LASTN_LENGTH_APPEND: "LASTN (length l2) (l1 @ l2) = l2"
+  sorry
+
+lemma BUTLASTN_CONS: "n <= length l ==> BUTLASTN n (x # l) = x # BUTLASTN n l"
+  sorry
+
+lemma BUTLASTN_LENGTH_CONS: "BUTLASTN (length l) (x # l) = [x]"
+  sorry
+
+lemma LAST_LASTN_LAST: "[| n <= length l; 0 < n |] ==> last (LASTN n l) = last l"
+  sorry
+
+lemma BUTLASTN_LASTN_NIL: "n <= length l ==> BUTLASTN n (LASTN n l) = []"
+  sorry
+
+lemma LASTN_BUTLASTN: "n + m <= length l ==> LASTN n (BUTLASTN m l) = BUTLASTN m (LASTN (n + m) l)"
+  sorry
+
+lemma BUTLASTN_LASTN: "m <= n & n <= length l
+==> BUTLASTN m (LASTN n l) = LASTN (n - m) (BUTLASTN m l)"
+  sorry
+
+lemma LASTN_1: "l ~= [] ==> LASTN 1 l = [last l]"
+  sorry
+
+lemma BUTLASTN_1: "l ~= [] ==> BUTLASTN 1 l = butlast l"
+  sorry
+
+lemma BUTLASTN_APPEND1: "length l2 <= n ==> BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1"
+  sorry
+
+lemma LASTN_APPEND2: "n <= length l2 ==> LASTN n (l1 @ l2) = LASTN n l2"
+  sorry
+
+lemma LASTN_APPEND1: "length l2 <= n ==> LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2"
+  sorry
+
+lemma LASTN_MAP: "n <= length l ==> LASTN n (map f l) = map f (LASTN n l)"
+  sorry
+
+lemma BUTLASTN_MAP: "n <= length l ==> BUTLASTN n (map f l) = map f (BUTLASTN n l)"
+  sorry
+
+lemma ALL_EL_LASTN: "[| list_all P l; m <= length l |] ==> list_all P (LASTN m l)"
+  sorry
+
+lemma ALL_EL_BUTLASTN: "[| list_all P l; m <= length l |] ==> list_all P (BUTLASTN m l)"
+  sorry
+
+lemma LENGTH_FIRSTN: "n <= length l ==> length (FIRSTN n l) = n"
+  sorry
+
+lemma FIRSTN_FIRSTN: "[| m <= length l; n <= m |] ==> FIRSTN n (FIRSTN m l) = FIRSTN n l"
+  sorry
+
+lemma LENGTH_BUTFIRSTN: "n <= length l ==> length (BUTFIRSTN n l) = length l - n"
+  sorry
+
+lemma BUTFIRSTN_LENGTH_NIL: "BUTFIRSTN (length l) l = []"
+  sorry
+
+lemma BUTFIRSTN_APPEND1: "n <= length l1 ==> BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2"
+  sorry
+
+lemma BUTFIRSTN_APPEND2: "length l1 <= n ==> BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2"
+  sorry
+
+lemma BUTFIRSTN_BUTFIRSTN: "n + m <= length l ==> BUTFIRSTN n (BUTFIRSTN m l) = BUTFIRSTN (n + m) l"
+  sorry
+
+lemma APPEND_FIRSTN_BUTFIRSTN: "n <= length l ==> FIRSTN n l @ BUTFIRSTN n l = l"
+  sorry
+
+lemma LASTN_SEG: "n <= length l ==> LASTN n l = SEG n (length l - n) l"
+  sorry
+
+lemma FIRSTN_SEG: "n <= length l ==> FIRSTN n l = SEG n 0 l"
+  sorry
+
+lemma BUTFIRSTN_SEG: "n <= length l ==> BUTFIRSTN n l = SEG (length l - n) n l"
+  sorry
+
+lemma BUTFIRSTN_SNOC: "n <= length l ==> BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l)"
+  sorry
+
+lemma APPEND_BUTLASTN_BUTFIRSTN: "m + n = length l ==> BUTLASTN m l @ BUTFIRSTN n l = l"
+  sorry
+
+lemma SEG_SEG: "n1 + m1 <= length l & n2 + m2 <= n1
+==> SEG n2 m2 (SEG n1 m1 l) = SEG n2 (m1 + m2) l"
+  sorry
+
+lemma SEG_APPEND1: "n + m <= length l1 ==> SEG n m (l1 @ l2) = SEG n m l1"
+  sorry
+
+lemma SEG_APPEND2: "length l1 <= m & n <= length l2
+==> SEG n m (l1 @ l2) = SEG n (m - length l1) l2"
+  sorry
+
+lemma SEG_FIRSTN_BUTFISTN: "n + m <= length l ==> SEG n m l = FIRSTN n (BUTFIRSTN m l)"
+  sorry
+
+lemma SEG_APPEND: "m < length l1 & length l1 <= n + m & n + m <= length l1 + length l2
+==> SEG n m (l1 @ l2) =
+    SEG (length l1 - m) m l1 @ SEG (n + m - length l1) 0 l2"
+  sorry
+
+lemma SEG_LENGTH_SNOC: "SEG 1 (length x) (SNOC xa x) = [xa]"
+  sorry
+
+lemma SEG_SNOC: "n + m <= length l ==> SEG n m (SNOC x l) = SEG n m l"
+  sorry
+
+lemma ELL_SEG: "n < length l ==> ELL n l = hd (SEG 1 (PRE (length l - n)) l)"
+  sorry
+
+lemma SNOC_FOLDR: "SNOC x l = foldr op # l [x]"
+  sorry
+
+lemma IS_EL_FOLDR_MAP: "List.member xa x = foldr op | (map (op = x) xa) False"
+  sorry
+
+lemma IS_EL_FOLDL_MAP: "List.member xa x = foldl op | False (map (op = x) xa)"
+  sorry
+
+lemma FILTER_FILTER: "filter P (filter Q l) = [x<-l. P x & Q x]"
+  sorry
+
+lemma FCOMM_FOLDR_FLAT: "[| FCOMM g f; LEFT_ID g e |]
+==> foldr f (concat l) e = foldr g (map (FOLDR f e) l) e"
+  sorry
+
+lemma FCOMM_FOLDL_FLAT: "[| FCOMM f g; RIGHT_ID g e |]
+==> foldl f e (concat l) = foldl g e (map (foldl f e) l)"
+  sorry
+
+lemma FOLDR_MAP_REVERSE: "(!!(a::'a) (b::'a) c::'a. (f::'a => 'a => 'a) a (f b c) = f b (f a c))
+==> foldr f (map (g::'b => 'a) (rev (l::'b list))) (e::'a) =
+    foldr f (map g l) e"
+  sorry
+
+lemma FOLDR_FILTER_REVERSE: "(!!(a::'a) (b::'a) c::'a. (f::'a => 'a => 'a) a (f b c) = f b (f a c))
+==> foldr f (filter (P::'a => bool) (rev (l::'a list))) (e::'a) =
+    foldr f (filter P l) e"
+  sorry
+
+lemma COMM_ASSOC_FOLDR_REVERSE: "[| COMM f; ASSOC f |] ==> foldr f (rev l) e = foldr f l e"
+  sorry
+
+lemma COMM_ASSOC_FOLDL_REVERSE: "[| COMM f; ASSOC f |] ==> foldl f e (rev l) = foldl f e l"
+  sorry
+
+lemma ELL_LAST: "~ List.null l ==> ELL 0 l = last l"
+  sorry
+
+lemma ELL_0_SNOC: "ELL 0 (SNOC x l) = x"
+  sorry
+
+lemma ELL_SNOC: "0 < n ==> ELL n (SNOC x l) = ELL (PRE n) l"
+  sorry
+
+lemma ELL_SUC_SNOC: "ELL (Suc n) (SNOC x xa) = ELL n xa"
+  sorry
+
+lemma ELL_CONS: "n < length l ==> ELL n (x # l) = ELL n l"
+  sorry
+
+lemma ELL_LENGTH_CONS: "ELL (length l) (x # l) = x"
+  sorry
+
+lemma ELL_LENGTH_SNOC: "ELL (length l) (SNOC x l) = (if List.null l then x else hd l)"
+  sorry
+
+lemma ELL_APPEND2: "n < length l2 ==> ELL n (l1 @ l2) = ELL n l2"
+  sorry
+
+lemma ELL_APPEND1: "length l2 <= n ==> ELL n (l1 @ l2) = ELL (n - length l2) l1"
+  sorry
+
+lemma ELL_PRE_LENGTH: "l ~= [] ==> ELL (PRE (length l)) l = hd l"
+  sorry
+
+lemma EL_LENGTH_SNOC: "EL (length l) (SNOC x l) = x"
+  sorry
+
+lemma EL_PRE_LENGTH: "l ~= [] ==> EL (PRE (length l)) l = last l"
+  sorry
+
+lemma EL_SNOC: "n < length l ==> EL n (SNOC x l) = EL n l"
+  sorry
+
+lemma EL_ELL: "n < length l ==> EL n l = ELL (PRE (length l - n)) l"
+  sorry
+
+lemma EL_LENGTH_APPEND: "~ List.null l2 ==> EL (length l1) (l1 @ l2) = hd l2"
+  sorry
+
+lemma ELL_EL: "n < length l ==> ELL n l = EL (PRE (length l - n)) l"
+  sorry
+
+lemma ELL_MAP: "n < length l ==> ELL n (map f l) = f (ELL n l)"
+  sorry
+
+lemma LENGTH_BUTLAST: "l ~= [] ==> length (butlast l) = PRE (length l)"
+  sorry
+
+lemma BUTFIRSTN_LENGTH_APPEND: "BUTFIRSTN (length l1) (l1 @ l2) = l2"
+  sorry
+
+lemma FIRSTN_APPEND1: "n <= length l1 ==> FIRSTN n (l1 @ l2) = FIRSTN n l1"
+  sorry
+
+lemma FIRSTN_APPEND2: "length l1 <= n ==> FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2"
+  sorry
+
+lemma FIRSTN_LENGTH_APPEND: "FIRSTN (length l1) (l1 @ l2) = l1"
+  sorry
+
+lemma REVERSE_FLAT: "rev (concat l) = concat (rev (map rev l))"
+  sorry
+
+lemma MAP_FILTER: "(!!x. P (f x) = P x) ==> map f (filter P l) = filter P (map f l)"
+  sorry
+
+lemma FLAT_REVERSE: "concat (rev l) = rev (concat (map rev l))"
+  sorry
+
+lemma FLAT_FLAT: "concat (concat l) = concat (map concat l)"
+  sorry
+
+lemma ALL_EL_SEG: "[| list_all P l; m + k <= length l |] ==> list_all P (SEG m k l)"
+  sorry
+
+lemma ALL_EL_FIRSTN: "[| list_all P l; m <= length l |] ==> list_all P (FIRSTN m l)"
+  sorry
+
+lemma ALL_EL_BUTFIRSTN: "[| list_all P l; m <= length l |] ==> list_all P (BUTFIRSTN m l)"
+  sorry
+
+lemma SOME_EL_SEG: "[| m + k <= length l; list_ex P (SEG m k l) |] ==> list_ex P l"
+  sorry
+
+lemma SOME_EL_FIRSTN: "[| m <= length l; list_ex P (FIRSTN m l) |] ==> list_ex P l"
+  sorry
+
+lemma SOME_EL_BUTFIRSTN: "[| m <= length l; list_ex P (BUTFIRSTN m l) |] ==> list_ex P l"
+  sorry
+
+lemma SOME_EL_LASTN: "[| m <= length l; list_ex P (LASTN m l) |] ==> list_ex P l"
+  sorry
+
+lemma SOME_EL_BUTLASTN: "[| m <= length l; list_ex P (BUTLASTN m l) |] ==> list_ex P l"
+  sorry
+
+lemma IS_EL_REVERSE: "List.member (rev l) x = List.member l x"
+  sorry
+
+lemma IS_EL_FILTER: "P x ==> List.member (filter P l) x = List.member l x"
+  sorry
+
+lemma IS_EL_SEG: "[| n + m <= length l; List.member (SEG n m l) x |] ==> List.member l x"
+  sorry
+
+lemma IS_EL_SOME_EL: "List.member l x = list_ex (op = x) l"
+  sorry
+
+lemma IS_EL_FIRSTN: "[| x <= length xa; List.member (FIRSTN x xa) xb |] ==> List.member xa xb"
+  sorry
+
+lemma IS_EL_BUTFIRSTN: "[| x <= length xa; List.member (BUTFIRSTN x xa) xb |] ==> List.member xa xb"
+  sorry
+
+lemma IS_EL_BUTLASTN: "[| x <= length xa; List.member (BUTLASTN x xa) xb |] ==> List.member xa xb"
+  sorry
+
+lemma IS_EL_LASTN: "[| x <= length xa; List.member (LASTN x xa) xb |] ==> List.member xa xb"
+  sorry
+
+lemma ZIP_SNOC: "length l1 = length l2
+==> zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2)"
+  sorry
+
+lemma UNZIP_SNOC: "unzip (SNOC x l) =
+(SNOC (fst x) (fst (unzip l)), SNOC (snd x) (snd (unzip l)))"
+  sorry
+
+lemma LENGTH_UNZIP_FST: "length (UNZIP_FST x) = length x"
+  sorry
+
+lemma LENGTH_UNZIP_SND: "length (UNZIP_SND (x::('a * 'b) list)) = length x"
+  sorry
+
+lemma SUM_APPEND: "HOL4Compat.sum (l1 @ l2) = HOL4Compat.sum l1 + HOL4Compat.sum l2"
+  sorry
+
+lemma SUM_REVERSE: "HOL4Compat.sum (rev l) = HOL4Compat.sum l"
+  sorry
+
+lemma SUM_FLAT: "HOL4Compat.sum (concat l) = HOL4Compat.sum (map HOL4Compat.sum l)"
+  sorry
+
+lemma EL_APPEND1: "n < length l1 ==> EL n (l1 @ l2) = EL n l1"
+  sorry
+
+lemma EL_APPEND2: "length l1 <= n ==> EL n (l1 @ l2) = EL (n - length l1) l2"
+  sorry
+
+lemma EL_MAP: "n < length l ==> EL n (map f l) = f (EL n l)"
+  sorry
+
+lemma EL_CONS: "0 < n ==> EL n (x # l) = EL (PRE n) l"
+  sorry
+
+lemma EL_SEG: "n < length l ==> EL n l = hd (SEG 1 n l)"
+  sorry
+
+lemma EL_IS_EL: "n < length l ==> List.member l (EL n l)"
+  sorry
+
+lemma TL_SNOC: "tl (SNOC x l) = (if List.null l then [] else SNOC x (tl l))"
+  sorry
+
+lemma EL_REVERSE: "n < length l ==> EL n (rev l) = EL (PRE (length l - n)) l"
+  sorry
+
+lemma EL_REVERSE_ELL: "n < length l ==> EL n (rev l) = ELL n l"
+  sorry
+
+lemma ELL_LENGTH_APPEND: "~ List.null l1 ==> ELL (length l2) (l1 @ l2) = last l1"
+  sorry
+
+lemma ELL_IS_EL: "n < length l ==> List.member l (ELL n l)"
+  sorry
+
+lemma ELL_REVERSE: "n < length l ==> ELL n (rev l) = ELL (PRE (length l - n)) l"
+  sorry
+
+lemma ELL_REVERSE_EL: "n < length l ==> ELL n (rev l) = EL n l"
+  sorry
+
+lemma FIRSTN_BUTLASTN: "n <= length l ==> FIRSTN n l = BUTLASTN (length l - n) l"
+  sorry
+
+lemma BUTLASTN_FIRSTN: "n <= length l ==> BUTLASTN n l = FIRSTN (length l - n) l"
+  sorry
+
+lemma LASTN_BUTFIRSTN: "n <= length l ==> LASTN n l = BUTFIRSTN (length l - n) l"
+  sorry
+
+lemma BUTFIRSTN_LASTN: "n <= length l ==> BUTFIRSTN n l = LASTN (length l - n) l"
+  sorry
+
+lemma SEG_LASTN_BUTLASTN: "n + m <= length l ==> SEG n m l = LASTN n (BUTLASTN (length l - (n + m)) l)"
+  sorry
+
+lemma BUTFIRSTN_REVERSE: "n <= length l ==> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)"
+  sorry
+
+lemma BUTLASTN_REVERSE: "n <= length l ==> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)"
+  sorry
+
+lemma LASTN_REVERSE: "n <= length l ==> LASTN n (rev l) = rev (FIRSTN n l)"
+  sorry
+
+lemma FIRSTN_REVERSE: "n <= length l ==> FIRSTN n (rev l) = rev (LASTN n l)"
+  sorry
+
+lemma SEG_REVERSE: "n + m <= length l ==> SEG n m (rev l) = rev (SEG n (length l - (n + m)) l)"
+  sorry
+
+lemma LENGTH_GENLIST: "length (GENLIST f n) = n"
+  sorry
+
+lemma LENGTH_REPLICATE: "length (REPLICATE n x) = n"
+  sorry
+
+lemma IS_EL_REPLICATE: "0 < n ==> List.member (REPLICATE n x) x"
+  sorry
+
+lemma ALL_EL_REPLICATE: "list_all (op = x) (REPLICATE n x)"
+  sorry
+
+lemma AND_EL_FOLDL: "AND_EL l = foldl op & True l"
+  sorry
+
+lemma AND_EL_FOLDR: "AND_EL l = foldr op & l True"
+  sorry
+
+lemma OR_EL_FOLDL: "OR_EL l = foldl op | False l"
+  sorry
+
+lemma OR_EL_FOLDR: "OR_EL l = foldr op | l False"
+  sorry
 
 ;end_setup
 
 ;setup_theory state_transformer
 
-definition UNIT :: "'b => 'a => 'b * 'a" where 
+definition
+  UNIT :: "'b => 'a => 'b * 'a"  where
   "(op ==::('b::type => 'a::type => 'b::type * 'a::type)
         => ('b::type => 'a::type => 'b::type * 'a::type) => prop)
  (UNIT::'b::type => 'a::type => 'b::type * 'a::type)
  (Pair::'b::type => 'a::type => 'b::type * 'a::type)"
 
-lemma UNIT_DEF: "ALL x::'b::type. UNIT x = Pair x"
-  by (import state_transformer UNIT_DEF)
-
-definition BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" where 
-  "(op ==::(('a::type => 'b::type * 'a::type)
-         => ('b::type => 'a::type => 'c::type * 'a::type)
-            => 'a::type => 'c::type * 'a::type)
-        => (('a::type => 'b::type * 'a::type)
-            => ('b::type => 'a::type => 'c::type * 'a::type)
-               => 'a::type => 'c::type * 'a::type)
-           => prop)
- (BIND::('a::type => 'b::type * 'a::type)
-        => ('b::type => 'a::type => 'c::type * 'a::type)
-           => 'a::type => 'c::type * 'a::type)
- (%(g::'a::type => 'b::type * 'a::type)
-     f::'b::type => 'a::type => 'c::type * 'a::type.
-     (op o::('b::type * 'a::type => 'c::type * 'a::type)
-            => ('a::type => 'b::type * 'a::type)
-               => 'a::type => 'c::type * 'a::type)
-      ((split::('b::type => 'a::type => 'c::type * 'a::type)
-               => 'b::type * 'a::type => 'c::type * 'a::type)
-        f)
-      g)"
-
-lemma BIND_DEF: "(All::(('a::type => 'b::type * 'a::type) => bool) => bool)
- (%g::'a::type => 'b::type * 'a::type.
-     (All::(('b::type => 'a::type => 'c::type * 'a::type) => bool) => bool)
-      (%f::'b::type => 'a::type => 'c::type * 'a::type.
-          (op =::('a::type => 'c::type * 'a::type)
-                 => ('a::type => 'c::type * 'a::type) => bool)
-           ((BIND::('a::type => 'b::type * 'a::type)
-                   => ('b::type => 'a::type => 'c::type * 'a::type)
-                      => 'a::type => 'c::type * 'a::type)
-             g f)
-           ((op o::('b::type * 'a::type => 'c::type * 'a::type)
-                   => ('a::type => 'b::type * 'a::type)
-                      => 'a::type => 'c::type * 'a::type)
-             ((split::('b::type => 'a::type => 'c::type * 'a::type)
-                      => 'b::type * 'a::type => 'c::type * 'a::type)
-               f)
-             g)))"
-  by (import state_transformer BIND_DEF)
-
-definition MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" where 
-  "MMAP ==
-%(f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type.
-   BIND m (UNIT o f)"
-
-lemma MMAP_DEF: "ALL (f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type.
-   MMAP f m = BIND m (UNIT o f)"
-  by (import state_transformer MMAP_DEF)
-
-definition JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" where 
-  "JOIN ==
-%z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. BIND z I"
-
-lemma JOIN_DEF: "ALL z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type.
-   JOIN z = BIND z I"
-  by (import state_transformer JOIN_DEF)
-
-lemma BIND_LEFT_UNIT: "ALL (k::'a::type => 'b::type => 'c::type * 'b::type) x::'a::type.
-   BIND (UNIT x) k = k x"
-  by (import state_transformer BIND_LEFT_UNIT)
-
-lemma UNIT_UNCURRY: "ALL x::'a::type * 'b::type. split UNIT x = x"
-  by (import state_transformer UNIT_UNCURRY)
-
-lemma BIND_RIGHT_UNIT: "ALL k::'a::type => 'b::type * 'a::type. BIND k UNIT = k"
-  by (import state_transformer BIND_RIGHT_UNIT)
-
-lemma BIND_ASSOC: "ALL (x::'a::type => 'b::type * 'a::type)
-   (xa::'b::type => 'a::type => 'c::type * 'a::type)
-   xb::'c::type => 'a::type => 'd::type * 'a::type.
-   BIND x (%a::'b::type. BIND (xa a) xb) = BIND (BIND x xa) xb"
-  by (import state_transformer BIND_ASSOC)
+lemma UNIT_DEF: "UNIT x = Pair x"
+  sorry
+
+definition
+  BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a"  where
+  "BIND == %g f. (%(x, y). f x y) o g"
+
+lemma BIND_DEF: "BIND (g::'a => 'b * 'a) (f::'b => 'a => 'c * 'a) =
+(%(x::'b, y::'a). f x y) o g"
+  sorry
+
+definition
+  MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a"  where
+  "MMAP == %(f::'c => 'b) m::'a => 'c * 'a. BIND m (UNIT o f)"
+
+lemma MMAP_DEF: "MMAP f m = BIND m (UNIT o f)"
+  sorry
+
+definition
+  JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a"  where
+  "JOIN == %z. BIND z I"
+
+lemma JOIN_DEF: "JOIN z = BIND z I"
+  sorry
+
+lemma BIND_LEFT_UNIT: "BIND (UNIT (x::'a)) (k::'a => 'b => 'c * 'b) = k x"
+  sorry
+
+lemma UNIT_UNCURRY: "prod_case UNIT x = x"
+  sorry
+
+lemma BIND_RIGHT_UNIT: "BIND k UNIT = k"
+  sorry
+
+lemma BIND_ASSOC: "BIND (x::'a => 'b * 'a)
+ (%a::'b. BIND ((xa::'b => 'a => 'c * 'a) a) (xb::'c => 'a => 'd * 'a)) =
+BIND (BIND x xa) xb"
+  sorry
 
 lemma MMAP_ID: "MMAP I = I"
-  by (import state_transformer MMAP_ID)
-
-lemma MMAP_COMP: "ALL (f::'c::type => 'd::type) g::'b::type => 'c::type.
-   MMAP (f o g) = MMAP f o MMAP g"
-  by (import state_transformer MMAP_COMP)
-
-lemma MMAP_UNIT: "ALL f::'b::type => 'c::type. MMAP f o UNIT = UNIT o f"
-  by (import state_transformer MMAP_UNIT)
-
-lemma MMAP_JOIN: "ALL f::'b::type => 'c::type. MMAP f o JOIN = JOIN o MMAP (MMAP f)"
-  by (import state_transformer MMAP_JOIN)
+  sorry
+
+lemma MMAP_COMP: "MMAP ((f::'c => 'd) o (g::'b => 'c)) = MMAP f o MMAP g"
+  sorry
+
+lemma MMAP_UNIT: "MMAP (f::'b => 'c) o UNIT = UNIT o f"
+  sorry
+
+lemma MMAP_JOIN: "MMAP f o JOIN = JOIN o MMAP (MMAP f)"
+  sorry
 
 lemma JOIN_UNIT: "JOIN o UNIT = I"
-  by (import state_transformer JOIN_UNIT)
+  sorry
 
 lemma JOIN_MMAP_UNIT: "JOIN o MMAP UNIT = I"
-  by (import state_transformer JOIN_MMAP_UNIT)
+  sorry
 
 lemma JOIN_MAP_JOIN: "JOIN o MMAP JOIN = JOIN o JOIN"
-  by (import state_transformer JOIN_MAP_JOIN)
-
-lemma JOIN_MAP: "ALL (x::'a::type => 'b::type * 'a::type)
-   xa::'b::type => 'a::type => 'c::type * 'a::type.
-   BIND x xa = JOIN (MMAP xa x)"
-  by (import state_transformer JOIN_MAP)
-
-lemma FST_o_UNIT: "ALL x::'a::type. fst o UNIT x = K x"
-  by (import state_transformer FST_o_UNIT)
-
-lemma SND_o_UNIT: "ALL x::'a::type. snd o UNIT x = I"
-  by (import state_transformer SND_o_UNIT)
-
-lemma FST_o_MMAP: "ALL (x::'a::type => 'b::type) xa::'c::type => 'a::type * 'c::type.
-   fst o MMAP x xa = x o (fst o xa)"
-  by (import state_transformer FST_o_MMAP)
+  sorry
+
+lemma JOIN_MAP: "BIND (x::'a => 'b * 'a) (xa::'b => 'a => 'c * 'a) = JOIN (MMAP xa x)"
+  sorry
+
+lemma FST_o_UNIT: "fst o UNIT (x::'a) = K x"
+  sorry
+
+lemma SND_o_UNIT: "snd o UNIT (x::'a) = I"
+  sorry
+
+lemma FST_o_MMAP: "fst o MMAP (x::'a => 'b) (xa::'c => 'a * 'c) = x o (fst o xa)"
+  sorry
 
 ;end_setup
 
--- a/src/HOL/Import/HOL/HOL4Prob.thy	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/HOL4Prob.thy	Tue Sep 06 22:41:35 2011 -0700
@@ -4,357 +4,211 @@
 
 ;setup_theory prob_extra
 
-lemma BOOL_BOOL_CASES_THM: "ALL f::bool => bool.
-   f = (%b::bool. False) |
-   f = (%b::bool. True) | f = (%b::bool. b) | f = Not"
+lemma BOOL_BOOL_CASES_THM: "f = (%b. False) | f = (%b. True) | f = (%b. b) | f = Not"
   by (import prob_extra BOOL_BOOL_CASES_THM)
 
 lemma EVEN_ODD_BASIC: "EVEN 0 & ~ EVEN 1 & EVEN 2 & ~ ODD 0 & ODD 1 & ~ ODD 2"
   by (import prob_extra EVEN_ODD_BASIC)
 
-lemma EVEN_ODD_EXISTS_EQ: "ALL n::nat.
-   EVEN n = (EX m::nat. n = 2 * m) & ODD n = (EX m::nat. n = Suc (2 * m))"
+lemma EVEN_ODD_EXISTS_EQ: "EVEN n = (EX m. n = 2 * m) & ODD n = (EX m. n = Suc (2 * m))"
   by (import prob_extra EVEN_ODD_EXISTS_EQ)
 
-lemma DIV_THEN_MULT: "ALL (p::nat) q::nat. Suc q * (p div Suc q) <= p"
+lemma DIV_THEN_MULT: "Suc q * (p div Suc q) <= p"
   by (import prob_extra DIV_THEN_MULT)
 
-lemma DIV_TWO_UNIQUE: "ALL (n::nat) (q::nat) r::nat.
-   n = 2 * q + r & (r = 0 | r = 1) --> q = n div 2 & r = n mod 2"
+lemma DIV_TWO_UNIQUE: "(n::nat) = (2::nat) * (q::nat) + (r::nat) & (r = (0::nat) | r = (1::nat))
+==> q = n div (2::nat) & r = n mod (2::nat)"
   by (import prob_extra DIV_TWO_UNIQUE)
 
-lemma DIVISION_TWO: "ALL n::nat. n = 2 * (n div 2) + n mod 2 & (n mod 2 = 0 | n mod 2 = 1)"
+lemma DIVISION_TWO: "(n::nat) = (2::nat) * (n div (2::nat)) + n mod (2::nat) &
+(n mod (2::nat) = (0::nat) | n mod (2::nat) = (1::nat))"
   by (import prob_extra DIVISION_TWO)
 
-lemma DIV_TWO: "ALL n::nat. n = 2 * (n div 2) + n mod 2"
+lemma DIV_TWO: "(n::nat) = (2::nat) * (n div (2::nat)) + n mod (2::nat)"
   by (import prob_extra DIV_TWO)
 
-lemma MOD_TWO: "ALL n::nat. n mod 2 = (if EVEN n then 0 else 1)"
+lemma MOD_TWO: "n mod 2 = (if EVEN n then 0 else 1)"
   by (import prob_extra MOD_TWO)
 
-lemma DIV_TWO_BASIC: "(op &::bool => bool => bool)
- ((op =::nat => nat => bool)
-   ((op div::nat => nat => nat) (0::nat)
-     ((number_of \<Colon> int => nat)
-       ((Int.Bit0 \<Colon> int => int)
-         ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-   (0::nat))
- ((op &::bool => bool => bool)
-   ((op =::nat => nat => bool)
-     ((op div::nat => nat => nat) (1::nat)
-       ((number_of \<Colon> int => nat)
-         ((Int.Bit0 \<Colon> int => int)
-           ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-     (0::nat))
-   ((op =::nat => nat => bool)
-     ((op div::nat => nat => nat)
-       ((number_of \<Colon> int => nat)
-         ((Int.Bit0 \<Colon> int => int)
-           ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-       ((number_of \<Colon> int => nat)
-         ((Int.Bit0 \<Colon> int => int)
-           ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-     (1::nat)))"
+lemma DIV_TWO_BASIC: "(0::nat) div (2::nat) = (0::nat) &
+(1::nat) div (2::nat) = (0::nat) & (2::nat) div (2::nat) = (1::nat)"
   by (import prob_extra DIV_TWO_BASIC)
 
-lemma DIV_TWO_MONO: "ALL (m::nat) n::nat. m div 2 < n div 2 --> m < n"
+lemma DIV_TWO_MONO: "(m::nat) div (2::nat) < (n::nat) div (2::nat) ==> m < n"
   by (import prob_extra DIV_TWO_MONO)
 
-lemma DIV_TWO_MONO_EVEN: "ALL (m::nat) n::nat. EVEN n --> (m div 2 < n div 2) = (m < n)"
+lemma DIV_TWO_MONO_EVEN: "EVEN n ==> (m div 2 < n div 2) = (m < n)"
   by (import prob_extra DIV_TWO_MONO_EVEN)
 
-lemma DIV_TWO_CANCEL: "ALL n::nat. 2 * n div 2 = n & Suc (2 * n) div 2 = n"
+lemma DIV_TWO_CANCEL: "2 * n div 2 = n & Suc (2 * n) div 2 = n"
   by (import prob_extra DIV_TWO_CANCEL)
 
-lemma EXP_DIV_TWO: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op =::nat => nat => bool)
-      ((op div::nat => nat => nat)
-        ((op ^::nat => nat => nat)
-          ((number_of \<Colon> int => nat)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-          ((Suc::nat => nat) n))
-        ((number_of \<Colon> int => nat)
-          ((Int.Bit0 \<Colon> int => int)
-            ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-      ((op ^::nat => nat => nat)
-        ((number_of \<Colon> int => nat)
-          ((Int.Bit0 \<Colon> int => int)
-            ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-        n))"
+lemma EXP_DIV_TWO: "(2::nat) ^ Suc (n::nat) div (2::nat) = (2::nat) ^ n"
   by (import prob_extra EXP_DIV_TWO)
 
-lemma EVEN_EXP_TWO: "ALL n::nat. EVEN (2 ^ n) = (n ~= 0)"
+lemma EVEN_EXP_TWO: "EVEN (2 ^ n) = (n ~= 0)"
   by (import prob_extra EVEN_EXP_TWO)
 
-lemma DIV_TWO_EXP: "ALL (n::nat) k::nat. (k div 2 < 2 ^ n) = (k < 2 ^ Suc n)"
+lemma DIV_TWO_EXP: "((k::nat) div (2::nat) < (2::nat) ^ (n::nat)) = (k < (2::nat) ^ Suc n)"
   by (import prob_extra DIV_TWO_EXP)
 
 consts
   inf :: "(real => bool) => real" 
 
 defs
-  inf_primdef: "inf == %P::real => bool. - sup (IMAGE uminus P)"
+  inf_primdef: "prob_extra.inf == %P. - real.sup (IMAGE uminus P)"
 
-lemma inf_def: "ALL P::real => bool. inf P = - sup (IMAGE uminus P)"
+lemma inf_def: "prob_extra.inf P = - real.sup (IMAGE uminus P)"
   by (import prob_extra inf_def)
 
-lemma INF_DEF_ALT: "ALL P::real => bool. inf P = - sup (%r::real. P (- r))"
+lemma INF_DEF_ALT: "prob_extra.inf P = - real.sup (%r. P (- r))"
   by (import prob_extra INF_DEF_ALT)
 
-lemma REAL_SUP_EXISTS_UNIQUE: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x <= z) -->
-   (EX! s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s))"
+lemma REAL_SUP_EXISTS_UNIQUE: "Ex (P::real => bool) & (EX z::real. ALL x::real. P x --> x <= z)
+==> EX! s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)"
   by (import prob_extra REAL_SUP_EXISTS_UNIQUE)
 
-lemma REAL_SUP_MAX: "ALL (P::real => bool) z::real.
-   P z & (ALL x::real. P x --> x <= z) --> sup P = z"
+lemma REAL_SUP_MAX: "P z & (ALL x. P x --> x <= z) ==> real.sup P = z"
   by (import prob_extra REAL_SUP_MAX)
 
-lemma REAL_INF_MIN: "ALL (P::real => bool) z::real.
-   P z & (ALL x::real. P x --> z <= x) --> inf P = z"
+lemma REAL_INF_MIN: "P z & (ALL x. P x --> z <= x) ==> prob_extra.inf P = z"
   by (import prob_extra REAL_INF_MIN)
 
-lemma HALF_POS: "(op <::real => real => bool) (0::real)
- ((op /::real => real => real) (1::real)
-   ((number_of \<Colon> int => real)
-     ((Int.Bit0 \<Colon> int => int)
-       ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))"
-  by (import prob_extra HALF_POS)
-
-lemma HALF_CANCEL: "(op =::real => real => bool)
- ((op *::real => real => real)
-   ((number_of \<Colon> int => real)
-     ((Int.Bit0 \<Colon> int => int)
-       ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-   ((op /::real => real => real) (1::real)
-     ((number_of \<Colon> int => real)
-       ((Int.Bit0 \<Colon> int => int)
-         ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))))
- (1::real)"
+lemma HALF_CANCEL: "(2::real) * ((1::real) / (2::real)) = (1::real)"
   by (import prob_extra HALF_CANCEL)
 
-lemma POW_HALF_POS: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op <::real => real => bool) (0::real)
-      ((op ^::real => nat => real)
-        ((op /::real => real => real) (1::real)
-          ((number_of \<Colon> int => real)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-        n))"
+lemma POW_HALF_POS: "(0::real) < ((1::real) / (2::real)) ^ (n::nat)"
   by (import prob_extra POW_HALF_POS)
 
-lemma POW_HALF_MONO: "(All::(nat => bool) => bool)
- (%m::nat.
-     (All::(nat => bool) => bool)
-      (%n::nat.
-          (op -->::bool => bool => bool) ((op <=::nat => nat => bool) m n)
-           ((op <=::real => real => bool)
-             ((op ^::real => nat => real)
-               ((op /::real => real => real) (1::real)
-                 ((number_of \<Colon> int => real)
-                   ((Int.Bit0 \<Colon> int => int)
-                     ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-               n)
-             ((op ^::real => nat => real)
-               ((op /::real => real => real) (1::real)
-                 ((number_of \<Colon> int => real)
-                   ((Int.Bit0 \<Colon> int => int)
-                     ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-               m))))"
+lemma POW_HALF_MONO: "(m::nat) <= (n::nat)
+==> ((1::real) / (2::real)) ^ n <= ((1::real) / (2::real)) ^ m"
   by (import prob_extra POW_HALF_MONO)
 
-lemma POW_HALF_TWICE: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op =::real => real => bool)
-      ((op ^::real => nat => real)
-        ((op /::real => real => real) (1::real)
-          ((number_of \<Colon> int => real)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-        n)
-      ((op *::real => real => real)
-        ((number_of \<Colon> int => real)
-          ((Int.Bit0 \<Colon> int => int)
-            ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-        ((op ^::real => nat => real)
-          ((op /::real => real => real) (1::real)
-            ((number_of \<Colon> int => real)
-              ((Int.Bit0 \<Colon> int => int)
-                ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-          ((Suc::nat => nat) n))))"
+lemma POW_HALF_TWICE: "((1::real) / (2::real)) ^ (n::nat) =
+(2::real) * ((1::real) / (2::real)) ^ Suc n"
   by (import prob_extra POW_HALF_TWICE)
 
-lemma X_HALF_HALF: "ALL x::real. 1 / 2 * x + 1 / 2 * x = x"
+lemma X_HALF_HALF: "(1::real) / (2::real) * (x::real) + (1::real) / (2::real) * x = x"
   by (import prob_extra X_HALF_HALF)
 
-lemma REAL_SUP_LE_X: "ALL (P::real => bool) x::real.
-   Ex P & (ALL r::real. P r --> r <= x) --> sup P <= x"
+lemma REAL_SUP_LE_X: "Ex P & (ALL r. P r --> r <= x) ==> real.sup P <= x"
   by (import prob_extra REAL_SUP_LE_X)
 
-lemma REAL_X_LE_SUP: "ALL (P::real => bool) x::real.
-   Ex P &
-   (EX z::real. ALL r::real. P r --> r <= z) &
-   (EX r::real. P r & x <= r) -->
-   x <= sup P"
+lemma REAL_X_LE_SUP: "Ex P & (EX z. ALL r. P r --> r <= z) & (EX r. P r & x <= r)
+==> x <= real.sup P"
   by (import prob_extra REAL_X_LE_SUP)
 
-lemma ABS_BETWEEN_LE: "ALL (x::real) (y::real) d::real.
-   (0 <= d & x - d <= y & y <= x + d) = (abs (y - x) <= d)"
+lemma ABS_BETWEEN_LE: "((0::real) <= (d::real) & (x::real) - d <= (y::real) & y <= x + d) =
+(abs (y - x) <= d)"
   by (import prob_extra ABS_BETWEEN_LE)
 
-lemma ONE_MINUS_HALF: "(op =::real => real => bool)
- ((op -::real => real => real) (1::real)
-   ((op /::real => real => real) (1::real)
-     ((number_of \<Colon> int => real)
-       ((Int.Bit0 \<Colon> int => int)
-         ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))))
- ((op /::real => real => real) (1::real)
-   ((number_of \<Colon> int => real)
-     ((Int.Bit0 \<Colon> int => int)
-       ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))"
+lemma ONE_MINUS_HALF: "(1::real) - (1::real) / (2::real) = (1::real) / (2::real)"
   by (import prob_extra ONE_MINUS_HALF)
 
-lemma HALF_LT_1: "(op <::real => real => bool)
- ((op /::real => real => real) (1::real)
-   ((number_of \<Colon> int => real)
-     ((Int.Bit0 \<Colon> int => int)
-       ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
- (1::real)"
+lemma HALF_LT_1: "(1::real) / (2::real) < (1::real)"
   by (import prob_extra HALF_LT_1)
 
-lemma POW_HALF_EXP: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op =::real => real => bool)
-      ((op ^::real => nat => real)
-        ((op /::real => real => real) (1::real)
-          ((number_of \<Colon> int => real)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-        n)
-      ((inverse::real => real)
-        ((real::nat => real)
-          ((op ^::nat => nat => nat)
-            ((number_of \<Colon> int => nat)
-              ((Int.Bit0 \<Colon> int => int)
-                ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-            n))))"
+lemma POW_HALF_EXP: "((1::real) / (2::real)) ^ (n::nat) = inverse (real ((2::nat) ^ n))"
   by (import prob_extra POW_HALF_EXP)
 
-lemma INV_SUC_POS: "ALL n::nat. 0 < 1 / real (Suc n)"
+lemma INV_SUC_POS: "0 < 1 / real (Suc n)"
   by (import prob_extra INV_SUC_POS)
 
-lemma INV_SUC_MAX: "ALL x::nat. 1 / real (Suc x) <= 1"
+lemma INV_SUC_MAX: "1 / real (Suc x) <= 1"
   by (import prob_extra INV_SUC_MAX)
 
-lemma INV_SUC: "ALL n::nat. 0 < 1 / real (Suc n) & 1 / real (Suc n) <= 1"
+lemma INV_SUC: "0 < 1 / real (Suc n) & 1 / real (Suc n) <= 1"
   by (import prob_extra INV_SUC)
 
-lemma ABS_UNIT_INTERVAL: "ALL (x::real) y::real.
-   0 <= x & x <= 1 & 0 <= y & y <= 1 --> abs (x - y) <= 1"
+lemma ABS_UNIT_INTERVAL: "(0::real) <= (x::real) &
+x <= (1::real) & (0::real) <= (y::real) & y <= (1::real)
+==> abs (x - y) <= (1::real)"
   by (import prob_extra ABS_UNIT_INTERVAL)
 
-lemma MEM_NIL: "ALL l::'a::type list. (ALL x::'a::type. ~ x mem l) = (l = [])"
+lemma MEM_NIL: "(ALL x. ~ List.member l x) = (l = [])"
   by (import prob_extra MEM_NIL)
 
-lemma MAP_MEM: "ALL (f::'a::type => 'b::type) (l::'a::type list) x::'b::type.
-   x mem map f l = (EX y::'a::type. y mem l & x = f y)"
+lemma MAP_MEM: "List.member (map (f::'a => 'b) (l::'a list)) (x::'b) =
+(EX y::'a. List.member l y & x = f y)"
   by (import prob_extra MAP_MEM)
 
-lemma MEM_NIL_MAP_CONS: "ALL (x::'a::type) l::'a::type list list. ~ [] mem map (op # x) l"
+lemma MEM_NIL_MAP_CONS: "~ List.member (map (op # x) l) []"
   by (import prob_extra MEM_NIL_MAP_CONS)
 
-lemma FILTER_TRUE: "ALL l::'a::type list. [x::'a::type<-l. True] = l"
+lemma FILTER_TRUE: "[x<-l. True] = l"
   by (import prob_extra FILTER_TRUE)
 
-lemma FILTER_FALSE: "ALL l::'a::type list. [x::'a::type<-l. False] = []"
+lemma FILTER_FALSE: "[x<-l. False] = []"
   by (import prob_extra FILTER_FALSE)
 
-lemma FILTER_MEM: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
-   x mem filter P l --> P x"
+lemma FILTER_MEM: "List.member (filter P l) x ==> P x"
   by (import prob_extra FILTER_MEM)
 
-lemma MEM_FILTER: "ALL (P::'a::type => bool) (l::'a::type list) x::'a::type.
-   x mem filter P l --> x mem l"
+lemma MEM_FILTER: "List.member (filter P l) x ==> List.member l x"
   by (import prob_extra MEM_FILTER)
 
-lemma FILTER_OUT_ELT: "ALL (x::'a::type) l::'a::type list. x mem l | [y::'a::type<-l. y ~= x] = l"
+lemma FILTER_OUT_ELT: "List.member l x | [y<-l. y ~= x] = l"
   by (import prob_extra FILTER_OUT_ELT)
 
-lemma IS_PREFIX_NIL: "ALL x::'a::type list. IS_PREFIX x [] & IS_PREFIX [] x = (x = [])"
+lemma IS_PREFIX_NIL: "IS_PREFIX x [] & IS_PREFIX [] x = (x = [])"
   by (import prob_extra IS_PREFIX_NIL)
 
-lemma IS_PREFIX_REFL: "ALL x::'a::type list. IS_PREFIX x x"
+lemma IS_PREFIX_REFL: "IS_PREFIX x x"
   by (import prob_extra IS_PREFIX_REFL)
 
-lemma IS_PREFIX_ANTISYM: "ALL (x::'a::type list) y::'a::type list.
-   IS_PREFIX y x & IS_PREFIX x y --> x = y"
+lemma IS_PREFIX_ANTISYM: "IS_PREFIX y x & IS_PREFIX x y ==> x = y"
   by (import prob_extra IS_PREFIX_ANTISYM)
 
-lemma IS_PREFIX_TRANS: "ALL (x::'a::type list) (y::'a::type list) z::'a::type list.
-   IS_PREFIX x y & IS_PREFIX y z --> IS_PREFIX x z"
+lemma IS_PREFIX_TRANS: "IS_PREFIX x y & IS_PREFIX y z ==> IS_PREFIX x z"
   by (import prob_extra IS_PREFIX_TRANS)
 
-lemma IS_PREFIX_BUTLAST: "ALL (x::'a::type) y::'a::type list. IS_PREFIX (x # y) (butlast (x # y))"
+lemma IS_PREFIX_BUTLAST: "IS_PREFIX (x # y) (butlast (x # y))"
   by (import prob_extra IS_PREFIX_BUTLAST)
 
-lemma IS_PREFIX_LENGTH: "ALL (x::'a::type list) y::'a::type list.
-   IS_PREFIX y x --> length x <= length y"
+lemma IS_PREFIX_LENGTH: "IS_PREFIX y x ==> length x <= length y"
   by (import prob_extra IS_PREFIX_LENGTH)
 
-lemma IS_PREFIX_LENGTH_ANTI: "ALL (x::'a::type list) y::'a::type list.
-   IS_PREFIX y x & length x = length y --> x = y"
+lemma IS_PREFIX_LENGTH_ANTI: "IS_PREFIX y x & length x = length y ==> x = y"
   by (import prob_extra IS_PREFIX_LENGTH_ANTI)
 
-lemma IS_PREFIX_SNOC: "ALL (x::'a::type) (y::'a::type list) z::'a::type list.
-   IS_PREFIX (SNOC x y) z = (IS_PREFIX y z | z = SNOC x y)"
+lemma IS_PREFIX_SNOC: "IS_PREFIX (SNOC x y) z = (IS_PREFIX y z | z = SNOC x y)"
   by (import prob_extra IS_PREFIX_SNOC)
 
-lemma FOLDR_MAP: "ALL (f::'b::type => 'c::type => 'c::type) (e::'c::type)
-   (g::'a::type => 'b::type) l::'a::type list.
-   foldr f (map g l) e = foldr (%x::'a::type. f (g x)) l e"
+lemma FOLDR_MAP: "foldr (f::'b => 'c => 'c) (map (g::'a => 'b) (l::'a list)) (e::'c) =
+foldr (%x::'a. f (g x)) l e"
   by (import prob_extra FOLDR_MAP)
 
-lemma LAST_MEM: "ALL (h::'a::type) t::'a::type list. last (h # t) mem h # t"
+lemma LAST_MEM: "List.member (h # t) (last (h # t))"
   by (import prob_extra LAST_MEM)
 
-lemma LAST_MAP_CONS: "ALL (b::bool) (h::bool list) t::bool list list.
-   EX x::bool list. last (map (op # b) (h # t)) = b # x"
+lemma LAST_MAP_CONS: "EX x::bool list.
+   last (map (op # (b::bool)) ((h::bool list) # (t::bool list list))) =
+   b # x"
   by (import prob_extra LAST_MAP_CONS)
 
-lemma EXISTS_LONGEST: "ALL (x::'a::type list) y::'a::type list list.
-   EX z::'a::type list.
-      z mem x # y &
-      (ALL w::'a::type list. w mem x # y --> length w <= length z)"
+lemma EXISTS_LONGEST: "EX z. List.member (x # y) z &
+      (ALL w. List.member (x # y) w --> length w <= length z)"
   by (import prob_extra EXISTS_LONGEST)
 
-lemma UNION_DEF_ALT: "ALL (s::'a::type => bool) t::'a::type => bool.
-   pred_set.UNION s t = (%x::'a::type. s x | t x)"
+lemma UNION_DEF_ALT: "pred_set.UNION s t = (%x. s x | t x)"
   by (import prob_extra UNION_DEF_ALT)
 
-lemma INTER_UNION_RDISTRIB: "ALL (p::'a::type => bool) (q::'a::type => bool) r::'a::type => bool.
-   pred_set.INTER (pred_set.UNION p q) r =
-   pred_set.UNION (pred_set.INTER p r) (pred_set.INTER q r)"
+lemma INTER_UNION_RDISTRIB: "pred_set.INTER (pred_set.UNION p q) r =
+pred_set.UNION (pred_set.INTER p r) (pred_set.INTER q r)"
   by (import prob_extra INTER_UNION_RDISTRIB)
 
-lemma SUBSET_EQ: "ALL (x::'a::type => bool) xa::'a::type => bool.
-   (x = xa) = (SUBSET x xa & SUBSET xa x)"
+lemma SUBSET_EQ: "(x = xa) = (SUBSET x xa & SUBSET xa x)"
   by (import prob_extra SUBSET_EQ)
 
-lemma INTER_IS_EMPTY: "ALL (s::'a::type => bool) t::'a::type => bool.
-   (pred_set.INTER s t = EMPTY) = (ALL x::'a::type. ~ s x | ~ t x)"
+lemma INTER_IS_EMPTY: "(pred_set.INTER s t = EMPTY) = (ALL x. ~ s x | ~ t x)"
   by (import prob_extra INTER_IS_EMPTY)
 
-lemma UNION_DISJOINT_SPLIT: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
-   pred_set.UNION s t = pred_set.UNION s u &
-   pred_set.INTER s t = EMPTY & pred_set.INTER s u = EMPTY -->
-   t = u"
+lemma UNION_DISJOINT_SPLIT: "pred_set.UNION s t = pred_set.UNION s u &
+pred_set.INTER s t = EMPTY & pred_set.INTER s u = EMPTY
+==> t = u"
   by (import prob_extra UNION_DISJOINT_SPLIT)
 
-lemma GSPEC_DEF_ALT: "ALL f::'a::type => 'b::type * bool.
-   GSPEC f = (%v::'b::type. EX x::'a::type. (v, True) = f x)"
+lemma GSPEC_DEF_ALT: "GSPEC (f::'a => 'b * bool) = (%v::'b. EX x::'a. (v, True) = f x)"
   by (import prob_extra GSPEC_DEF_ALT)
 
 ;end_setup
@@ -365,158 +219,56 @@
   alg_twin :: "bool list => bool list => bool" 
 
 defs
-  alg_twin_primdef: "alg_twin ==
-%(x::bool list) y::bool list.
-   EX l::bool list. x = SNOC True l & y = SNOC False l"
+  alg_twin_primdef: "alg_twin == %x y. EX l. x = SNOC True l & y = SNOC False l"
 
-lemma alg_twin_def: "ALL (x::bool list) y::bool list.
-   alg_twin x y = (EX l::bool list. x = SNOC True l & y = SNOC False l)"
+lemma alg_twin_def: "alg_twin x y = (EX l. x = SNOC True l & y = SNOC False l)"
   by (import prob_canon alg_twin_def)
 
-definition alg_order_tupled :: "bool list * bool list => bool" where 
-  "(op ==::(bool list * bool list => bool)
-        => (bool list * bool list => bool) => prop)
- (alg_order_tupled::bool list * bool list => bool)
- ((WFREC::(bool list * bool list => bool list * bool list => bool)
-          => ((bool list * bool list => bool)
-              => bool list * bool list => bool)
-             => bool list * bool list => bool)
-   ((Eps::((bool list * bool list => bool list * bool list => bool) => bool)
-          => bool list * bool list => bool list * bool list => bool)
-     (%R::bool list * bool list => bool list * bool list => bool.
-         (op &::bool => bool => bool)
-          ((WF::(bool list * bool list => bool list * bool list => bool)
-                => bool)
-            R)
-          ((All::(bool => bool) => bool)
-            (%h'::bool.
-                (All::(bool => bool) => bool)
-                 (%h::bool.
-                     (All::(bool list => bool) => bool)
-                      (%t'::bool list.
-                          (All::(bool list => bool) => bool)
-                           (%t::bool list.
-                               R ((Pair::bool list
-   => bool list => bool list * bool list)
-                                   t t')
-                                ((Pair::bool list
-  => bool list => bool list * bool list)
-                                  ((op #::bool => bool list => bool list) h
-                                    t)
-                                  ((op #::bool => bool list => bool list) h'
-                                    t')))))))))
-   (%alg_order_tupled::bool list * bool list => bool.
-       (split::(bool list => bool list => bool)
-               => bool list * bool list => bool)
-        (%(v::bool list) v1::bool list.
-            (list_case::bool
-                        => (bool => bool list => bool) => bool list => bool)
-             ((list_case::bool
-                          => (bool => bool list => bool)
-                             => bool list => bool)
-               (True::bool) (%(v8::bool) v9::bool list. True::bool) v1)
-             (%(v4::bool) v5::bool list.
-                 (list_case::bool
-                             => (bool => bool list => bool)
-                                => bool list => bool)
-                  (False::bool)
-                  (%(v10::bool) v11::bool list.
-                      (op |::bool => bool => bool)
-                       ((op &::bool => bool => bool)
-                         ((op =::bool => bool => bool) v4 (True::bool))
-                         ((op =::bool => bool => bool) v10 (False::bool)))
-                       ((op &::bool => bool => bool)
-                         ((op =::bool => bool => bool) v4 v10)
-                         (alg_order_tupled
-                           ((Pair::bool list
-                                   => bool list => bool list * bool list)
-                             v5 v11))))
-                  v1)
-             v)))"
+definition
+  alg_order_tupled :: "bool list * bool list => bool"  where
+  "alg_order_tupled ==
+WFREC (SOME R. WF R & (ALL h' h t' t. R (t, t') (h # t, h' # t')))
+ (%alg_order_tupled (v, v1).
+     case v of [] => case v1 of [] => True | _ => True
+     | v4 # v5 =>
+         case v1 of [] => False
+         | v10 # v11 =>
+             v4 = True & v10 = False |
+             v4 = v10 & alg_order_tupled (v5, v11))"
 
-lemma alg_order_tupled_primitive_def: "(op =::(bool list * bool list => bool)
-       => (bool list * bool list => bool) => bool)
- (alg_order_tupled::bool list * bool list => bool)
- ((WFREC::(bool list * bool list => bool list * bool list => bool)
-          => ((bool list * bool list => bool)
-              => bool list * bool list => bool)
-             => bool list * bool list => bool)
-   ((Eps::((bool list * bool list => bool list * bool list => bool) => bool)
-          => bool list * bool list => bool list * bool list => bool)
-     (%R::bool list * bool list => bool list * bool list => bool.
-         (op &::bool => bool => bool)
-          ((WF::(bool list * bool list => bool list * bool list => bool)
-                => bool)
-            R)
-          ((All::(bool => bool) => bool)
-            (%h'::bool.
-                (All::(bool => bool) => bool)
-                 (%h::bool.
-                     (All::(bool list => bool) => bool)
-                      (%t'::bool list.
-                          (All::(bool list => bool) => bool)
-                           (%t::bool list.
-                               R ((Pair::bool list
-   => bool list => bool list * bool list)
-                                   t t')
-                                ((Pair::bool list
-  => bool list => bool list * bool list)
-                                  ((op #::bool => bool list => bool list) h
-                                    t)
-                                  ((op #::bool => bool list => bool list) h'
-                                    t')))))))))
-   (%alg_order_tupled::bool list * bool list => bool.
-       (split::(bool list => bool list => bool)
-               => bool list * bool list => bool)
-        (%(v::bool list) v1::bool list.
-            (list_case::bool
-                        => (bool => bool list => bool) => bool list => bool)
-             ((list_case::bool
-                          => (bool => bool list => bool)
-                             => bool list => bool)
-               (True::bool) (%(v8::bool) v9::bool list. True::bool) v1)
-             (%(v4::bool) v5::bool list.
-                 (list_case::bool
-                             => (bool => bool list => bool)
-                                => bool list => bool)
-                  (False::bool)
-                  (%(v10::bool) v11::bool list.
-                      (op |::bool => bool => bool)
-                       ((op &::bool => bool => bool)
-                         ((op =::bool => bool => bool) v4 (True::bool))
-                         ((op =::bool => bool => bool) v10 (False::bool)))
-                       ((op &::bool => bool => bool)
-                         ((op =::bool => bool => bool) v4 v10)
-                         (alg_order_tupled
-                           ((Pair::bool list
-                                   => bool list => bool list * bool list)
-                             v5 v11))))
-                  v1)
-             v)))"
+lemma alg_order_tupled_primitive_def: "alg_order_tupled =
+WFREC (SOME R. WF R & (ALL h' h t' t. R (t, t') (h # t, h' # t')))
+ (%alg_order_tupled (v, v1).
+     case v of [] => case v1 of [] => True | _ => True
+     | v4 # v5 =>
+         case v1 of [] => False
+         | v10 # v11 =>
+             v4 = True & v10 = False |
+             v4 = v10 & alg_order_tupled (v5, v11))"
   by (import prob_canon alg_order_tupled_primitive_def)
 
 consts
   alg_order :: "bool list => bool list => bool" 
 
 defs
-  alg_order_primdef: "alg_order == %(x::bool list) x1::bool list. alg_order_tupled (x, x1)"
+  alg_order_primdef: "alg_order == %x x1. alg_order_tupled (x, x1)"
 
-lemma alg_order_curried_def: "ALL (x::bool list) x1::bool list. alg_order x x1 = alg_order_tupled (x, x1)"
+lemma alg_order_curried_def: "alg_order x x1 = alg_order_tupled (x, x1)"
   by (import prob_canon alg_order_curried_def)
 
-lemma alg_order_ind: "ALL P::bool list => bool list => bool.
-   (ALL (x::bool) xa::bool list. P [] (x # xa)) &
-   P [] [] &
-   (ALL (x::bool) xa::bool list. P (x # xa) []) &
-   (ALL (x::bool) (xa::bool list) (xb::bool) xc::bool list.
-       P xa xc --> P (x # xa) (xb # xc)) -->
-   (ALL x::bool list. All (P x))"
+lemma alg_order_ind: "(ALL (x::bool) xa::bool list.
+    (P::bool list => bool list => bool) [] (x # xa)) &
+P [] [] &
+(ALL (x::bool) xa::bool list. P (x # xa) []) &
+(ALL (x::bool) (xa::bool list) (xb::bool) xc::bool list.
+    P xa xc --> P (x # xa) (xb # xc))
+==> P (x::bool list) (xa::bool list)"
   by (import prob_canon alg_order_ind)
 
-lemma alg_order_def: "alg_order [] ((v6::bool) # (v7::bool list)) = True &
+lemma alg_order_def: "alg_order [] (v6 # v7) = True &
 alg_order [] [] = True &
-alg_order ((v2::bool) # (v3::bool list)) [] = False &
-alg_order ((h::bool) # (t::bool list)) ((h'::bool) # (t'::bool list)) =
+alg_order (v2 # v3) [] = False &
+alg_order (h # t) (h' # t') =
 (h = True & h' = False | h = h' & alg_order t t')"
   by (import prob_canon alg_order_def)
 
@@ -525,42 +277,28 @@
 
 defs
   alg_sorted_primdef: "alg_sorted ==
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_sorted::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_sorted.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               alg_order v2 v6 & alg_sorted (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. alg_order v2 v6 & alg_sorted (v6 # v7))))"
 
 lemma alg_sorted_primitive_def: "alg_sorted =
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_sorted::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_sorted.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               alg_order v2 v6 & alg_sorted (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. alg_order v2 v6 & alg_sorted (v6 # v7))))"
   by (import prob_canon alg_sorted_primitive_def)
 
-lemma alg_sorted_ind: "ALL P::bool list list => bool.
-   (ALL (x::bool list) (y::bool list) z::bool list list.
-       P (y # z) --> P (x # y # z)) &
-   (ALL v::bool list. P [v]) & P [] -->
-   All P"
+lemma alg_sorted_ind: "(ALL (x::bool list) (y::bool list) z::bool list list.
+    (P::bool list list => bool) (y # z) --> P (x # y # z)) &
+(ALL v::bool list. P [v]) & P []
+==> P (x::bool list list)"
   by (import prob_canon alg_sorted_ind)
 
-lemma alg_sorted_def: "alg_sorted ((x::bool list) # (y::bool list) # (z::bool list list)) =
-(alg_order x y & alg_sorted (y # z)) &
-alg_sorted [v::bool list] = True & alg_sorted [] = True"
+lemma alg_sorted_def: "alg_sorted (x # y # z) = (alg_order x y & alg_sorted (y # z)) &
+alg_sorted [v] = True & alg_sorted [] = True"
   by (import prob_canon alg_sorted_def)
 
 consts
@@ -568,42 +306,28 @@
 
 defs
   alg_prefixfree_primdef: "alg_prefixfree ==
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_prefixfree::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_prefixfree.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))"
 
 lemma alg_prefixfree_primitive_def: "alg_prefixfree =
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_prefixfree::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_prefixfree.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))"
   by (import prob_canon alg_prefixfree_primitive_def)
 
-lemma alg_prefixfree_ind: "ALL P::bool list list => bool.
-   (ALL (x::bool list) (y::bool list) z::bool list list.
-       P (y # z) --> P (x # y # z)) &
-   (ALL v::bool list. P [v]) & P [] -->
-   All P"
+lemma alg_prefixfree_ind: "(ALL (x::bool list) (y::bool list) z::bool list list.
+    (P::bool list list => bool) (y # z) --> P (x # y # z)) &
+(ALL v::bool list. P [v]) & P []
+==> P (x::bool list list)"
   by (import prob_canon alg_prefixfree_ind)
 
-lemma alg_prefixfree_def: "alg_prefixfree ((x::bool list) # (y::bool list) # (z::bool list list)) =
-(~ IS_PREFIX y x & alg_prefixfree (y # z)) &
-alg_prefixfree [v::bool list] = True & alg_prefixfree [] = True"
+lemma alg_prefixfree_def: "alg_prefixfree (x # y # z) = (~ IS_PREFIX y x & alg_prefixfree (y # z)) &
+alg_prefixfree [v] = True & alg_prefixfree [] = True"
   by (import prob_canon alg_prefixfree_def)
 
 consts
@@ -611,60 +335,44 @@
 
 defs
   alg_twinfree_primdef: "alg_twinfree ==
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_twinfree::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_twinfree.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))"
 
 lemma alg_twinfree_primitive_def: "alg_twinfree =
-WFREC
- (SOME R::bool list list => bool list list => bool.
-     WF R &
-     (ALL (x::bool list) (z::bool list list) y::bool list.
-         R (y # z) (x # y # z)))
- (%alg_twinfree::bool list list => bool.
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_twinfree.
      list_case True
-      (%v2::bool list.
-          list_case True
-           (%(v6::bool list) v7::bool list list.
-               ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))"
+      (%v2. list_case True
+             (%v6 v7. ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))"
   by (import prob_canon alg_twinfree_primitive_def)
 
-lemma alg_twinfree_ind: "ALL P::bool list list => bool.
-   (ALL (x::bool list) (y::bool list) z::bool list list.
-       P (y # z) --> P (x # y # z)) &
-   (ALL v::bool list. P [v]) & P [] -->
-   All P"
+lemma alg_twinfree_ind: "(ALL (x::bool list) (y::bool list) z::bool list list.
+    (P::bool list list => bool) (y # z) --> P (x # y # z)) &
+(ALL v::bool list. P [v]) & P []
+==> P (x::bool list list)"
   by (import prob_canon alg_twinfree_ind)
 
-lemma alg_twinfree_def: "alg_twinfree ((x::bool list) # (y::bool list) # (z::bool list list)) =
-(~ alg_twin x y & alg_twinfree (y # z)) &
-alg_twinfree [v::bool list] = True & alg_twinfree [] = True"
+lemma alg_twinfree_def: "alg_twinfree (x # y # z) = (~ alg_twin x y & alg_twinfree (y # z)) &
+alg_twinfree [v] = True & alg_twinfree [] = True"
   by (import prob_canon alg_twinfree_def)
 
 consts
   alg_longest :: "bool list list => nat" 
 
 defs
-  alg_longest_primdef: "alg_longest ==
-FOLDR (%(h::bool list) t::nat. if t <= length h then length h else t) 0"
+  alg_longest_primdef: "alg_longest == FOLDR (%h t. if t <= length h then length h else t) 0"
 
-lemma alg_longest_def: "alg_longest =
-FOLDR (%(h::bool list) t::nat. if t <= length h then length h else t) 0"
+lemma alg_longest_def: "alg_longest = FOLDR (%h t. if t <= length h then length h else t) 0"
   by (import prob_canon alg_longest_def)
 
 consts
   alg_canon_prefs :: "bool list => bool list list => bool list list" 
 
-specification (alg_canon_prefs_primdef: alg_canon_prefs) alg_canon_prefs_def: "(ALL l::bool list. alg_canon_prefs l [] = [l]) &
-(ALL (l::bool list) (h::bool list) t::bool list list.
+specification (alg_canon_prefs_primdef: alg_canon_prefs) alg_canon_prefs_def: "(ALL l. alg_canon_prefs l [] = [l]) &
+(ALL l h t.
     alg_canon_prefs l (h # t) =
     (if IS_PREFIX h l then alg_canon_prefs l t else l # h # t))"
   by (import prob_canon alg_canon_prefs_def)
@@ -672,8 +380,8 @@
 consts
   alg_canon_find :: "bool list => bool list list => bool list list" 
 
-specification (alg_canon_find_primdef: alg_canon_find) alg_canon_find_def: "(ALL l::bool list. alg_canon_find l [] = [l]) &
-(ALL (l::bool list) (h::bool list) t::bool list list.
+specification (alg_canon_find_primdef: alg_canon_find) alg_canon_find_def: "(ALL l. alg_canon_find l [] = [l]) &
+(ALL l h t.
     alg_canon_find l (h # t) =
     (if alg_order h l
      then if IS_PREFIX l h then h # t else h # alg_canon_find l t
@@ -692,8 +400,8 @@
 consts
   alg_canon_merge :: "bool list => bool list list => bool list list" 
 
-specification (alg_canon_merge_primdef: alg_canon_merge) alg_canon_merge_def: "(ALL l::bool list. alg_canon_merge l [] = [l]) &
-(ALL (l::bool list) (h::bool list) t::bool list list.
+specification (alg_canon_merge_primdef: alg_canon_merge) alg_canon_merge_def: "(ALL l. alg_canon_merge l [] = [l]) &
+(ALL l h t.
     alg_canon_merge l (h # t) =
     (if alg_twin l h then alg_canon_merge (butlast h) t else l # h # t))"
   by (import prob_canon alg_canon_merge_def)
@@ -711,365 +419,308 @@
   alg_canon :: "bool list list => bool list list" 
 
 defs
-  alg_canon_primdef: "alg_canon == %l::bool list list. alg_canon2 (alg_canon1 l)"
+  alg_canon_primdef: "alg_canon == %l. alg_canon2 (alg_canon1 l)"
 
-lemma alg_canon_def: "ALL l::bool list list. alg_canon l = alg_canon2 (alg_canon1 l)"
+lemma alg_canon_def: "alg_canon l = alg_canon2 (alg_canon1 l)"
   by (import prob_canon alg_canon_def)
 
 consts
   algebra_canon :: "bool list list => bool" 
 
 defs
-  algebra_canon_primdef: "algebra_canon == %l::bool list list. alg_canon l = l"
+  algebra_canon_primdef: "algebra_canon == %l. alg_canon l = l"
 
-lemma algebra_canon_def: "ALL l::bool list list. algebra_canon l = (alg_canon l = l)"
+lemma algebra_canon_def: "algebra_canon l = (alg_canon l = l)"
   by (import prob_canon algebra_canon_def)
 
-lemma ALG_TWIN_NIL: "ALL l::bool list. ~ alg_twin l [] & ~ alg_twin [] l"
+lemma ALG_TWIN_NIL: "~ alg_twin l [] & ~ alg_twin [] l"
   by (import prob_canon ALG_TWIN_NIL)
 
-lemma ALG_TWIN_SING: "ALL (x::bool) l::bool list.
-   alg_twin [x] l = (x = True & l = [False]) &
-   alg_twin l [x] = (l = [True] & x = False)"
+lemma ALG_TWIN_SING: "alg_twin [x] l = (x = True & l = [False]) &
+alg_twin l [x] = (l = [True] & x = False)"
   by (import prob_canon ALG_TWIN_SING)
 
-lemma ALG_TWIN_CONS: "ALL (x::bool) (y::bool) (z::bool list) (h::bool) t::bool list.
-   alg_twin (x # y # z) (h # t) = (x = h & alg_twin (y # z) t) &
-   alg_twin (h # t) (x # y # z) = (x = h & alg_twin t (y # z))"
+lemma ALG_TWIN_CONS: "alg_twin (x # y # z) (h # t) = (x = h & alg_twin (y # z) t) &
+alg_twin (h # t) (x # y # z) = (x = h & alg_twin t (y # z))"
   by (import prob_canon ALG_TWIN_CONS)
 
-lemma ALG_TWIN_REDUCE: "ALL (h::bool) (t::bool list) t'::bool list.
-   alg_twin (h # t) (h # t') = alg_twin t t'"
+lemma ALG_TWIN_REDUCE: "alg_twin (h # t) (h # t') = alg_twin t t'"
   by (import prob_canon ALG_TWIN_REDUCE)
 
-lemma ALG_TWINS_PREFIX: "ALL (x::bool list) l::bool list.
-   IS_PREFIX x l -->
-   x = l | IS_PREFIX x (SNOC True l) | IS_PREFIX x (SNOC False l)"
+lemma ALG_TWINS_PREFIX: "IS_PREFIX x l
+==> x = l | IS_PREFIX x (SNOC True l) | IS_PREFIX x (SNOC False l)"
   by (import prob_canon ALG_TWINS_PREFIX)
 
-lemma ALG_ORDER_NIL: "ALL x::bool list. alg_order [] x & alg_order x [] = (x = [])"
+lemma ALG_ORDER_NIL: "alg_order [] x & alg_order x [] = (x = [])"
   by (import prob_canon ALG_ORDER_NIL)
 
-lemma ALG_ORDER_REFL: "ALL x::bool list. alg_order x x"
+lemma ALG_ORDER_REFL: "alg_order x x"
   by (import prob_canon ALG_ORDER_REFL)
 
-lemma ALG_ORDER_ANTISYM: "ALL (x::bool list) y::bool list. alg_order x y & alg_order y x --> x = y"
+lemma ALG_ORDER_ANTISYM: "alg_order x y & alg_order y x ==> x = y"
   by (import prob_canon ALG_ORDER_ANTISYM)
 
-lemma ALG_ORDER_TRANS: "ALL (x::bool list) (y::bool list) z::bool list.
-   alg_order x y & alg_order y z --> alg_order x z"
+lemma ALG_ORDER_TRANS: "alg_order x y & alg_order y z ==> alg_order x z"
   by (import prob_canon ALG_ORDER_TRANS)
 
-lemma ALG_ORDER_TOTAL: "ALL (x::bool list) y::bool list. alg_order x y | alg_order y x"
+lemma ALG_ORDER_TOTAL: "alg_order x y | alg_order y x"
   by (import prob_canon ALG_ORDER_TOTAL)
 
-lemma ALG_ORDER_PREFIX: "ALL (x::bool list) y::bool list. IS_PREFIX y x --> alg_order x y"
+lemma ALG_ORDER_PREFIX: "IS_PREFIX y x ==> alg_order x y"
   by (import prob_canon ALG_ORDER_PREFIX)
 
-lemma ALG_ORDER_PREFIX_ANTI: "ALL (x::bool list) y::bool list. alg_order x y & IS_PREFIX x y --> x = y"
+lemma ALG_ORDER_PREFIX_ANTI: "alg_order x y & IS_PREFIX x y ==> x = y"
   by (import prob_canon ALG_ORDER_PREFIX_ANTI)
 
-lemma ALG_ORDER_PREFIX_MONO: "ALL (x::bool list) (y::bool list) z::bool list.
-   alg_order x y & alg_order y z & IS_PREFIX z x --> IS_PREFIX y x"
+lemma ALG_ORDER_PREFIX_MONO: "alg_order x y & alg_order y z & IS_PREFIX z x ==> IS_PREFIX y x"
   by (import prob_canon ALG_ORDER_PREFIX_MONO)
 
-lemma ALG_ORDER_PREFIX_TRANS: "ALL (x::bool list) (y::bool list) z::bool list.
-   alg_order x y & IS_PREFIX y z --> alg_order x z | IS_PREFIX x z"
+lemma ALG_ORDER_PREFIX_TRANS: "alg_order x y & IS_PREFIX y z ==> alg_order x z | IS_PREFIX x z"
   by (import prob_canon ALG_ORDER_PREFIX_TRANS)
 
-lemma ALG_ORDER_SNOC: "ALL (x::bool) l::bool list. ~ alg_order (SNOC x l) l"
+lemma ALG_ORDER_SNOC: "~ alg_order (SNOC x l) l"
   by (import prob_canon ALG_ORDER_SNOC)
 
-lemma ALG_SORTED_MIN: "ALL (h::bool list) t::bool list list.
-   alg_sorted (h # t) --> (ALL x::bool list. x mem t --> alg_order h x)"
+lemma ALG_SORTED_MIN: "[| alg_sorted (h # t); List.member t x |] ==> alg_order h x"
   by (import prob_canon ALG_SORTED_MIN)
 
-lemma ALG_SORTED_DEF_ALT: "ALL (h::bool list) t::bool list list.
-   alg_sorted (h # t) =
-   ((ALL x::bool list. x mem t --> alg_order h x) & alg_sorted t)"
+lemma ALG_SORTED_DEF_ALT: "alg_sorted (h # t) =
+((ALL x. List.member t x --> alg_order h x) & alg_sorted t)"
   by (import prob_canon ALG_SORTED_DEF_ALT)
 
-lemma ALG_SORTED_TL: "ALL (h::bool list) t::bool list list. alg_sorted (h # t) --> alg_sorted t"
+lemma ALG_SORTED_TL: "alg_sorted (h # t) ==> alg_sorted t"
   by (import prob_canon ALG_SORTED_TL)
 
-lemma ALG_SORTED_MONO: "ALL (x::bool list) (y::bool list) z::bool list list.
-   alg_sorted (x # y # z) --> alg_sorted (x # z)"
+lemma ALG_SORTED_MONO: "alg_sorted (x # y # z) ==> alg_sorted (x # z)"
   by (import prob_canon ALG_SORTED_MONO)
 
-lemma ALG_SORTED_TLS: "ALL (l::bool list list) b::bool. alg_sorted (map (op # b) l) = alg_sorted l"
+lemma ALG_SORTED_TLS: "alg_sorted (map (op # b) l) = alg_sorted l"
   by (import prob_canon ALG_SORTED_TLS)
 
-lemma ALG_SORTED_STEP: "ALL (l1::bool list list) l2::bool list list.
-   alg_sorted (map (op # True) l1 @ map (op # False) l2) =
-   (alg_sorted l1 & alg_sorted l2)"
+lemma ALG_SORTED_STEP: "alg_sorted (map (op # True) l1 @ map (op # False) l2) =
+(alg_sorted l1 & alg_sorted l2)"
   by (import prob_canon ALG_SORTED_STEP)
 
-lemma ALG_SORTED_APPEND: "ALL (h::bool list) (h'::bool list) (t::bool list list) t'::bool list list.
-   alg_sorted ((h # t) @ h' # t') =
-   (alg_sorted (h # t) & alg_sorted (h' # t') & alg_order (last (h # t)) h')"
+lemma ALG_SORTED_APPEND: "alg_sorted ((h # t) @ h' # t') =
+(alg_sorted (h # t) & alg_sorted (h' # t') & alg_order (last (h # t)) h')"
   by (import prob_canon ALG_SORTED_APPEND)
 
-lemma ALG_SORTED_FILTER: "ALL (P::bool list => bool) b::bool list list.
-   alg_sorted b --> alg_sorted (filter P b)"
+lemma ALG_SORTED_FILTER: "alg_sorted b ==> alg_sorted (filter P b)"
   by (import prob_canon ALG_SORTED_FILTER)
 
-lemma ALG_PREFIXFREE_TL: "ALL (h::bool list) t::bool list list.
-   alg_prefixfree (h # t) --> alg_prefixfree t"
+lemma ALG_PREFIXFREE_TL: "alg_prefixfree (h # t) ==> alg_prefixfree t"
   by (import prob_canon ALG_PREFIXFREE_TL)
 
-lemma ALG_PREFIXFREE_MONO: "ALL (x::bool list) (y::bool list) z::bool list list.
-   alg_sorted (x # y # z) & alg_prefixfree (x # y # z) -->
-   alg_prefixfree (x # z)"
+lemma ALG_PREFIXFREE_MONO: "alg_sorted (x # y # z) & alg_prefixfree (x # y # z)
+==> alg_prefixfree (x # z)"
   by (import prob_canon ALG_PREFIXFREE_MONO)
 
-lemma ALG_PREFIXFREE_ELT: "ALL (h::bool list) t::bool list list.
-   alg_sorted (h # t) & alg_prefixfree (h # t) -->
-   (ALL x::bool list. x mem t --> ~ IS_PREFIX x h & ~ IS_PREFIX h x)"
+lemma ALG_PREFIXFREE_ELT: "[| alg_sorted (h # t) & alg_prefixfree (h # t); List.member t x |]
+==> ~ IS_PREFIX x h & ~ IS_PREFIX h x"
   by (import prob_canon ALG_PREFIXFREE_ELT)
 
-lemma ALG_PREFIXFREE_TLS: "ALL (l::bool list list) b::bool.
-   alg_prefixfree (map (op # b) l) = alg_prefixfree l"
+lemma ALG_PREFIXFREE_TLS: "alg_prefixfree (map (op # b) l) = alg_prefixfree l"
   by (import prob_canon ALG_PREFIXFREE_TLS)
 
-lemma ALG_PREFIXFREE_STEP: "ALL (l1::bool list list) l2::bool list list.
-   alg_prefixfree (map (op # True) l1 @ map (op # False) l2) =
-   (alg_prefixfree l1 & alg_prefixfree l2)"
+lemma ALG_PREFIXFREE_STEP: "alg_prefixfree (map (op # True) l1 @ map (op # False) l2) =
+(alg_prefixfree l1 & alg_prefixfree l2)"
   by (import prob_canon ALG_PREFIXFREE_STEP)
 
-lemma ALG_PREFIXFREE_APPEND: "ALL (h::bool list) (h'::bool list) (t::bool list list) t'::bool list list.
-   alg_prefixfree ((h # t) @ h' # t') =
-   (alg_prefixfree (h # t) &
-    alg_prefixfree (h' # t') & ~ IS_PREFIX h' (last (h # t)))"
+lemma ALG_PREFIXFREE_APPEND: "alg_prefixfree ((h # t) @ h' # t') =
+(alg_prefixfree (h # t) &
+ alg_prefixfree (h' # t') & ~ IS_PREFIX h' (last (h # t)))"
   by (import prob_canon ALG_PREFIXFREE_APPEND)
 
-lemma ALG_PREFIXFREE_FILTER: "ALL (P::bool list => bool) b::bool list list.
-   alg_sorted b & alg_prefixfree b --> alg_prefixfree (filter P b)"
+lemma ALG_PREFIXFREE_FILTER: "alg_sorted b & alg_prefixfree b ==> alg_prefixfree (filter P b)"
   by (import prob_canon ALG_PREFIXFREE_FILTER)
 
-lemma ALG_TWINFREE_TL: "ALL (h::bool list) t::bool list list.
-   alg_twinfree (h # t) --> alg_twinfree t"
+lemma ALG_TWINFREE_TL: "alg_twinfree (h # t) ==> alg_twinfree t"
   by (import prob_canon ALG_TWINFREE_TL)
 
-lemma ALG_TWINFREE_TLS: "ALL (l::bool list list) b::bool.
-   alg_twinfree (map (op # b) l) = alg_twinfree l"
+lemma ALG_TWINFREE_TLS: "alg_twinfree (map (op # b) l) = alg_twinfree l"
   by (import prob_canon ALG_TWINFREE_TLS)
 
-lemma ALG_TWINFREE_STEP1: "ALL (l1::bool list list) l2::bool list list.
-   alg_twinfree (map (op # True) l1 @ map (op # False) l2) -->
-   alg_twinfree l1 & alg_twinfree l2"
+lemma ALG_TWINFREE_STEP1: "alg_twinfree (map (op # True) l1 @ map (op # False) l2)
+==> alg_twinfree l1 & alg_twinfree l2"
   by (import prob_canon ALG_TWINFREE_STEP1)
 
-lemma ALG_TWINFREE_STEP2: "ALL (l1::bool list list) l2::bool list list.
-   (~ [] mem l1 | ~ [] mem l2) & alg_twinfree l1 & alg_twinfree l2 -->
-   alg_twinfree (map (op # True) l1 @ map (op # False) l2)"
+lemma ALG_TWINFREE_STEP2: "(~ List.member l1 [] | ~ List.member l2 []) &
+alg_twinfree l1 & alg_twinfree l2
+==> alg_twinfree (map (op # True) l1 @ map (op # False) l2)"
   by (import prob_canon ALG_TWINFREE_STEP2)
 
-lemma ALG_TWINFREE_STEP: "ALL (l1::bool list list) l2::bool list list.
-   ~ [] mem l1 | ~ [] mem l2 -->
-   alg_twinfree (map (op # True) l1 @ map (op # False) l2) =
-   (alg_twinfree l1 & alg_twinfree l2)"
+lemma ALG_TWINFREE_STEP: "~ List.member l1 [] | ~ List.member l2 []
+==> alg_twinfree (map (op # True) l1 @ map (op # False) l2) =
+    (alg_twinfree l1 & alg_twinfree l2)"
   by (import prob_canon ALG_TWINFREE_STEP)
 
-lemma ALG_LONGEST_HD: "ALL (h::bool list) t::bool list list. length h <= alg_longest (h # t)"
+lemma ALG_LONGEST_HD: "length h <= alg_longest (h # t)"
   by (import prob_canon ALG_LONGEST_HD)
 
-lemma ALG_LONGEST_TL: "ALL (h::bool list) t::bool list list. alg_longest t <= alg_longest (h # t)"
+lemma ALG_LONGEST_TL: "alg_longest t <= alg_longest (h # t)"
   by (import prob_canon ALG_LONGEST_TL)
 
-lemma ALG_LONGEST_TLS: "ALL (h::bool list) (t::bool list list) b::bool.
-   alg_longest (map (op # b) (h # t)) = Suc (alg_longest (h # t))"
+lemma ALG_LONGEST_TLS: "alg_longest (map (op # b) (h # t)) = Suc (alg_longest (h # t))"
   by (import prob_canon ALG_LONGEST_TLS)
 
-lemma ALG_LONGEST_APPEND: "ALL (l1::bool list list) l2::bool list list.
-   alg_longest l1 <= alg_longest (l1 @ l2) &
-   alg_longest l2 <= alg_longest (l1 @ l2)"
+lemma ALG_LONGEST_APPEND: "alg_longest l1 <= alg_longest (l1 @ l2) &
+alg_longest l2 <= alg_longest (l1 @ l2)"
   by (import prob_canon ALG_LONGEST_APPEND)
 
-lemma ALG_CANON_PREFS_HD: "ALL (l::bool list) b::bool list list. hd (alg_canon_prefs l b) = l"
+lemma ALG_CANON_PREFS_HD: "hd (alg_canon_prefs l b) = l"
   by (import prob_canon ALG_CANON_PREFS_HD)
 
-lemma ALG_CANON_PREFS_DELETES: "ALL (l::bool list) (b::bool list list) x::bool list.
-   x mem alg_canon_prefs l b --> x mem l # b"
+lemma ALG_CANON_PREFS_DELETES: "List.member (alg_canon_prefs l b) x ==> List.member (l # b) x"
   by (import prob_canon ALG_CANON_PREFS_DELETES)
 
-lemma ALG_CANON_PREFS_SORTED: "ALL (l::bool list) b::bool list list.
-   alg_sorted (l # b) --> alg_sorted (alg_canon_prefs l b)"
+lemma ALG_CANON_PREFS_SORTED: "alg_sorted (l # b) ==> alg_sorted (alg_canon_prefs l b)"
   by (import prob_canon ALG_CANON_PREFS_SORTED)
 
-lemma ALG_CANON_PREFS_PREFIXFREE: "ALL (l::bool list) b::bool list list.
-   alg_sorted b & alg_prefixfree b --> alg_prefixfree (alg_canon_prefs l b)"
+lemma ALG_CANON_PREFS_PREFIXFREE: "alg_sorted b & alg_prefixfree b ==> alg_prefixfree (alg_canon_prefs l b)"
   by (import prob_canon ALG_CANON_PREFS_PREFIXFREE)
 
-lemma ALG_CANON_PREFS_CONSTANT: "ALL (l::bool list) b::bool list list.
-   alg_prefixfree (l # b) --> alg_canon_prefs l b = l # b"
+lemma ALG_CANON_PREFS_CONSTANT: "alg_prefixfree (l # b) ==> alg_canon_prefs l b = l # b"
   by (import prob_canon ALG_CANON_PREFS_CONSTANT)
 
-lemma ALG_CANON_FIND_HD: "ALL (l::bool list) (h::bool list) t::bool list list.
-   hd (alg_canon_find l (h # t)) = l | hd (alg_canon_find l (h # t)) = h"
+lemma ALG_CANON_FIND_HD: "hd (alg_canon_find l (h # t)) = l | hd (alg_canon_find l (h # t)) = h"
   by (import prob_canon ALG_CANON_FIND_HD)
 
-lemma ALG_CANON_FIND_DELETES: "ALL (l::bool list) (b::bool list list) x::bool list.
-   x mem alg_canon_find l b --> x mem l # b"
+lemma ALG_CANON_FIND_DELETES: "List.member (alg_canon_find l b) x ==> List.member (l # b) x"
   by (import prob_canon ALG_CANON_FIND_DELETES)
 
-lemma ALG_CANON_FIND_SORTED: "ALL (l::bool list) b::bool list list.
-   alg_sorted b --> alg_sorted (alg_canon_find l b)"
+lemma ALG_CANON_FIND_SORTED: "alg_sorted b ==> alg_sorted (alg_canon_find l b)"
   by (import prob_canon ALG_CANON_FIND_SORTED)
 
-lemma ALG_CANON_FIND_PREFIXFREE: "ALL (l::bool list) b::bool list list.
-   alg_sorted b & alg_prefixfree b --> alg_prefixfree (alg_canon_find l b)"
+lemma ALG_CANON_FIND_PREFIXFREE: "alg_sorted b & alg_prefixfree b ==> alg_prefixfree (alg_canon_find l b)"
   by (import prob_canon ALG_CANON_FIND_PREFIXFREE)
 
-lemma ALG_CANON_FIND_CONSTANT: "ALL (l::bool list) b::bool list list.
-   alg_sorted (l # b) & alg_prefixfree (l # b) -->
-   alg_canon_find l b = l # b"
+lemma ALG_CANON_FIND_CONSTANT: "alg_sorted (l # b) & alg_prefixfree (l # b) ==> alg_canon_find l b = l # b"
   by (import prob_canon ALG_CANON_FIND_CONSTANT)
 
-lemma ALG_CANON1_SORTED: "ALL x::bool list list. alg_sorted (alg_canon1 x)"
+lemma ALG_CANON1_SORTED: "alg_sorted (alg_canon1 x)"
   by (import prob_canon ALG_CANON1_SORTED)
 
-lemma ALG_CANON1_PREFIXFREE: "ALL l::bool list list. alg_prefixfree (alg_canon1 l)"
+lemma ALG_CANON1_PREFIXFREE: "alg_prefixfree (alg_canon1 l)"
   by (import prob_canon ALG_CANON1_PREFIXFREE)
 
-lemma ALG_CANON1_CONSTANT: "ALL l::bool list list. alg_sorted l & alg_prefixfree l --> alg_canon1 l = l"
+lemma ALG_CANON1_CONSTANT: "alg_sorted l & alg_prefixfree l ==> alg_canon1 l = l"
   by (import prob_canon ALG_CANON1_CONSTANT)
 
-lemma ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE: "ALL (l::bool list) b::bool list list.
-   alg_sorted (l # b) & alg_prefixfree (l # b) & alg_twinfree b -->
-   alg_sorted (alg_canon_merge l b) &
-   alg_prefixfree (alg_canon_merge l b) & alg_twinfree (alg_canon_merge l b)"
+lemma ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE: "alg_sorted (l # b) & alg_prefixfree (l # b) & alg_twinfree b
+==> alg_sorted (alg_canon_merge l b) &
+    alg_prefixfree (alg_canon_merge l b) &
+    alg_twinfree (alg_canon_merge l b)"
   by (import prob_canon ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE)
 
-lemma ALG_CANON_MERGE_PREFIXFREE_PRESERVE: "ALL (l::bool list) (b::bool list list) h::bool list.
-   (ALL x::bool list. x mem l # b --> ~ IS_PREFIX h x & ~ IS_PREFIX x h) -->
-   (ALL x::bool list.
-       x mem alg_canon_merge l b --> ~ IS_PREFIX h x & ~ IS_PREFIX x h)"
+lemma ALG_CANON_MERGE_PREFIXFREE_PRESERVE: "[| !!x. List.member (l # b) x ==> ~ IS_PREFIX h x & ~ IS_PREFIX x h;
+   List.member (alg_canon_merge l b) x |]
+==> ~ IS_PREFIX h x & ~ IS_PREFIX x h"
   by (import prob_canon ALG_CANON_MERGE_PREFIXFREE_PRESERVE)
 
-lemma ALG_CANON_MERGE_SHORTENS: "ALL (l::bool list) (b::bool list list) x::bool list.
-   x mem alg_canon_merge l b -->
-   (EX y::bool list. y mem l # b & IS_PREFIX y x)"
+lemma ALG_CANON_MERGE_SHORTENS: "List.member (alg_canon_merge l b) x
+==> EX y. List.member (l # b) y & IS_PREFIX y x"
   by (import prob_canon ALG_CANON_MERGE_SHORTENS)
 
-lemma ALG_CANON_MERGE_CONSTANT: "ALL (l::bool list) b::bool list list.
-   alg_twinfree (l # b) --> alg_canon_merge l b = l # b"
+lemma ALG_CANON_MERGE_CONSTANT: "alg_twinfree (l # b) ==> alg_canon_merge l b = l # b"
   by (import prob_canon ALG_CANON_MERGE_CONSTANT)
 
-lemma ALG_CANON2_PREFIXFREE_PRESERVE: "ALL (x::bool list list) xa::bool list.
-   (ALL xb::bool list.
-       xb mem x --> ~ IS_PREFIX xa xb & ~ IS_PREFIX xb xa) -->
-   (ALL xb::bool list.
-       xb mem alg_canon2 x --> ~ IS_PREFIX xa xb & ~ IS_PREFIX xb xa)"
+lemma ALG_CANON2_PREFIXFREE_PRESERVE: "[| !!xb. List.member x xb ==> ~ IS_PREFIX xa xb & ~ IS_PREFIX xb xa;
+   List.member (alg_canon2 x) xb |]
+==> ~ IS_PREFIX xa xb & ~ IS_PREFIX xb xa"
   by (import prob_canon ALG_CANON2_PREFIXFREE_PRESERVE)
 
-lemma ALG_CANON2_SHORTENS: "ALL (x::bool list list) xa::bool list.
-   xa mem alg_canon2 x --> (EX y::bool list. y mem x & IS_PREFIX y xa)"
+lemma ALG_CANON2_SHORTENS: "List.member (alg_canon2 x) xa ==> EX y. List.member x y & IS_PREFIX y xa"
   by (import prob_canon ALG_CANON2_SHORTENS)
 
-lemma ALG_CANON2_SORTED_PREFIXFREE_TWINFREE: "ALL x::bool list list.
-   alg_sorted x & alg_prefixfree x -->
-   alg_sorted (alg_canon2 x) &
-   alg_prefixfree (alg_canon2 x) & alg_twinfree (alg_canon2 x)"
+lemma ALG_CANON2_SORTED_PREFIXFREE_TWINFREE: "alg_sorted x & alg_prefixfree x
+==> alg_sorted (alg_canon2 x) &
+    alg_prefixfree (alg_canon2 x) & alg_twinfree (alg_canon2 x)"
   by (import prob_canon ALG_CANON2_SORTED_PREFIXFREE_TWINFREE)
 
-lemma ALG_CANON2_CONSTANT: "ALL l::bool list list. alg_twinfree l --> alg_canon2 l = l"
+lemma ALG_CANON2_CONSTANT: "alg_twinfree l ==> alg_canon2 l = l"
   by (import prob_canon ALG_CANON2_CONSTANT)
 
-lemma ALG_CANON_SORTED_PREFIXFREE_TWINFREE: "ALL l::bool list list.
-   alg_sorted (alg_canon l) &
-   alg_prefixfree (alg_canon l) & alg_twinfree (alg_canon l)"
+lemma ALG_CANON_SORTED_PREFIXFREE_TWINFREE: "alg_sorted (alg_canon l) &
+alg_prefixfree (alg_canon l) & alg_twinfree (alg_canon l)"
   by (import prob_canon ALG_CANON_SORTED_PREFIXFREE_TWINFREE)
 
-lemma ALG_CANON_CONSTANT: "ALL l::bool list list.
-   alg_sorted l & alg_prefixfree l & alg_twinfree l --> alg_canon l = l"
+lemma ALG_CANON_CONSTANT: "alg_sorted l & alg_prefixfree l & alg_twinfree l ==> alg_canon l = l"
   by (import prob_canon ALG_CANON_CONSTANT)
 
-lemma ALG_CANON_IDEMPOT: "ALL l::bool list list. alg_canon (alg_canon l) = alg_canon l"
+lemma ALG_CANON_IDEMPOT: "alg_canon (alg_canon l) = alg_canon l"
   by (import prob_canon ALG_CANON_IDEMPOT)
 
-lemma ALGEBRA_CANON_DEF_ALT: "ALL l::bool list list.
-   algebra_canon l = (alg_sorted l & alg_prefixfree l & alg_twinfree l)"
+lemma ALGEBRA_CANON_DEF_ALT: "algebra_canon l = (alg_sorted l & alg_prefixfree l & alg_twinfree l)"
   by (import prob_canon ALGEBRA_CANON_DEF_ALT)
 
-lemma ALGEBRA_CANON_BASIC: "algebra_canon [] &
-algebra_canon [[]] & (ALL x::bool list. algebra_canon [x])"
+lemma ALGEBRA_CANON_BASIC: "algebra_canon [] & algebra_canon [[]] & (ALL x. algebra_canon [x])"
   by (import prob_canon ALGEBRA_CANON_BASIC)
 
-lemma ALG_CANON_BASIC: "alg_canon [] = [] &
-alg_canon [[]] = [[]] & (ALL x::bool list. alg_canon [x] = [x])"
+lemma ALG_CANON_BASIC: "alg_canon [] = [] & alg_canon [[]] = [[]] & (ALL x. alg_canon [x] = [x])"
   by (import prob_canon ALG_CANON_BASIC)
 
-lemma ALGEBRA_CANON_TL: "ALL (h::bool list) t::bool list list.
-   algebra_canon (h # t) --> algebra_canon t"
+lemma ALGEBRA_CANON_TL: "algebra_canon (h # t) ==> algebra_canon t"
   by (import prob_canon ALGEBRA_CANON_TL)
 
-lemma ALGEBRA_CANON_NIL_MEM: "ALL l::bool list list. (algebra_canon l & [] mem l) = (l = [[]])"
+lemma ALGEBRA_CANON_NIL_MEM: "(algebra_canon l & List.member l []) = (l = [[]])"
   by (import prob_canon ALGEBRA_CANON_NIL_MEM)
 
-lemma ALGEBRA_CANON_TLS: "ALL (l::bool list list) b::bool.
-   algebra_canon (map (op # b) l) = algebra_canon l"
+lemma ALGEBRA_CANON_TLS: "algebra_canon (map (op # b) l) = algebra_canon l"
   by (import prob_canon ALGEBRA_CANON_TLS)
 
-lemma ALGEBRA_CANON_STEP1: "ALL (l1::bool list list) l2::bool list list.
-   algebra_canon (map (op # True) l1 @ map (op # False) l2) -->
-   algebra_canon l1 & algebra_canon l2"
+lemma ALGEBRA_CANON_STEP1: "algebra_canon (map (op # True) l1 @ map (op # False) l2)
+==> algebra_canon l1 & algebra_canon l2"
   by (import prob_canon ALGEBRA_CANON_STEP1)
 
-lemma ALGEBRA_CANON_STEP2: "ALL (l1::bool list list) l2::bool list list.
-   (l1 ~= [[]] | l2 ~= [[]]) & algebra_canon l1 & algebra_canon l2 -->
-   algebra_canon (map (op # True) l1 @ map (op # False) l2)"
+lemma ALGEBRA_CANON_STEP2: "(l1 ~= [[]] | l2 ~= [[]]) & algebra_canon l1 & algebra_canon l2
+==> algebra_canon (map (op # True) l1 @ map (op # False) l2)"
   by (import prob_canon ALGEBRA_CANON_STEP2)
 
-lemma ALGEBRA_CANON_STEP: "ALL (l1::bool list list) l2::bool list list.
-   l1 ~= [[]] | l2 ~= [[]] -->
-   algebra_canon (map (op # True) l1 @ map (op # False) l2) =
-   (algebra_canon l1 & algebra_canon l2)"
+lemma ALGEBRA_CANON_STEP: "l1 ~= [[]] | l2 ~= [[]]
+==> algebra_canon (map (op # True) l1 @ map (op # False) l2) =
+    (algebra_canon l1 & algebra_canon l2)"
   by (import prob_canon ALGEBRA_CANON_STEP)
 
-lemma ALGEBRA_CANON_CASES_THM: "ALL l::bool list list.
-   algebra_canon l -->
-   l = [] |
-   l = [[]] |
-   (EX (l1::bool list list) l2::bool list list.
-       algebra_canon l1 &
-       algebra_canon l2 & l = map (op # True) l1 @ map (op # False) l2)"
+lemma ALGEBRA_CANON_CASES_THM: "algebra_canon l
+==> l = [] |
+    l = [[]] |
+    (EX l1 l2.
+        algebra_canon l1 &
+        algebra_canon l2 & l = map (op # True) l1 @ map (op # False) l2)"
   by (import prob_canon ALGEBRA_CANON_CASES_THM)
 
-lemma ALGEBRA_CANON_CASES: "ALL P::bool list list => bool.
-   P [] &
+lemma ALGEBRA_CANON_CASES: "[| P [] &
    P [[]] &
-   (ALL (l1::bool list list) l2::bool list list.
+   (ALL l1 l2.
        algebra_canon l1 &
        algebra_canon l2 &
        algebra_canon (map (op # True) l1 @ map (op # False) l2) -->
-       P (map (op # True) l1 @ map (op # False) l2)) -->
-   (ALL l::bool list list. algebra_canon l --> P l)"
+       P (map (op # True) l1 @ map (op # False) l2));
+   algebra_canon l |]
+==> P l"
   by (import prob_canon ALGEBRA_CANON_CASES)
 
-lemma ALGEBRA_CANON_INDUCTION: "ALL P::bool list list => bool.
-   P [] &
+lemma ALGEBRA_CANON_INDUCTION: "[| P [] &
    P [[]] &
-   (ALL (l1::bool list list) l2::bool list list.
+   (ALL l1 l2.
        algebra_canon l1 &
        algebra_canon l2 &
        P l1 &
        P l2 & algebra_canon (map (op # True) l1 @ map (op # False) l2) -->
-       P (map (op # True) l1 @ map (op # False) l2)) -->
-   (ALL l::bool list list. algebra_canon l --> P l)"
+       P (map (op # True) l1 @ map (op # False) l2));
+   algebra_canon l |]
+==> P l"
   by (import prob_canon ALGEBRA_CANON_INDUCTION)
 
-lemma MEM_NIL_STEP: "ALL (l1::bool list list) l2::bool list list.
-   ~ [] mem map (op # True) l1 @ map (op # False) l2"
+lemma MEM_NIL_STEP: "~ List.member (map (op # True) l1 @ map (op # False) l2) []"
   by (import prob_canon MEM_NIL_STEP)
 
-lemma ALG_SORTED_PREFIXFREE_MEM_NIL: "ALL l::bool list list.
-   (alg_sorted l & alg_prefixfree l & [] mem l) = (l = [[]])"
+lemma ALG_SORTED_PREFIXFREE_MEM_NIL: "(alg_sorted l & alg_prefixfree l & List.member l []) = (l = [[]])"
   by (import prob_canon ALG_SORTED_PREFIXFREE_MEM_NIL)
 
-lemma ALG_SORTED_PREFIXFREE_EQUALITY: "ALL (l::bool list list) l'::bool list list.
-   (ALL x::bool list. x mem l = x mem l') &
-   alg_sorted l & alg_sorted l' & alg_prefixfree l & alg_prefixfree l' -->
-   l = l'"
+lemma ALG_SORTED_PREFIXFREE_EQUALITY: "(ALL x. List.member l x = List.member l' x) &
+alg_sorted l & alg_sorted l' & alg_prefixfree l & alg_prefixfree l'
+==> l = l'"
   by (import prob_canon ALG_SORTED_PREFIXFREE_EQUALITY)
 
 ;end_setup
@@ -1080,34 +731,33 @@
   SHD :: "(nat => bool) => bool" 
 
 defs
-  SHD_primdef: "SHD == %f::nat => bool. f 0"
+  SHD_primdef: "SHD == %f. f 0"
 
-lemma SHD_def: "ALL f::nat => bool. SHD f = f 0"
+lemma SHD_def: "SHD f = f 0"
   by (import boolean_sequence SHD_def)
 
 consts
   STL :: "(nat => bool) => nat => bool" 
 
 defs
-  STL_primdef: "STL == %(f::nat => bool) n::nat. f (Suc n)"
+  STL_primdef: "STL == %f n. f (Suc n)"
 
-lemma STL_def: "ALL (f::nat => bool) n::nat. STL f n = f (Suc n)"
+lemma STL_def: "STL f n = f (Suc n)"
   by (import boolean_sequence STL_def)
 
 consts
   SCONS :: "bool => (nat => bool) => nat => bool" 
 
-specification (SCONS_primdef: SCONS) SCONS_def: "(ALL (h::bool) t::nat => bool. SCONS h t 0 = h) &
-(ALL (h::bool) (t::nat => bool) n::nat. SCONS h t (Suc n) = t n)"
+specification (SCONS_primdef: SCONS) SCONS_def: "(ALL h t. SCONS h t 0 = h) & (ALL h t n. SCONS h t (Suc n) = t n)"
   by (import boolean_sequence SCONS_def)
 
 consts
   SDEST :: "(nat => bool) => bool * (nat => bool)" 
 
 defs
-  SDEST_primdef: "SDEST == %s::nat => bool. (SHD s, STL s)"
+  SDEST_primdef: "SDEST == %s. (SHD s, STL s)"
 
-lemma SDEST_def: "SDEST = (%s::nat => bool. (SHD s, STL s))"
+lemma SDEST_def: "SDEST = (%s. (SHD s, STL s))"
   by (import boolean_sequence SDEST_def)
 
 consts
@@ -1122,32 +772,32 @@
 consts
   STAKE :: "nat => (nat => bool) => bool list" 
 
-specification (STAKE_primdef: STAKE) STAKE_def: "(ALL s::nat => bool. STAKE 0 s = []) &
-(ALL (n::nat) s::nat => bool. STAKE (Suc n) s = SHD s # STAKE n (STL s))"
+specification (STAKE_primdef: STAKE) STAKE_def: "(ALL s. STAKE 0 s = []) &
+(ALL n s. STAKE (Suc n) s = SHD s # STAKE n (STL s))"
   by (import boolean_sequence STAKE_def)
 
 consts
   SDROP :: "nat => (nat => bool) => nat => bool" 
 
-specification (SDROP_primdef: SDROP) SDROP_def: "SDROP 0 = I & (ALL n::nat. SDROP (Suc n) = SDROP n o STL)"
+specification (SDROP_primdef: SDROP) SDROP_def: "SDROP 0 = I & (ALL n. SDROP (Suc n) = SDROP n o STL)"
   by (import boolean_sequence SDROP_def)
 
-lemma SCONS_SURJ: "ALL x::nat => bool. EX (xa::bool) t::nat => bool. x = SCONS xa t"
+lemma SCONS_SURJ: "EX xa t. x = SCONS xa t"
   by (import boolean_sequence SCONS_SURJ)
 
-lemma SHD_STL_ISO: "ALL (h::bool) t::nat => bool. EX x::nat => bool. SHD x = h & STL x = t"
+lemma SHD_STL_ISO: "EX x. SHD x = h & STL x = t"
   by (import boolean_sequence SHD_STL_ISO)
 
-lemma SHD_SCONS: "ALL (h::bool) t::nat => bool. SHD (SCONS h t) = h"
+lemma SHD_SCONS: "SHD (SCONS h t) = h"
   by (import boolean_sequence SHD_SCONS)
 
-lemma STL_SCONS: "ALL (h::bool) t::nat => bool. STL (SCONS h t) = t"
+lemma STL_SCONS: "STL (SCONS h t) = t"
   by (import boolean_sequence STL_SCONS)
 
-lemma SHD_SCONST: "ALL b::bool. SHD (SCONST b) = b"
+lemma SHD_SCONST: "SHD (SCONST b) = b"
   by (import boolean_sequence SHD_SCONST)
 
-lemma STL_SCONST: "ALL b::bool. STL (SCONST b) = SCONST b"
+lemma STL_SCONST: "STL (SCONST b) = SCONST b"
   by (import boolean_sequence STL_SCONST)
 
 ;end_setup
@@ -1157,16 +807,15 @@
 consts
   alg_embed :: "bool list => (nat => bool) => bool" 
 
-specification (alg_embed_primdef: alg_embed) alg_embed_def: "(ALL s::nat => bool. alg_embed [] s = True) &
-(ALL (h::bool) (t::bool list) s::nat => bool.
-    alg_embed (h # t) s = (h = SHD s & alg_embed t (STL s)))"
+specification (alg_embed_primdef: alg_embed) alg_embed_def: "(ALL s. alg_embed [] s = True) &
+(ALL h t s. alg_embed (h # t) s = (h = SHD s & alg_embed t (STL s)))"
   by (import prob_algebra alg_embed_def)
 
 consts
   algebra_embed :: "bool list list => (nat => bool) => bool" 
 
 specification (algebra_embed_primdef: algebra_embed) algebra_embed_def: "algebra_embed [] = EMPTY &
-(ALL (h::bool list) t::bool list list.
+(ALL h t.
     algebra_embed (h # t) = pred_set.UNION (alg_embed h) (algebra_embed t))"
   by (import prob_algebra algebra_embed_def)
 
@@ -1174,157 +823,127 @@
   measurable :: "((nat => bool) => bool) => bool" 
 
 defs
-  measurable_primdef: "measurable ==
-%s::(nat => bool) => bool. EX b::bool list list. s = algebra_embed b"
+  measurable_primdef: "measurable == %s. EX b. s = algebra_embed b"
 
-lemma measurable_def: "ALL s::(nat => bool) => bool.
-   measurable s = (EX b::bool list list. s = algebra_embed b)"
+lemma measurable_def: "measurable s = (EX b. s = algebra_embed b)"
   by (import prob_algebra measurable_def)
 
-lemma HALVES_INTER: "pred_set.INTER (%x::nat => bool. SHD x = True)
- (%x::nat => bool. SHD x = False) =
-EMPTY"
+lemma HALVES_INTER: "pred_set.INTER (%x. SHD x = True) (%x. SHD x = False) = EMPTY"
   by (import prob_algebra HALVES_INTER)
 
-lemma INTER_STL: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool.
-   pred_set.INTER p q o STL = pred_set.INTER (p o STL) (q o STL)"
+lemma INTER_STL: "pred_set.INTER p q o STL = pred_set.INTER (p o STL) (q o STL)"
   by (import prob_algebra INTER_STL)
 
-lemma COMPL_SHD: "ALL b::bool.
-   COMPL (%x::nat => bool. SHD x = b) = (%x::nat => bool. SHD x = (~ b))"
+lemma COMPL_SHD: "COMPL (%x. SHD x = b) = (%x. SHD x = (~ b))"
   by (import prob_algebra COMPL_SHD)
 
 lemma ALG_EMBED_BASIC: "alg_embed [] = pred_set.UNIV &
-(ALL (h::bool) t::bool list.
-    alg_embed (h # t) =
-    pred_set.INTER (%x::nat => bool. SHD x = h) (alg_embed t o STL))"
+(ALL h t.
+    alg_embed (h # t) = pred_set.INTER (%x. SHD x = h) (alg_embed t o STL))"
   by (import prob_algebra ALG_EMBED_BASIC)
 
-lemma ALG_EMBED_NIL: "ALL c::bool list. All (alg_embed c) = (c = [])"
+lemma ALG_EMBED_NIL: "All (alg_embed c) = (c = [])"
   by (import prob_algebra ALG_EMBED_NIL)
 
-lemma ALG_EMBED_POPULATED: "ALL b::bool list. Ex (alg_embed b)"
+lemma ALG_EMBED_POPULATED: "Ex (alg_embed b)"
   by (import prob_algebra ALG_EMBED_POPULATED)
 
-lemma ALG_EMBED_PREFIX: "ALL (b::bool list) (c::bool list) s::nat => bool.
-   alg_embed b s & alg_embed c s --> IS_PREFIX b c | IS_PREFIX c b"
+lemma ALG_EMBED_PREFIX: "alg_embed b s & alg_embed c s ==> IS_PREFIX b c | IS_PREFIX c b"
   by (import prob_algebra ALG_EMBED_PREFIX)
 
-lemma ALG_EMBED_PREFIX_SUBSET: "ALL (b::bool list) c::bool list.
-   SUBSET (alg_embed b) (alg_embed c) = IS_PREFIX b c"
+lemma ALG_EMBED_PREFIX_SUBSET: "SUBSET (alg_embed b) (alg_embed c) = IS_PREFIX b c"
   by (import prob_algebra ALG_EMBED_PREFIX_SUBSET)
 
-lemma ALG_EMBED_TWINS: "ALL l::bool list.
-   pred_set.UNION (alg_embed (SNOC True l)) (alg_embed (SNOC False l)) =
-   alg_embed l"
+lemma ALG_EMBED_TWINS: "pred_set.UNION (alg_embed (SNOC True l)) (alg_embed (SNOC False l)) =
+alg_embed l"
   by (import prob_algebra ALG_EMBED_TWINS)
 
 lemma ALGEBRA_EMBED_BASIC: "algebra_embed [] = EMPTY &
 algebra_embed [[]] = pred_set.UNIV &
-(ALL b::bool. algebra_embed [[b]] = (%s::nat => bool. SHD s = b))"
+(ALL b. algebra_embed [[b]] = (%s. SHD s = b))"
   by (import prob_algebra ALGEBRA_EMBED_BASIC)
 
-lemma ALGEBRA_EMBED_MEM: "ALL (b::bool list list) x::nat => bool.
-   algebra_embed b x --> (EX l::bool list. l mem b & alg_embed l x)"
+lemma ALGEBRA_EMBED_MEM: "algebra_embed b x ==> EX l. List.member b l & alg_embed l x"
   by (import prob_algebra ALGEBRA_EMBED_MEM)
 
-lemma ALGEBRA_EMBED_APPEND: "ALL (l1::bool list list) l2::bool list list.
-   algebra_embed (l1 @ l2) =
-   pred_set.UNION (algebra_embed l1) (algebra_embed l2)"
+lemma ALGEBRA_EMBED_APPEND: "algebra_embed (l1 @ l2) =
+pred_set.UNION (algebra_embed l1) (algebra_embed l2)"
   by (import prob_algebra ALGEBRA_EMBED_APPEND)
 
-lemma ALGEBRA_EMBED_TLS: "ALL (l::bool list list) b::bool.
-   algebra_embed (map (op # b) l) (SCONS (h::bool) (t::nat => bool)) =
-   (h = b & algebra_embed l t)"
+lemma ALGEBRA_EMBED_TLS: "algebra_embed (map (op # b) l) (SCONS h t) = (h = b & algebra_embed l t)"
   by (import prob_algebra ALGEBRA_EMBED_TLS)
 
-lemma ALG_CANON_PREFS_EMBED: "ALL (l::bool list) b::bool list list.
-   algebra_embed (alg_canon_prefs l b) = algebra_embed (l # b)"
+lemma ALG_CANON_PREFS_EMBED: "algebra_embed (alg_canon_prefs l b) = algebra_embed (l # b)"
   by (import prob_algebra ALG_CANON_PREFS_EMBED)
 
-lemma ALG_CANON_FIND_EMBED: "ALL (l::bool list) b::bool list list.
-   algebra_embed (alg_canon_find l b) = algebra_embed (l # b)"
+lemma ALG_CANON_FIND_EMBED: "algebra_embed (alg_canon_find l b) = algebra_embed (l # b)"
   by (import prob_algebra ALG_CANON_FIND_EMBED)
 
-lemma ALG_CANON1_EMBED: "ALL x::bool list list. algebra_embed (alg_canon1 x) = algebra_embed x"
+lemma ALG_CANON1_EMBED: "algebra_embed (alg_canon1 x) = algebra_embed x"
   by (import prob_algebra ALG_CANON1_EMBED)
 
-lemma ALG_CANON_MERGE_EMBED: "ALL (l::bool list) b::bool list list.
-   algebra_embed (alg_canon_merge l b) = algebra_embed (l # b)"
+lemma ALG_CANON_MERGE_EMBED: "algebra_embed (alg_canon_merge l b) = algebra_embed (l # b)"
   by (import prob_algebra ALG_CANON_MERGE_EMBED)
 
-lemma ALG_CANON2_EMBED: "ALL x::bool list list. algebra_embed (alg_canon2 x) = algebra_embed x"
+lemma ALG_CANON2_EMBED: "algebra_embed (alg_canon2 x) = algebra_embed x"
   by (import prob_algebra ALG_CANON2_EMBED)
 
-lemma ALG_CANON_EMBED: "ALL l::bool list list. algebra_embed (alg_canon l) = algebra_embed l"
+lemma ALG_CANON_EMBED: "algebra_embed (alg_canon l) = algebra_embed l"
   by (import prob_algebra ALG_CANON_EMBED)
 
-lemma ALGEBRA_CANON_UNIV: "ALL l::bool list list.
-   algebra_canon l --> algebra_embed l = pred_set.UNIV --> l = [[]]"
+lemma ALGEBRA_CANON_UNIV: "[| algebra_canon l; algebra_embed l = pred_set.UNIV |] ==> l = [[]]"
   by (import prob_algebra ALGEBRA_CANON_UNIV)
 
-lemma ALG_CANON_REP: "ALL (b::bool list list) c::bool list list.
-   (alg_canon b = alg_canon c) = (algebra_embed b = algebra_embed c)"
+lemma ALG_CANON_REP: "(alg_canon b = alg_canon c) = (algebra_embed b = algebra_embed c)"
   by (import prob_algebra ALG_CANON_REP)
 
-lemma ALGEBRA_CANON_EMBED_EMPTY: "ALL l::bool list list.
-   algebra_canon l --> (ALL v::nat => bool. ~ algebra_embed l v) = (l = [])"
+lemma ALGEBRA_CANON_EMBED_EMPTY: "algebra_canon l ==> (ALL v. ~ algebra_embed l v) = (l = [])"
   by (import prob_algebra ALGEBRA_CANON_EMBED_EMPTY)
 
-lemma ALGEBRA_CANON_EMBED_UNIV: "ALL l::bool list list.
-   algebra_canon l --> All (algebra_embed l) = (l = [[]])"
+lemma ALGEBRA_CANON_EMBED_UNIV: "algebra_canon l ==> All (algebra_embed l) = (l = [[]])"
   by (import prob_algebra ALGEBRA_CANON_EMBED_UNIV)
 
-lemma MEASURABLE_ALGEBRA: "ALL b::bool list list. measurable (algebra_embed b)"
+lemma MEASURABLE_ALGEBRA: "measurable (algebra_embed b)"
   by (import prob_algebra MEASURABLE_ALGEBRA)
 
 lemma MEASURABLE_BASIC: "measurable EMPTY &
-measurable pred_set.UNIV &
-(ALL b::bool. measurable (%s::nat => bool. SHD s = b))"
+measurable pred_set.UNIV & (ALL b. measurable (%s. SHD s = b))"
   by (import prob_algebra MEASURABLE_BASIC)
 
-lemma MEASURABLE_SHD: "ALL b::bool. measurable (%s::nat => bool. SHD s = b)"
+lemma MEASURABLE_SHD: "measurable (%s. SHD s = b)"
   by (import prob_algebra MEASURABLE_SHD)
 
-lemma ALGEBRA_EMBED_COMPL: "ALL l::bool list list.
-   EX l'::bool list list. COMPL (algebra_embed l) = algebra_embed l'"
+lemma ALGEBRA_EMBED_COMPL: "EX l'. COMPL (algebra_embed l) = algebra_embed l'"
   by (import prob_algebra ALGEBRA_EMBED_COMPL)
 
-lemma MEASURABLE_COMPL: "ALL s::(nat => bool) => bool. measurable (COMPL s) = measurable s"
+lemma MEASURABLE_COMPL: "measurable (COMPL s) = measurable s"
   by (import prob_algebra MEASURABLE_COMPL)
 
-lemma MEASURABLE_UNION: "ALL (s::(nat => bool) => bool) t::(nat => bool) => bool.
-   measurable s & measurable t --> measurable (pred_set.UNION s t)"
+lemma MEASURABLE_UNION: "measurable s & measurable t ==> measurable (pred_set.UNION s t)"
   by (import prob_algebra MEASURABLE_UNION)
 
-lemma MEASURABLE_INTER: "ALL (s::(nat => bool) => bool) t::(nat => bool) => bool.
-   measurable s & measurable t --> measurable (pred_set.INTER s t)"
+lemma MEASURABLE_INTER: "measurable s & measurable t ==> measurable (pred_set.INTER s t)"
   by (import prob_algebra MEASURABLE_INTER)
 
-lemma MEASURABLE_STL: "ALL p::(nat => bool) => bool. measurable (p o STL) = measurable p"
+lemma MEASURABLE_STL: "measurable (p o STL) = measurable p"
   by (import prob_algebra MEASURABLE_STL)
 
-lemma MEASURABLE_SDROP: "ALL (n::nat) p::(nat => bool) => bool.
-   measurable (p o SDROP n) = measurable p"
+lemma MEASURABLE_SDROP: "measurable (p o SDROP n) = measurable p"
   by (import prob_algebra MEASURABLE_SDROP)
 
-lemma MEASURABLE_INTER_HALVES: "ALL p::(nat => bool) => bool.
-   (measurable (pred_set.INTER (%x::nat => bool. SHD x = True) p) &
-    measurable (pred_set.INTER (%x::nat => bool. SHD x = False) p)) =
-   measurable p"
+lemma MEASURABLE_INTER_HALVES: "(measurable (pred_set.INTER (%x. SHD x = True) p) &
+ measurable (pred_set.INTER (%x. SHD x = False) p)) =
+measurable p"
   by (import prob_algebra MEASURABLE_INTER_HALVES)
 
-lemma MEASURABLE_HALVES: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool.
-   measurable
-    (pred_set.UNION (pred_set.INTER (%x::nat => bool. SHD x = True) p)
-      (pred_set.INTER (%x::nat => bool. SHD x = False) q)) =
-   (measurable (pred_set.INTER (%x::nat => bool. SHD x = True) p) &
-    measurable (pred_set.INTER (%x::nat => bool. SHD x = False) q))"
+lemma MEASURABLE_HALVES: "measurable
+ (pred_set.UNION (pred_set.INTER (%x. SHD x = True) p)
+   (pred_set.INTER (%x. SHD x = False) q)) =
+(measurable (pred_set.INTER (%x. SHD x = True) p) &
+ measurable (pred_set.INTER (%x. SHD x = False) q))"
   by (import prob_algebra MEASURABLE_HALVES)
 
-lemma MEASURABLE_INTER_SHD: "ALL (b::bool) p::(nat => bool) => bool.
-   measurable (pred_set.INTER (%x::nat => bool. SHD x = b) (p o STL)) =
-   measurable p"
+lemma MEASURABLE_INTER_SHD: "measurable (pred_set.INTER (%x. SHD x = b) (p o STL)) = measurable p"
   by (import prob_algebra MEASURABLE_INTER_SHD)
 
 ;end_setup
@@ -1335,8 +954,7 @@
   alg_measure :: "bool list list => real" 
 
 specification (alg_measure_primdef: alg_measure) alg_measure_def: "alg_measure [] = 0 &
-(ALL (l::bool list) rest::bool list list.
-    alg_measure (l # rest) = (1 / 2) ^ length l + alg_measure rest)"
+(ALL l rest. alg_measure (l # rest) = (1 / 2) ^ length l + alg_measure rest)"
   by (import prob alg_measure_def)
 
 consts
@@ -1344,16 +962,12 @@
 
 defs
   algebra_measure_primdef: "algebra_measure ==
-%b::bool list list.
-   inf (%r::real.
-           EX c::bool list list.
-              algebra_embed b = algebra_embed c & alg_measure c = r)"
+%b. prob_extra.inf
+     (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)"
 
-lemma algebra_measure_def: "ALL b::bool list list.
-   algebra_measure b =
-   inf (%r::real.
-           EX c::bool list list.
-              algebra_embed b = algebra_embed c & alg_measure c = r)"
+lemma algebra_measure_def: "algebra_measure b =
+prob_extra.inf
+ (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)"
   by (import prob algebra_measure_def)
 
 consts
@@ -1361,209 +975,148 @@
 
 defs
   prob_primdef: "prob ==
-%s::(nat => bool) => bool.
-   sup (%r::real.
-           EX b::bool list list.
-              algebra_measure b = r & SUBSET (algebra_embed b) s)"
+%s. real.sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)"
 
-lemma prob_def: "ALL s::(nat => bool) => bool.
-   prob s =
-   sup (%r::real.
-           EX b::bool list list.
-              algebra_measure b = r & SUBSET (algebra_embed b) s)"
+lemma prob_def: "prob s =
+real.sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)"
   by (import prob prob_def)
 
-lemma ALG_TWINS_MEASURE: "(All::(bool list => bool) => bool)
- (%l::bool list.
-     (op =::real => real => bool)
-      ((op +::real => real => real)
-        ((op ^::real => nat => real)
-          ((op /::real => real => real) (1::real)
-            ((number_of \<Colon> int => real)
-              ((Int.Bit0 \<Colon> int => int)
-                ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-          ((size::bool list => nat)
-            ((SNOC::bool => bool list => bool list) (True::bool) l)))
-        ((op ^::real => nat => real)
-          ((op /::real => real => real) (1::real)
-            ((number_of \<Colon> int => real)
-              ((Int.Bit0 \<Colon> int => int)
-                ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-          ((size::bool list => nat)
-            ((SNOC::bool => bool list => bool list) (False::bool) l))))
-      ((op ^::real => nat => real)
-        ((op /::real => real => real) (1::real)
-          ((number_of \<Colon> int => real)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-        ((size::bool list => nat) l)))"
+lemma ALG_TWINS_MEASURE: "((1::real) / (2::real)) ^ length (SNOC True (l::bool list)) +
+((1::real) / (2::real)) ^ length (SNOC False l) =
+((1::real) / (2::real)) ^ length l"
   by (import prob ALG_TWINS_MEASURE)
 
 lemma ALG_MEASURE_BASIC: "alg_measure [] = 0 &
-alg_measure [[]] = 1 & (ALL b::bool. alg_measure [[b]] = 1 / 2)"
+alg_measure [[]] = 1 & (ALL b. alg_measure [[b]] = 1 / 2)"
   by (import prob ALG_MEASURE_BASIC)
 
-lemma ALG_MEASURE_POS: "ALL l::bool list list. 0 <= alg_measure l"
+lemma ALG_MEASURE_POS: "0 <= alg_measure l"
   by (import prob ALG_MEASURE_POS)
 
-lemma ALG_MEASURE_APPEND: "ALL (l1::bool list list) l2::bool list list.
-   alg_measure (l1 @ l2) = alg_measure l1 + alg_measure l2"
+lemma ALG_MEASURE_APPEND: "alg_measure (l1 @ l2) = alg_measure l1 + alg_measure l2"
   by (import prob ALG_MEASURE_APPEND)
 
-lemma ALG_MEASURE_TLS: "ALL (l::bool list list) b::bool.
-   2 * alg_measure (map (op # b) l) = alg_measure l"
+lemma ALG_MEASURE_TLS: "2 * alg_measure (map (op # b) l) = alg_measure l"
   by (import prob ALG_MEASURE_TLS)
 
-lemma ALG_CANON_PREFS_MONO: "ALL (l::bool list) b::bool list list.
-   alg_measure (alg_canon_prefs l b) <= alg_measure (l # b)"
+lemma ALG_CANON_PREFS_MONO: "alg_measure (alg_canon_prefs l b) <= alg_measure (l # b)"
   by (import prob ALG_CANON_PREFS_MONO)
 
-lemma ALG_CANON_FIND_MONO: "ALL (l::bool list) b::bool list list.
-   alg_measure (alg_canon_find l b) <= alg_measure (l # b)"
+lemma ALG_CANON_FIND_MONO: "alg_measure (alg_canon_find l b) <= alg_measure (l # b)"
   by (import prob ALG_CANON_FIND_MONO)
 
-lemma ALG_CANON1_MONO: "ALL x::bool list list. alg_measure (alg_canon1 x) <= alg_measure x"
+lemma ALG_CANON1_MONO: "alg_measure (alg_canon1 x) <= alg_measure x"
   by (import prob ALG_CANON1_MONO)
 
-lemma ALG_CANON_MERGE_MONO: "ALL (l::bool list) b::bool list list.
-   alg_measure (alg_canon_merge l b) <= alg_measure (l # b)"
+lemma ALG_CANON_MERGE_MONO: "alg_measure (alg_canon_merge l b) <= alg_measure (l # b)"
   by (import prob ALG_CANON_MERGE_MONO)
 
-lemma ALG_CANON2_MONO: "ALL x::bool list list. alg_measure (alg_canon2 x) <= alg_measure x"
+lemma ALG_CANON2_MONO: "alg_measure (alg_canon2 x) <= alg_measure x"
   by (import prob ALG_CANON2_MONO)
 
-lemma ALG_CANON_MONO: "ALL l::bool list list. alg_measure (alg_canon l) <= alg_measure l"
+lemma ALG_CANON_MONO: "alg_measure (alg_canon l) <= alg_measure l"
   by (import prob ALG_CANON_MONO)
 
-lemma ALGEBRA_MEASURE_DEF_ALT: "ALL l::bool list list. algebra_measure l = alg_measure (alg_canon l)"
+lemma ALGEBRA_MEASURE_DEF_ALT: "algebra_measure l = alg_measure (alg_canon l)"
   by (import prob ALGEBRA_MEASURE_DEF_ALT)
 
 lemma ALGEBRA_MEASURE_BASIC: "algebra_measure [] = 0 &
-algebra_measure [[]] = 1 & (ALL b::bool. algebra_measure [[b]] = 1 / 2)"
+algebra_measure [[]] = 1 & (ALL b. algebra_measure [[b]] = 1 / 2)"
   by (import prob ALGEBRA_MEASURE_BASIC)
 
-lemma ALGEBRA_CANON_MEASURE_MAX: "ALL l::bool list list. algebra_canon l --> alg_measure l <= 1"
+lemma ALGEBRA_CANON_MEASURE_MAX: "algebra_canon l ==> alg_measure l <= 1"
   by (import prob ALGEBRA_CANON_MEASURE_MAX)
 
-lemma ALGEBRA_MEASURE_MAX: "ALL l::bool list list. algebra_measure l <= 1"
+lemma ALGEBRA_MEASURE_MAX: "algebra_measure l <= 1"
   by (import prob ALGEBRA_MEASURE_MAX)
 
-lemma ALGEBRA_MEASURE_MONO_EMBED: "ALL (x::bool list list) xa::bool list list.
-   SUBSET (algebra_embed x) (algebra_embed xa) -->
-   algebra_measure x <= algebra_measure xa"
+lemma ALGEBRA_MEASURE_MONO_EMBED: "SUBSET (algebra_embed x) (algebra_embed xa)
+==> algebra_measure x <= algebra_measure xa"
   by (import prob ALGEBRA_MEASURE_MONO_EMBED)
 
-lemma ALG_MEASURE_COMPL: "ALL l::bool list list.
-   algebra_canon l -->
-   (ALL c::bool list list.
-       algebra_canon c -->
-       COMPL (algebra_embed l) = algebra_embed c -->
-       alg_measure l + alg_measure c = 1)"
+lemma ALG_MEASURE_COMPL: "[| algebra_canon l; algebra_canon c;
+   COMPL (algebra_embed l) = algebra_embed c |]
+==> alg_measure l + alg_measure c = 1"
   by (import prob ALG_MEASURE_COMPL)
 
-lemma ALG_MEASURE_ADDITIVE: "ALL l::bool list list.
-   algebra_canon l -->
-   (ALL c::bool list list.
-       algebra_canon c -->
-       (ALL d::bool list list.
-           algebra_canon d -->
-           pred_set.INTER (algebra_embed c) (algebra_embed d) = EMPTY &
-           algebra_embed l =
-           pred_set.UNION (algebra_embed c) (algebra_embed d) -->
-           alg_measure l = alg_measure c + alg_measure d))"
+lemma ALG_MEASURE_ADDITIVE: "[| algebra_canon l; algebra_canon c; algebra_canon d;
+   pred_set.INTER (algebra_embed c) (algebra_embed d) = EMPTY &
+   algebra_embed l = pred_set.UNION (algebra_embed c) (algebra_embed d) |]
+==> alg_measure l = alg_measure c + alg_measure d"
   by (import prob ALG_MEASURE_ADDITIVE)
 
-lemma PROB_ALGEBRA: "ALL l::bool list list. prob (algebra_embed l) = algebra_measure l"
+lemma PROB_ALGEBRA: "prob (algebra_embed l) = algebra_measure l"
   by (import prob PROB_ALGEBRA)
 
 lemma PROB_BASIC: "prob EMPTY = 0 &
-prob pred_set.UNIV = 1 &
-(ALL b::bool. prob (%s::nat => bool. SHD s = b) = 1 / 2)"
+prob pred_set.UNIV = 1 & (ALL b. prob (%s. SHD s = b) = 1 / 2)"
   by (import prob PROB_BASIC)
 
-lemma PROB_ADDITIVE: "ALL (s::(nat => bool) => bool) t::(nat => bool) => bool.
-   measurable s & measurable t & pred_set.INTER s t = EMPTY -->
-   prob (pred_set.UNION s t) = prob s + prob t"
+lemma PROB_ADDITIVE: "measurable s & measurable t & pred_set.INTER s t = EMPTY
+==> prob (pred_set.UNION s t) = prob s + prob t"
   by (import prob PROB_ADDITIVE)
 
-lemma PROB_COMPL: "ALL s::(nat => bool) => bool. measurable s --> prob (COMPL s) = 1 - prob s"
+lemma PROB_COMPL: "measurable s ==> prob (COMPL s) = 1 - prob s"
   by (import prob PROB_COMPL)
 
-lemma PROB_SUP_EXISTS1: "ALL s::(nat => bool) => bool.
-   EX (x::real) b::bool list list.
-      algebra_measure b = x & SUBSET (algebra_embed b) s"
+lemma PROB_SUP_EXISTS1: "EX x b. algebra_measure b = x & SUBSET (algebra_embed b) s"
   by (import prob PROB_SUP_EXISTS1)
 
-lemma PROB_SUP_EXISTS2: "ALL s::(nat => bool) => bool.
-   EX x::real.
-      ALL r::real.
-         (EX b::bool list list.
-             algebra_measure b = r & SUBSET (algebra_embed b) s) -->
+lemma PROB_SUP_EXISTS2: "EX x. ALL r.
+         (EX b. algebra_measure b = r & SUBSET (algebra_embed b) s) -->
          r <= x"
   by (import prob PROB_SUP_EXISTS2)
 
-lemma PROB_LE_X: "ALL (s::(nat => bool) => bool) x::real.
-   (ALL s'::(nat => bool) => bool.
-       measurable s' & SUBSET s' s --> prob s' <= x) -->
-   prob s <= x"
+lemma PROB_LE_X: "(!!s'. measurable s' & SUBSET s' s ==> prob s' <= x) ==> prob s <= x"
   by (import prob PROB_LE_X)
 
-lemma X_LE_PROB: "ALL (s::(nat => bool) => bool) x::real.
-   (EX s'::(nat => bool) => bool.
-       measurable s' & SUBSET s' s & x <= prob s') -->
-   x <= prob s"
+lemma X_LE_PROB: "EX s'. measurable s' & SUBSET s' s & x <= prob s' ==> x <= prob s"
   by (import prob X_LE_PROB)
 
-lemma PROB_SUBSET_MONO: "ALL (s::(nat => bool) => bool) t::(nat => bool) => bool.
-   SUBSET s t --> prob s <= prob t"
+lemma PROB_SUBSET_MONO: "SUBSET s t ==> prob s <= prob t"
   by (import prob PROB_SUBSET_MONO)
 
-lemma PROB_ALG: "ALL x::bool list. prob (alg_embed x) = (1 / 2) ^ length x"
+lemma PROB_ALG: "prob (alg_embed x) = (1 / 2) ^ length x"
   by (import prob PROB_ALG)
 
-lemma PROB_STL: "ALL p::(nat => bool) => bool. measurable p --> prob (p o STL) = prob p"
+lemma PROB_STL: "measurable p ==> prob (p o STL) = prob p"
   by (import prob PROB_STL)
 
-lemma PROB_SDROP: "ALL (n::nat) p::(nat => bool) => bool.
-   measurable p --> prob (p o SDROP n) = prob p"
+lemma PROB_SDROP: "measurable p ==> prob (p o SDROP n) = prob p"
   by (import prob PROB_SDROP)
 
-lemma PROB_INTER_HALVES: "ALL p::(nat => bool) => bool.
-   measurable p -->
-   prob (pred_set.INTER (%x::nat => bool. SHD x = True) p) +
-   prob (pred_set.INTER (%x::nat => bool. SHD x = False) p) =
-   prob p"
+lemma PROB_INTER_HALVES: "measurable p
+==> prob (pred_set.INTER (%x. SHD x = True) p) +
+    prob (pred_set.INTER (%x. SHD x = False) p) =
+    prob p"
   by (import prob PROB_INTER_HALVES)
 
-lemma PROB_INTER_SHD: "ALL (b::bool) p::(nat => bool) => bool.
-   measurable p -->
-   prob (pred_set.INTER (%x::nat => bool. SHD x = b) (p o STL)) =
-   1 / 2 * prob p"
+lemma PROB_INTER_SHD: "measurable p
+==> prob (pred_set.INTER (%x. SHD x = b) (p o STL)) = 1 / 2 * prob p"
   by (import prob PROB_INTER_SHD)
 
-lemma ALGEBRA_MEASURE_POS: "ALL l::bool list list. 0 <= algebra_measure l"
+lemma ALGEBRA_MEASURE_POS: "0 <= algebra_measure l"
   by (import prob ALGEBRA_MEASURE_POS)
 
-lemma ALGEBRA_MEASURE_RANGE: "ALL l::bool list list. 0 <= algebra_measure l & algebra_measure l <= 1"
+lemma ALGEBRA_MEASURE_RANGE: "0 <= algebra_measure l & algebra_measure l <= 1"
   by (import prob ALGEBRA_MEASURE_RANGE)
 
-lemma PROB_POS: "ALL p::(nat => bool) => bool. 0 <= prob p"
+lemma PROB_POS: "0 <= prob p"
   by (import prob PROB_POS)
 
-lemma PROB_MAX: "ALL p::(nat => bool) => bool. prob p <= 1"
+lemma PROB_MAX: "prob p <= 1"
   by (import prob PROB_MAX)
 
-lemma PROB_RANGE: "ALL p::(nat => bool) => bool. 0 <= prob p & prob p <= 1"
+lemma PROB_RANGE: "0 <= prob p & prob p <= 1"
   by (import prob PROB_RANGE)
 
-lemma ABS_PROB: "ALL p::(nat => bool) => bool. abs (prob p) = prob p"
+lemma ABS_PROB: "abs (prob p) = prob p"
   by (import prob ABS_PROB)
 
-lemma PROB_SHD: "ALL b::bool. prob (%s::nat => bool. SHD s = b) = 1 / 2"
+lemma PROB_SHD: "prob (%s. SHD s = b) = 1 / 2"
   by (import prob PROB_SHD)
 
-lemma PROB_COMPL_LE1: "ALL (p::(nat => bool) => bool) r::real.
-   measurable p --> (prob (COMPL p) <= r) = (1 - r <= prob p)"
+lemma PROB_COMPL_LE1: "measurable p ==> (prob (COMPL p) <= r) = (1 - r <= prob p)"
   by (import prob PROB_COMPL_LE1)
 
 ;end_setup
@@ -1583,41 +1136,41 @@
   pseudo_linear_tl :: "nat => nat => nat => nat => nat" 
 
 defs
-  pseudo_linear_tl_primdef: "pseudo_linear_tl ==
-%(a::nat) (b::nat) (n::nat) x::nat. (a * x + b) mod (2 * n + 1)"
+  pseudo_linear_tl_primdef: "pseudo_linear_tl == %a b n x. (a * x + b) mod (2 * n + 1)"
 
-lemma pseudo_linear_tl_def: "ALL (a::nat) (b::nat) (n::nat) x::nat.
-   pseudo_linear_tl a b n x = (a * x + b) mod (2 * n + 1)"
+lemma pseudo_linear_tl_def: "pseudo_linear_tl a b n x = (a * x + b) mod (2 * n + 1)"
   by (import prob_pseudo pseudo_linear_tl_def)
 
-lemma PSEUDO_LINEAR1_EXECUTE: "EX x::nat => nat => bool.
-   (ALL xa::nat. SHD (x xa) = pseudo_linear_hd xa) &
-   (ALL xa::nat.
-       STL (x xa) =
-       x (pseudo_linear_tl
-           (NUMERAL
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1
-                 (NUMERAL_BIT1
-                   (NUMERAL_BIT2 (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO)))))))
-           (NUMERAL
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1
-                 (NUMERAL_BIT1
-                   (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO)))))))
-           (NUMERAL
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1
-                 (NUMERAL_BIT1
-                   (NUMERAL_BIT1 (NUMERAL_BIT2 (NUMERAL_BIT1 ALT_ZERO)))))))
-           xa))"
+lemma PSEUDO_LINEAR1_EXECUTE: "EX x. (ALL xa. SHD (x xa) = pseudo_linear_hd xa) &
+      (ALL xa.
+          STL (x xa) =
+          x (pseudo_linear_tl
+              (NUMERAL
+                (NUMERAL_BIT1
+                  (NUMERAL_BIT1
+                    (NUMERAL_BIT1
+                      (NUMERAL_BIT2
+                        (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO)))))))
+              (NUMERAL
+                (NUMERAL_BIT1
+                  (NUMERAL_BIT1
+                    (NUMERAL_BIT1
+                      (NUMERAL_BIT1
+                        (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO)))))))
+              (NUMERAL
+                (NUMERAL_BIT1
+                  (NUMERAL_BIT1
+                    (NUMERAL_BIT1
+                      (NUMERAL_BIT1
+                        (NUMERAL_BIT2 (NUMERAL_BIT1 ALT_ZERO)))))))
+              xa))"
   by (import prob_pseudo PSEUDO_LINEAR1_EXECUTE)
 
 consts
   pseudo_linear1 :: "nat => nat => bool" 
 
-specification (pseudo_linear1_primdef: pseudo_linear1) pseudo_linear1_def: "(ALL x::nat. SHD (pseudo_linear1 x) = pseudo_linear_hd x) &
-(ALL x::nat.
+specification (pseudo_linear1_primdef: pseudo_linear1) pseudo_linear1_def: "(ALL x. SHD (pseudo_linear1 x) = pseudo_linear_hd x) &
+(ALL x.
     STL (pseudo_linear1 x) =
     pseudo_linear1
      (pseudo_linear_tl
@@ -1657,13 +1210,11 @@
 
 defs
   indep_set_primdef: "indep_set ==
-%(p::(nat => bool) => bool) q::(nat => bool) => bool.
-   measurable p & measurable q & prob (pred_set.INTER p q) = prob p * prob q"
+%p q. measurable p &
+      measurable q & prob (pred_set.INTER p q) = prob p * prob q"
 
-lemma indep_set_def: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool.
-   indep_set p q =
-   (measurable p &
-    measurable q & prob (pred_set.INTER p q) = prob p * prob q)"
+lemma indep_set_def: "indep_set p q =
+(measurable p & measurable q & prob (pred_set.INTER p q) = prob p * prob q)"
   by (import prob_indep indep_set_def)
 
 consts
@@ -1671,24 +1222,19 @@
 
 defs
   alg_cover_set_primdef: "alg_cover_set ==
-%l::bool list list.
-   alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV"
+%l. alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV"
 
-lemma alg_cover_set_def: "ALL l::bool list list.
-   alg_cover_set l =
-   (alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV)"
+lemma alg_cover_set_def: "alg_cover_set l =
+(alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV)"
   by (import prob_indep alg_cover_set_def)
 
 consts
   alg_cover :: "bool list list => (nat => bool) => bool list" 
 
 defs
-  alg_cover_primdef: "alg_cover ==
-%(l::bool list list) x::nat => bool.
-   SOME b::bool list. b mem l & alg_embed b x"
+  alg_cover_primdef: "alg_cover == %l x. SOME b. List.member l b & alg_embed b x"
 
-lemma alg_cover_def: "ALL (l::bool list list) x::nat => bool.
-   alg_cover l x = (SOME b::bool list. b mem l & alg_embed b x)"
+lemma alg_cover_def: "alg_cover l x = (SOME b. List.member l b & alg_embed b x)"
   by (import prob_indep alg_cover_def)
 
 consts
@@ -1696,195 +1242,148 @@
 
 defs
   indep_primdef: "indep ==
-%f::(nat => bool) => 'a::type * (nat => bool).
-   EX (l::bool list list) r::bool list => 'a::type.
-      alg_cover_set l &
-      (ALL s::nat => bool.
-          f s =
-          (let c::bool list = alg_cover l s in (r c, SDROP (length c) s)))"
+%f. EX l r.
+       alg_cover_set l &
+       (ALL s. f s = (let c = alg_cover l s in (r c, SDROP (length c) s)))"
 
-lemma indep_def: "ALL f::(nat => bool) => 'a::type * (nat => bool).
-   indep f =
-   (EX (l::bool list list) r::bool list => 'a::type.
-       alg_cover_set l &
-       (ALL s::nat => bool.
-           f s =
-           (let c::bool list = alg_cover l s in (r c, SDROP (length c) s))))"
+lemma indep_def: "indep f =
+(EX l r.
+    alg_cover_set l &
+    (ALL s. f s = (let c = alg_cover l s in (r c, SDROP (length c) s))))"
   by (import prob_indep indep_def)
 
-lemma INDEP_SET_BASIC: "ALL p::(nat => bool) => bool.
-   measurable p --> indep_set EMPTY p & indep_set pred_set.UNIV p"
+lemma INDEP_SET_BASIC: "measurable p ==> indep_set EMPTY p & indep_set pred_set.UNIV p"
   by (import prob_indep INDEP_SET_BASIC)
 
-lemma INDEP_SET_SYM: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool.
-   indep_set p q = indep_set p q"
+lemma INDEP_SET_SYM: "indep_set p q = indep_set p q"
   by (import prob_indep INDEP_SET_SYM)
 
-lemma INDEP_SET_DISJOINT_DECOMP: "ALL (p::(nat => bool) => bool) (q::(nat => bool) => bool)
-   r::(nat => bool) => bool.
-   indep_set p r & indep_set q r & pred_set.INTER p q = EMPTY -->
-   indep_set (pred_set.UNION p q) r"
+lemma INDEP_SET_DISJOINT_DECOMP: "indep_set p r & indep_set q r & pred_set.INTER p q = EMPTY
+==> indep_set (pred_set.UNION p q) r"
   by (import prob_indep INDEP_SET_DISJOINT_DECOMP)
 
 lemma ALG_COVER_SET_BASIC: "~ alg_cover_set [] & alg_cover_set [[]] & alg_cover_set [[True], [False]]"
   by (import prob_indep ALG_COVER_SET_BASIC)
 
-lemma ALG_COVER_WELL_DEFINED: "ALL (l::bool list list) x::nat => bool.
-   alg_cover_set l --> alg_cover l x mem l & alg_embed (alg_cover l x) x"
+lemma ALG_COVER_WELL_DEFINED: "alg_cover_set l
+==> List.member l (alg_cover l x) & alg_embed (alg_cover l x) x"
   by (import prob_indep ALG_COVER_WELL_DEFINED)
 
 lemma ALG_COVER_UNIV: "alg_cover [[]] = K []"
   by (import prob_indep ALG_COVER_UNIV)
 
-lemma MAP_CONS_TL_FILTER: "ALL (l::bool list list) b::bool.
-   ~ [] mem l -->
-   map (op # b) (map tl [x::bool list<-l. hd x = b]) =
-   [x::bool list<-l. hd x = b]"
+lemma MAP_CONS_TL_FILTER: "~ List.member (l::bool list list) []
+==> map (op # (b::bool)) (map tl [x::bool list<-l. hd x = b]) =
+    [x::bool list<-l. hd x = b]"
   by (import prob_indep MAP_CONS_TL_FILTER)
 
-lemma ALG_COVER_SET_CASES_THM: "ALL l::bool list list.
-   alg_cover_set l =
-   (l = [[]] |
-    (EX (x::bool list list) xa::bool list list.
-        alg_cover_set x &
-        alg_cover_set xa & l = map (op # True) x @ map (op # False) xa))"
+lemma ALG_COVER_SET_CASES_THM: "alg_cover_set l =
+(l = [[]] |
+ (EX x xa.
+     alg_cover_set x &
+     alg_cover_set xa & l = map (op # True) x @ map (op # False) xa))"
   by (import prob_indep ALG_COVER_SET_CASES_THM)
 
-lemma ALG_COVER_SET_CASES: "ALL P::bool list list => bool.
-   P [[]] &
-   (ALL (l1::bool list list) l2::bool list list.
+lemma ALG_COVER_SET_CASES: "[| P [[]] &
+   (ALL l1 l2.
        alg_cover_set l1 &
        alg_cover_set l2 &
        alg_cover_set (map (op # True) l1 @ map (op # False) l2) -->
-       P (map (op # True) l1 @ map (op # False) l2)) -->
-   (ALL l::bool list list. alg_cover_set l --> P l)"
+       P (map (op # True) l1 @ map (op # False) l2));
+   alg_cover_set l |]
+==> P l"
   by (import prob_indep ALG_COVER_SET_CASES)
 
-lemma ALG_COVER_SET_INDUCTION: "ALL P::bool list list => bool.
-   P [[]] &
-   (ALL (l1::bool list list) l2::bool list list.
+lemma ALG_COVER_SET_INDUCTION: "[| P [[]] &
+   (ALL l1 l2.
        alg_cover_set l1 &
        alg_cover_set l2 &
        P l1 &
        P l2 & alg_cover_set (map (op # True) l1 @ map (op # False) l2) -->
-       P (map (op # True) l1 @ map (op # False) l2)) -->
-   (ALL l::bool list list. alg_cover_set l --> P l)"
+       P (map (op # True) l1 @ map (op # False) l2));
+   alg_cover_set l |]
+==> P l"
   by (import prob_indep ALG_COVER_SET_INDUCTION)
 
-lemma ALG_COVER_EXISTS_UNIQUE: "ALL l::bool list list.
-   alg_cover_set l -->
-   (ALL s::nat => bool. EX! x::bool list. x mem l & alg_embed x s)"
+lemma ALG_COVER_EXISTS_UNIQUE: "alg_cover_set l ==> EX! x. List.member l x & alg_embed x s"
   by (import prob_indep ALG_COVER_EXISTS_UNIQUE)
 
-lemma ALG_COVER_UNIQUE: "ALL (l::bool list list) (x::bool list) s::nat => bool.
-   alg_cover_set l & x mem l & alg_embed x s --> alg_cover l s = x"
+lemma ALG_COVER_UNIQUE: "alg_cover_set l & List.member l x & alg_embed x s ==> alg_cover l s = x"
   by (import prob_indep ALG_COVER_UNIQUE)
 
-lemma ALG_COVER_STEP: "ALL (l1::bool list list) (l2::bool list list) (h::bool) t::nat => bool.
-   alg_cover_set l1 & alg_cover_set l2 -->
-   alg_cover (map (op # True) l1 @ map (op # False) l2) (SCONS h t) =
-   (if h then True # alg_cover l1 t else False # alg_cover l2 t)"
+lemma ALG_COVER_STEP: "alg_cover_set l1 & alg_cover_set l2
+==> alg_cover (map (op # True) l1 @ map (op # False) l2) (SCONS h t) =
+    (if h then True # alg_cover l1 t else False # alg_cover l2 t)"
   by (import prob_indep ALG_COVER_STEP)
 
-lemma ALG_COVER_HEAD: "ALL l::bool list list.
-   alg_cover_set l -->
-   (ALL f::bool list => bool. f o alg_cover l = algebra_embed (filter f l))"
+lemma ALG_COVER_HEAD: "alg_cover_set l ==> f o alg_cover l = algebra_embed (filter f l)"
   by (import prob_indep ALG_COVER_HEAD)
 
-lemma ALG_COVER_TAIL_STEP: "ALL (l1::bool list list) (l2::bool list list) q::(nat => bool) => bool.
-   alg_cover_set l1 & alg_cover_set l2 -->
-   q o
-   (%x::nat => bool.
-       SDROP
-        (length (alg_cover (map (op # True) l1 @ map (op # False) l2) x))
-        x) =
-   pred_set.UNION
-    (pred_set.INTER (%x::nat => bool. SHD x = True)
-      (q o ((%x::nat => bool. SDROP (length (alg_cover l1 x)) x) o STL)))
-    (pred_set.INTER (%x::nat => bool. SHD x = False)
-      (q o ((%x::nat => bool. SDROP (length (alg_cover l2 x)) x) o STL)))"
+lemma ALG_COVER_TAIL_STEP: "alg_cover_set l1 & alg_cover_set l2
+==> q o
+    (%x. SDROP
+          (length (alg_cover (map (op # True) l1 @ map (op # False) l2) x))
+          x) =
+    pred_set.UNION
+     (pred_set.INTER (%x. SHD x = True)
+       (q o ((%x. SDROP (length (alg_cover l1 x)) x) o STL)))
+     (pred_set.INTER (%x. SHD x = False)
+       (q o ((%x. SDROP (length (alg_cover l2 x)) x) o STL)))"
   by (import prob_indep ALG_COVER_TAIL_STEP)
 
-lemma ALG_COVER_TAIL_MEASURABLE: "ALL l::bool list list.
-   alg_cover_set l -->
-   (ALL q::(nat => bool) => bool.
-       measurable
-        (q o (%x::nat => bool. SDROP (length (alg_cover l x)) x)) =
-       measurable q)"
+lemma ALG_COVER_TAIL_MEASURABLE: "alg_cover_set l
+==> measurable (q o (%x. SDROP (length (alg_cover l x)) x)) = measurable q"
   by (import prob_indep ALG_COVER_TAIL_MEASURABLE)
 
-lemma ALG_COVER_TAIL_PROB: "ALL l::bool list list.
-   alg_cover_set l -->
-   (ALL q::(nat => bool) => bool.
-       measurable q -->
-       prob (q o (%x::nat => bool. SDROP (length (alg_cover l x)) x)) =
-       prob q)"
+lemma ALG_COVER_TAIL_PROB: "[| alg_cover_set l; measurable q |]
+==> prob (q o (%x. SDROP (length (alg_cover l x)) x)) = prob q"
   by (import prob_indep ALG_COVER_TAIL_PROB)
 
-lemma INDEP_INDEP_SET_LEMMA: "ALL l::bool list list.
-   alg_cover_set l -->
-   (ALL q::(nat => bool) => bool.
-       measurable q -->
-       (ALL x::bool list.
-           x mem l -->
-           prob
-            (pred_set.INTER (alg_embed x)
-              (q o (%x::nat => bool. SDROP (length (alg_cover l x)) x))) =
-           (1 / 2) ^ length x * prob q))"
+lemma INDEP_INDEP_SET_LEMMA: "[| alg_cover_set l; measurable q; List.member l x |]
+==> prob
+     (pred_set.INTER (alg_embed x)
+       (q o (%x. SDROP (length (alg_cover l x)) x))) =
+    (1 / 2) ^ length x * prob q"
   by (import prob_indep INDEP_INDEP_SET_LEMMA)
 
-lemma INDEP_SET_LIST: "ALL (q::(nat => bool) => bool) l::bool list list.
-   alg_sorted l &
-   alg_prefixfree l &
-   measurable q &
-   (ALL x::bool list. x mem l --> indep_set (alg_embed x) q) -->
-   indep_set (algebra_embed l) q"
+lemma INDEP_SET_LIST: "alg_sorted l &
+alg_prefixfree l &
+measurable q & (ALL x. List.member l x --> indep_set (alg_embed x) q)
+==> indep_set (algebra_embed l) q"
   by (import prob_indep INDEP_SET_LIST)
 
-lemma INDEP_INDEP_SET: "ALL (f::(nat => bool) => 'a::type * (nat => bool)) (p::'a::type => bool)
-   q::(nat => bool) => bool.
-   indep f & measurable q --> indep_set (p o (fst o f)) (q o (snd o f))"
+lemma INDEP_INDEP_SET: "indep f & measurable q ==> indep_set (p o (fst o f)) (q o (snd o f))"
   by (import prob_indep INDEP_INDEP_SET)
 
-lemma INDEP_UNIT: "ALL x::'a::type. indep (UNIT x)"
+lemma INDEP_UNIT: "indep (UNIT x)"
   by (import prob_indep INDEP_UNIT)
 
 lemma INDEP_SDEST: "indep SDEST"
   by (import prob_indep INDEP_SDEST)
 
-lemma BIND_STEP: "ALL f::(nat => bool) => 'a::type * (nat => bool).
-   BIND SDEST (%k::bool. f o SCONS k) = f"
+lemma BIND_STEP: "BIND SDEST (%k. f o SCONS k) = f"
   by (import prob_indep BIND_STEP)
 
-lemma INDEP_BIND_SDEST: "ALL f::bool => (nat => bool) => 'a::type * (nat => bool).
-   (ALL x::bool. indep (f x)) --> indep (BIND SDEST f)"
+lemma INDEP_BIND_SDEST: "(!!x. indep (f x)) ==> indep (BIND SDEST f)"
   by (import prob_indep INDEP_BIND_SDEST)
 
-lemma INDEP_BIND: "ALL (f::(nat => bool) => 'a::type * (nat => bool))
-   g::'a::type => (nat => bool) => 'b::type * (nat => bool).
-   indep f & (ALL x::'a::type. indep (g x)) --> indep (BIND f g)"
+lemma INDEP_BIND: "indep f & (ALL x. indep (g x)) ==> indep (BIND f g)"
   by (import prob_indep INDEP_BIND)
 
-lemma INDEP_PROB: "ALL (f::(nat => bool) => 'a::type * (nat => bool)) (p::'a::type => bool)
-   q::(nat => bool) => bool.
-   indep f & measurable q -->
-   prob (pred_set.INTER (p o (fst o f)) (q o (snd o f))) =
-   prob (p o (fst o f)) * prob q"
+lemma INDEP_PROB: "indep f & measurable q
+==> prob (pred_set.INTER (p o (fst o f)) (q o (snd o f))) =
+    prob (p o (fst o f)) * prob q"
   by (import prob_indep INDEP_PROB)
 
-lemma INDEP_MEASURABLE1: "ALL (f::(nat => bool) => 'a::type * (nat => bool)) p::'a::type => bool.
-   indep f --> measurable (p o (fst o f))"
+lemma INDEP_MEASURABLE1: "indep f ==> measurable (p o (fst o f))"
   by (import prob_indep INDEP_MEASURABLE1)
 
-lemma INDEP_MEASURABLE2: "ALL (f::(nat => bool) => 'a::type * (nat => bool)) q::(nat => bool) => bool.
-   indep f & measurable q --> measurable (q o (snd o f))"
+lemma INDEP_MEASURABLE2: "indep f & measurable q ==> measurable (q o (snd o f))"
   by (import prob_indep INDEP_MEASURABLE2)
 
-lemma PROB_INDEP_BOUND: "ALL (f::(nat => bool) => nat * (nat => bool)) n::nat.
-   indep f -->
-   prob (%s::nat => bool. fst (f s) < Suc n) =
-   prob (%s::nat => bool. fst (f s) < n) +
-   prob (%s::nat => bool. fst (f s) = n)"
+lemma PROB_INDEP_BOUND: "indep f
+==> prob (%s. fst (f s) < Suc n) =
+    prob (%s. fst (f s) < n) + prob (%s. fst (f s) = n)"
   by (import prob_indep PROB_INDEP_BOUND)
 
 ;end_setup
@@ -1896,47 +1395,36 @@
 
 defs
   unif_bound_primdef: "unif_bound ==
-WFREC
- (SOME R::nat => nat => bool. WF R & (ALL v::nat. R (Suc v div 2) (Suc v)))
- (%unif_bound::nat => nat.
-     nat_case 0 (%v1::nat. Suc (unif_bound (Suc v1 div 2))))"
+WFREC (SOME R. WF R & (ALL v. R (Suc v div 2) (Suc v)))
+ (%unif_bound. nat_case 0 (%v1. Suc (unif_bound (Suc v1 div 2))))"
 
 lemma unif_bound_primitive_def: "unif_bound =
-WFREC
- (SOME R::nat => nat => bool. WF R & (ALL v::nat. R (Suc v div 2) (Suc v)))
- (%unif_bound::nat => nat.
-     nat_case 0 (%v1::nat. Suc (unif_bound (Suc v1 div 2))))"
+WFREC (SOME R. WF R & (ALL v. R (Suc v div 2) (Suc v)))
+ (%unif_bound. nat_case 0 (%v1. Suc (unif_bound (Suc v1 div 2))))"
   by (import prob_uniform unif_bound_primitive_def)
 
-lemma unif_bound_def: "unif_bound 0 = 0 &
-unif_bound (Suc (v::nat)) = Suc (unif_bound (Suc v div 2))"
+lemma unif_bound_def: "unif_bound 0 = 0 & unif_bound (Suc v) = Suc (unif_bound (Suc v div 2))"
   by (import prob_uniform unif_bound_def)
 
-lemma unif_bound_ind: "ALL P::nat => bool.
-   P 0 & (ALL v::nat. P (Suc v div 2) --> P (Suc v)) --> All P"
+lemma unif_bound_ind: "P 0 & (ALL v. P (Suc v div 2) --> P (Suc v)) ==> P x"
   by (import prob_uniform unif_bound_ind)
 
-definition unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" where 
+definition
+  unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)"  where
   "unif_tupled ==
-WFREC
- (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool.
-     WF R & (ALL (s::nat => bool) v2::nat. R (Suc v2 div 2, s) (Suc v2, s)))
- (%(unif_tupled::nat * (nat => bool) => nat * (nat => bool)) (v::nat,
-     v1::nat => bool).
+WFREC (SOME R. WF R & (ALL s v2. R (Suc v2 div 2, s) (Suc v2, s)))
+ (%unif_tupled (v, v1).
      case v of 0 => (0, v1)
-     | Suc (v3::nat) =>
-         let (m::nat, s'::nat => bool) = unif_tupled (Suc v3 div 2, v1)
+     | Suc v3 =>
+         let (m, s') = unif_tupled (Suc v3 div 2, v1)
          in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))"
 
 lemma unif_tupled_primitive_def: "unif_tupled =
-WFREC
- (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool.
-     WF R & (ALL (s::nat => bool) v2::nat. R (Suc v2 div 2, s) (Suc v2, s)))
- (%(unif_tupled::nat * (nat => bool) => nat * (nat => bool)) (v::nat,
-     v1::nat => bool).
+WFREC (SOME R. WF R & (ALL s v2. R (Suc v2 div 2, s) (Suc v2, s)))
+ (%unif_tupled (v, v1).
      case v of 0 => (0, v1)
-     | Suc (v3::nat) =>
-         let (m::nat, s'::nat => bool) = unif_tupled (Suc v3 div 2, v1)
+     | Suc v3 =>
+         let (m, s') = unif_tupled (Suc v3 div 2, v1)
          in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))"
   by (import prob_uniform unif_tupled_primitive_def)
 
@@ -1944,247 +1432,160 @@
   unif :: "nat => (nat => bool) => nat * (nat => bool)" 
 
 defs
-  unif_primdef: "unif == %(x::nat) x1::nat => bool. unif_tupled (x, x1)"
+  unif_primdef: "unif == %x x1. unif_tupled (x, x1)"
 
-lemma unif_curried_def: "ALL (x::nat) x1::nat => bool. unif x x1 = unif_tupled (x, x1)"
+lemma unif_curried_def: "unif x x1 = unif_tupled (x, x1)"
   by (import prob_uniform unif_curried_def)
 
-lemma unif_def: "unif 0 (s::nat => bool) = (0, s) &
-unif (Suc (v2::nat)) s =
-(let (m::nat, s'::nat => bool) = unif (Suc v2 div 2) s
+lemma unif_def: "unif 0 s = (0, s) &
+unif (Suc v2) s =
+(let (m, s') = unif (Suc v2 div 2) s
  in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))"
   by (import prob_uniform unif_def)
 
-lemma unif_ind: "ALL P::nat => (nat => bool) => bool.
-   All (P 0) &
-   (ALL (v2::nat) s::nat => bool. P (Suc v2 div 2) s --> P (Suc v2) s) -->
-   (ALL v::nat. All (P v))"
+lemma unif_ind: "All ((P::nat => (nat => bool) => bool) (0::nat)) &
+(ALL (v2::nat) s::nat => bool. P (Suc v2 div (2::nat)) s --> P (Suc v2) s)
+==> P (v::nat) (x::nat => bool)"
   by (import prob_uniform unif_ind)
 
-definition uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" where 
+definition
+  uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)"  where
   "uniform_tupled ==
 WFREC
- (SOME R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool.
+ (SOME R.
      WF R &
-     (ALL (t::nat) (s::nat => bool) (n::nat) (res::nat) s'::nat => bool.
+     (ALL t s n res s'.
          (res, s') = unif n s & ~ res < Suc n -->
          R (t, Suc n, s') (Suc t, Suc n, s)))
- (%(uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool))
-     (v::nat, v1::nat * (nat => bool)).
-     case v of
-     0 => (%(v3::nat, v4::nat => bool).
-              case v3 of 0 => ARB | Suc (v5::nat) => (0, v4))
-           v1
-     | Suc (v2::nat) =>
-         (%(v7::nat, v8::nat => bool).
-             case v7 of 0 => ARB
-             | Suc (v9::nat) =>
-                 let (res::nat, s'::nat => bool) = unif v9 v8
-                 in if res < Suc v9 then (res, s')
-                    else uniform_tupled (v2, Suc v9, s'))
-          v1)"
+ (%uniform_tupled (v, v1).
+     case v of 0 => case v1 of (0, v4) => ARB | (Suc v5, v4) => (0, v4)
+     | Suc v2 =>
+         case v1 of (0, v8) => ARB
+         | (Suc v9, v8) =>
+             let (res, s') = unif v9 v8
+             in if res < Suc v9 then (res, s')
+                else uniform_tupled (v2, Suc v9, s'))"
 
 lemma uniform_tupled_primitive_def: "uniform_tupled =
 WFREC
- (SOME R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool.
+ (SOME R.
      WF R &
-     (ALL (t::nat) (s::nat => bool) (n::nat) (res::nat) s'::nat => bool.
+     (ALL t s n res s'.
          (res, s') = unif n s & ~ res < Suc n -->
          R (t, Suc n, s') (Suc t, Suc n, s)))
- (%(uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool))
-     (v::nat, v1::nat * (nat => bool)).
-     case v of
-     0 => (%(v3::nat, v4::nat => bool).
-              case v3 of 0 => ARB | Suc (v5::nat) => (0, v4))
-           v1
-     | Suc (v2::nat) =>
-         (%(v7::nat, v8::nat => bool).
-             case v7 of 0 => ARB
-             | Suc (v9::nat) =>
-                 let (res::nat, s'::nat => bool) = unif v9 v8
-                 in if res < Suc v9 then (res, s')
-                    else uniform_tupled (v2, Suc v9, s'))
-          v1)"
+ (%uniform_tupled (v, v1).
+     case v of 0 => case v1 of (0, v4) => ARB | (Suc v5, v4) => (0, v4)
+     | Suc v2 =>
+         case v1 of (0, v8) => ARB
+         | (Suc v9, v8) =>
+             let (res, s') = unif v9 v8
+             in if res < Suc v9 then (res, s')
+                else uniform_tupled (v2, Suc v9, s'))"
   by (import prob_uniform uniform_tupled_primitive_def)
 
 consts
   uniform :: "nat => nat => (nat => bool) => nat * (nat => bool)" 
 
 defs
-  uniform_primdef: "uniform == %(x::nat) (x1::nat) x2::nat => bool. uniform_tupled (x, x1, x2)"
+  uniform_primdef: "uniform == %x x1 x2. uniform_tupled (x, x1, x2)"
 
-lemma uniform_curried_def: "ALL (x::nat) (x1::nat) x2::nat => bool.
-   uniform x x1 x2 = uniform_tupled (x, x1, x2)"
+lemma uniform_curried_def: "uniform x x1 x2 = uniform_tupled (x, x1, x2)"
   by (import prob_uniform uniform_curried_def)
 
-lemma uniform_ind: "ALL P::nat => nat => (nat => bool) => bool.
-   (ALL x::nat. All (P (Suc x) 0)) &
-   All (P 0 0) &
-   (ALL x::nat. All (P 0 (Suc x))) &
-   (ALL (x::nat) (xa::nat) xb::nat => bool.
-       (ALL (xc::nat) xd::nat => bool.
-           (xc, xd) = unif xa xb & ~ xc < Suc xa --> P x (Suc xa) xd) -->
-       P (Suc x) (Suc xa) xb) -->
-   (ALL (x::nat) xa::nat. All (P x xa))"
+lemma uniform_ind: "(ALL x. All (P (Suc x) 0)) &
+All (P 0 0) &
+(ALL x. All (P 0 (Suc x))) &
+(ALL x xa xb.
+    (ALL xc xd.
+        (xc, xd) = unif xa xb & ~ xc < Suc xa --> P x (Suc xa) xd) -->
+    P (Suc x) (Suc xa) xb)
+==> P x xa xb"
   by (import prob_uniform uniform_ind)
 
-lemma uniform_def: "uniform 0 (Suc (n::nat)) (s::nat => bool) = (0, s) &
-uniform (Suc (t::nat)) (Suc n) s =
-(let (xa::nat, x::nat => bool) = unif n s
+lemma uniform_def: "uniform 0 (Suc n) s = (0, s) &
+uniform (Suc t) (Suc n) s =
+(let (xa, x) = unif n s
  in if xa < Suc n then (xa, x) else uniform t (Suc n) x)"
   by (import prob_uniform uniform_def)
 
-lemma SUC_DIV_TWO_ZERO: "ALL n::nat. (Suc n div 2 = 0) = (n = 0)"
+lemma SUC_DIV_TWO_ZERO: "(Suc n div 2 = 0) = (n = 0)"
   by (import prob_uniform SUC_DIV_TWO_ZERO)
 
-lemma UNIF_BOUND_LOWER: "ALL n::nat. n < 2 ^ unif_bound n"
+lemma UNIF_BOUND_LOWER: "n < 2 ^ unif_bound n"
   by (import prob_uniform UNIF_BOUND_LOWER)
 
-lemma UNIF_BOUND_LOWER_SUC: "ALL n::nat. Suc n <= 2 ^ unif_bound n"
+lemma UNIF_BOUND_LOWER_SUC: "Suc n <= 2 ^ unif_bound n"
   by (import prob_uniform UNIF_BOUND_LOWER_SUC)
 
-lemma UNIF_BOUND_UPPER: "ALL n::nat. n ~= 0 --> 2 ^ unif_bound n <= 2 * n"
+lemma UNIF_BOUND_UPPER: "n ~= 0 ==> 2 ^ unif_bound n <= 2 * n"
   by (import prob_uniform UNIF_BOUND_UPPER)
 
-lemma UNIF_BOUND_UPPER_SUC: "ALL n::nat. 2 ^ unif_bound n <= Suc (2 * n)"
+lemma UNIF_BOUND_UPPER_SUC: "2 ^ unif_bound n <= Suc (2 * n)"
   by (import prob_uniform UNIF_BOUND_UPPER_SUC)
 
 lemma UNIF_DEF_MONAD: "unif 0 = UNIT 0 &
-(ALL n::nat.
+(ALL n.
     unif (Suc n) =
     BIND (unif (Suc n div 2))
-     (%m::nat.
-         BIND SDEST (%b::bool. UNIT (if b then 2 * m + 1 else 2 * m))))"
+     (%m. BIND SDEST (%b. UNIT (if b then 2 * m + 1 else 2 * m))))"
   by (import prob_uniform UNIF_DEF_MONAD)
 
-lemma UNIFORM_DEF_MONAD: "(ALL x::nat. uniform 0 (Suc x) = UNIT 0) &
-(ALL (x::nat) xa::nat.
+lemma UNIFORM_DEF_MONAD: "(ALL x. uniform 0 (Suc x) = UNIT 0) &
+(ALL x xa.
     uniform (Suc x) (Suc xa) =
-    BIND (unif xa)
-     (%m::nat. if m < Suc xa then UNIT m else uniform x (Suc xa)))"
+    BIND (unif xa) (%m. if m < Suc xa then UNIT m else uniform x (Suc xa)))"
   by (import prob_uniform UNIFORM_DEF_MONAD)
 
-lemma INDEP_UNIF: "ALL n::nat. indep (unif n)"
+lemma INDEP_UNIF: "indep (unif n)"
   by (import prob_uniform INDEP_UNIF)
 
-lemma INDEP_UNIFORM: "ALL (t::nat) n::nat. indep (uniform t (Suc n))"
+lemma INDEP_UNIFORM: "indep (uniform t (Suc n))"
   by (import prob_uniform INDEP_UNIFORM)
 
-lemma PROB_UNIF: "ALL (n::nat) k::nat.
-   prob (%s::nat => bool. fst (unif n s) = k) =
-   (if k < 2 ^ unif_bound n then (1 / 2) ^ unif_bound n else 0)"
+lemma PROB_UNIF: "prob (%s. fst (unif n s) = k) =
+(if k < 2 ^ unif_bound n then (1 / 2) ^ unif_bound n else 0)"
   by (import prob_uniform PROB_UNIF)
 
-lemma UNIF_RANGE: "ALL (n::nat) s::nat => bool. fst (unif n s) < 2 ^ unif_bound n"
+lemma UNIF_RANGE: "fst (unif n s) < 2 ^ unif_bound n"
   by (import prob_uniform UNIF_RANGE)
 
-lemma PROB_UNIF_PAIR: "ALL (n::nat) (k::nat) k'::nat.
-   (prob (%s::nat => bool. fst (unif n s) = k) =
-    prob (%s::nat => bool. fst (unif n s) = k')) =
-   ((k < 2 ^ unif_bound n) = (k' < 2 ^ unif_bound n))"
+lemma PROB_UNIF_PAIR: "(prob (%s. fst (unif n s) = k) = prob (%s. fst (unif n s) = k')) =
+((k < 2 ^ unif_bound n) = (k' < 2 ^ unif_bound n))"
   by (import prob_uniform PROB_UNIF_PAIR)
 
-lemma PROB_UNIF_BOUND: "ALL (n::nat) k::nat.
-   k <= 2 ^ unif_bound n -->
-   prob (%s::nat => bool. fst (unif n s) < k) =
-   real k * (1 / 2) ^ unif_bound n"
+lemma PROB_UNIF_BOUND: "k <= 2 ^ unif_bound n
+==> prob (%s. fst (unif n s) < k) = real k * (1 / 2) ^ unif_bound n"
   by (import prob_uniform PROB_UNIF_BOUND)
 
-lemma PROB_UNIF_GOOD: "ALL n::nat. 1 / 2 <= prob (%s::nat => bool. fst (unif n s) < Suc n)"
+lemma PROB_UNIF_GOOD: "1 / 2 <= prob (%s. fst (unif n s) < Suc n)"
   by (import prob_uniform PROB_UNIF_GOOD)
 
-lemma UNIFORM_RANGE: "ALL (t::nat) (n::nat) s::nat => bool. fst (uniform t (Suc n) s) < Suc n"
+lemma UNIFORM_RANGE: "fst (uniform t (Suc n) s) < Suc n"
   by (import prob_uniform UNIFORM_RANGE)
 
-lemma PROB_UNIFORM_LOWER_BOUND: "(All::(real => bool) => bool)
- (%b::real.
-     (op -->::bool => bool => bool)
-      ((All::(nat => bool) => bool)
-        (%k::nat.
-            (op -->::bool => bool => bool)
-             ((op <::nat => nat => bool) k ((Suc::nat => nat) (n::nat)))
-             ((op <::real => real => bool)
-               ((prob::((nat => bool) => bool) => real)
-                 (%s::nat => bool.
-                     (op =::nat => nat => bool)
-                      ((fst::nat * (nat => bool) => nat)
-                        ((uniform::nat
-                                   => nat
-=> (nat => bool) => nat * (nat => bool))
-                          (t::nat) ((Suc::nat => nat) n) s))
-                      k))
-               b)))
-      ((All::(nat => bool) => bool)
-        (%m::nat.
-            (op -->::bool => bool => bool)
-             ((op <::nat => nat => bool) m ((Suc::nat => nat) n))
-             ((op <::real => real => bool)
-               ((prob::((nat => bool) => bool) => real)
-                 (%s::nat => bool.
-                     (op <::nat => nat => bool)
-                      ((fst::nat * (nat => bool) => nat)
-                        ((uniform::nat
-                                   => nat
-=> (nat => bool) => nat * (nat => bool))
-                          t ((Suc::nat => nat) n) s))
-                      ((Suc::nat => nat) m)))
-               ((op *::real => real => real)
-                 ((real::nat => real) ((Suc::nat => nat) m)) b)))))"
+lemma PROB_UNIFORM_LOWER_BOUND: "[| !!k. k < Suc n ==> prob (%s. fst (uniform t (Suc n) s) = k) < b;
+   m < Suc n |]
+==> prob (%s. fst (uniform t (Suc n) s) < Suc m) < real (Suc m) * b"
   by (import prob_uniform PROB_UNIFORM_LOWER_BOUND)
 
-lemma PROB_UNIFORM_UPPER_BOUND: "(All::(real => bool) => bool)
- (%b::real.
-     (op -->::bool => bool => bool)
-      ((All::(nat => bool) => bool)
-        (%k::nat.
-            (op -->::bool => bool => bool)
-             ((op <::nat => nat => bool) k ((Suc::nat => nat) (n::nat)))
-             ((op <::real => real => bool) b
-               ((prob::((nat => bool) => bool) => real)
-                 (%s::nat => bool.
-                     (op =::nat => nat => bool)
-                      ((fst::nat * (nat => bool) => nat)
-                        ((uniform::nat
-                                   => nat
-=> (nat => bool) => nat * (nat => bool))
-                          (t::nat) ((Suc::nat => nat) n) s))
-                      k)))))
-      ((All::(nat => bool) => bool)
-        (%m::nat.
-            (op -->::bool => bool => bool)
-             ((op <::nat => nat => bool) m ((Suc::nat => nat) n))
-             ((op <::real => real => bool)
-               ((op *::real => real => real)
-                 ((real::nat => real) ((Suc::nat => nat) m)) b)
-               ((prob::((nat => bool) => bool) => real)
-                 (%s::nat => bool.
-                     (op <::nat => nat => bool)
-                      ((fst::nat * (nat => bool) => nat)
-                        ((uniform::nat
-                                   => nat
-=> (nat => bool) => nat * (nat => bool))
-                          t ((Suc::nat => nat) n) s))
-                      ((Suc::nat => nat) m)))))))"
+lemma PROB_UNIFORM_UPPER_BOUND: "[| !!k. k < Suc n ==> b < prob (%s. fst (uniform t (Suc n) s) = k);
+   m < Suc n |]
+==> real (Suc m) * b < prob (%s. fst (uniform t (Suc n) s) < Suc m)"
   by (import prob_uniform PROB_UNIFORM_UPPER_BOUND)
 
-lemma PROB_UNIFORM_PAIR_SUC: "ALL (t::nat) (n::nat) (k::nat) k'::nat.
-   k < Suc n & k' < Suc n -->
-   abs (prob (%s::nat => bool. fst (uniform t (Suc n) s) = k) -
-        prob (%s::nat => bool. fst (uniform t (Suc n) s) = k'))
-   <= (1 / 2) ^ t"
+lemma PROB_UNIFORM_PAIR_SUC: "k < Suc n & k' < Suc n
+==> abs (prob (%s. fst (uniform t (Suc n) s) = k) -
+         prob (%s. fst (uniform t (Suc n) s) = k'))
+    <= (1 / 2) ^ t"
   by (import prob_uniform PROB_UNIFORM_PAIR_SUC)
 
-lemma PROB_UNIFORM_SUC: "ALL (t::nat) (n::nat) k::nat.
-   k < Suc n -->
-   abs (prob (%s::nat => bool. fst (uniform t (Suc n) s) = k) -
-        1 / real (Suc n))
-   <= (1 / 2) ^ t"
+lemma PROB_UNIFORM_SUC: "k < Suc n
+==> abs (prob (%s. fst (uniform t (Suc n) s) = k) - 1 / real (Suc n))
+    <= (1 / 2) ^ t"
   by (import prob_uniform PROB_UNIFORM_SUC)
 
-lemma PROB_UNIFORM: "ALL (t::nat) (n::nat) k::nat.
-   k < n -->
-   abs (prob (%s::nat => bool. fst (uniform t n s) = k) - 1 / real n)
-   <= (1 / 2) ^ t"
+lemma PROB_UNIFORM: "k < n
+==> abs (prob (%s. fst (uniform t n s) = k) - 1 / real n) <= (1 / 2) ^ t"
   by (import prob_uniform PROB_UNIFORM)
 
 ;end_setup
--- a/src/HOL/Import/HOL/HOL4Real.thy	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/HOL4Real.thy	Tue Sep 06 22:41:35 2011 -0700
@@ -4,858 +4,702 @@
 
 ;setup_theory realax
 
-lemma HREAL_RDISTRIB: "ALL (x::hreal) (y::hreal) z::hreal.
-   hreal_mul (hreal_add x y) z = hreal_add (hreal_mul x z) (hreal_mul y z)"
+lemma HREAL_RDISTRIB: "hreal_mul (hreal_add x y) z = hreal_add (hreal_mul x z) (hreal_mul y z)"
   by (import realax HREAL_RDISTRIB)
 
-lemma HREAL_EQ_ADDL: "ALL (x::hreal) y::hreal. x ~= hreal_add x y"
+lemma HREAL_EQ_ADDL: "x ~= hreal_add x y"
   by (import realax HREAL_EQ_ADDL)
 
-lemma HREAL_EQ_LADD: "ALL (x::hreal) (y::hreal) z::hreal.
-   (hreal_add x y = hreal_add x z) = (y = z)"
+lemma HREAL_EQ_LADD: "(hreal_add x y = hreal_add x z) = (y = z)"
   by (import realax HREAL_EQ_LADD)
 
-lemma HREAL_LT_REFL: "ALL x::hreal. ~ hreal_lt x x"
+lemma HREAL_LT_REFL: "~ hreal_lt x x"
   by (import realax HREAL_LT_REFL)
 
-lemma HREAL_LT_ADDL: "ALL (x::hreal) y::hreal. hreal_lt x (hreal_add x y)"
+lemma HREAL_LT_ADDL: "hreal_lt x (hreal_add x y)"
   by (import realax HREAL_LT_ADDL)
 
-lemma HREAL_LT_NE: "ALL (x::hreal) y::hreal. hreal_lt x y --> x ~= y"
+lemma HREAL_LT_NE: "hreal_lt x y ==> x ~= y"
   by (import realax HREAL_LT_NE)
 
-lemma HREAL_LT_ADDR: "ALL (x::hreal) y::hreal. ~ hreal_lt (hreal_add x y) x"
+lemma HREAL_LT_ADDR: "~ hreal_lt (hreal_add x y) x"
   by (import realax HREAL_LT_ADDR)
 
-lemma HREAL_LT_GT: "ALL (x::hreal) y::hreal. hreal_lt x y --> ~ hreal_lt y x"
+lemma HREAL_LT_GT: "hreal_lt x y ==> ~ hreal_lt y x"
   by (import realax HREAL_LT_GT)
 
-lemma HREAL_LT_ADD2: "ALL (x1::hreal) (x2::hreal) (y1::hreal) y2::hreal.
-   hreal_lt x1 y1 & hreal_lt x2 y2 -->
-   hreal_lt (hreal_add x1 x2) (hreal_add y1 y2)"
+lemma HREAL_LT_ADD2: "hreal_lt x1 y1 & hreal_lt x2 y2
+==> hreal_lt (hreal_add x1 x2) (hreal_add y1 y2)"
   by (import realax HREAL_LT_ADD2)
 
-lemma HREAL_LT_LADD: "ALL (x::hreal) (y::hreal) z::hreal.
-   hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z"
+lemma HREAL_LT_LADD: "hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z"
   by (import realax HREAL_LT_LADD)
 
-definition treal_0 :: "hreal * hreal" where 
+definition
+  treal_0 :: "hreal * hreal"  where
   "treal_0 == (hreal_1, hreal_1)"
 
 lemma treal_0: "treal_0 = (hreal_1, hreal_1)"
   by (import realax treal_0)
 
-definition treal_1 :: "hreal * hreal" where 
+definition
+  treal_1 :: "hreal * hreal"  where
   "treal_1 == (hreal_add hreal_1 hreal_1, hreal_1)"
 
 lemma treal_1: "treal_1 = (hreal_add hreal_1 hreal_1, hreal_1)"
   by (import realax treal_1)
 
-definition treal_neg :: "hreal * hreal => hreal * hreal" where 
-  "treal_neg == %(x::hreal, y::hreal). (y, x)"
-
-lemma treal_neg: "ALL (x::hreal) y::hreal. treal_neg (x, y) = (y, x)"
+definition
+  treal_neg :: "hreal * hreal => hreal * hreal"  where
+  "treal_neg == %(x, y). (y, x)"
+
+lemma treal_neg: "treal_neg (x, y) = (y, x)"
   by (import realax treal_neg)
 
-definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
-  "treal_add ==
-%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
-   (hreal_add x1 x2, hreal_add y1 y2)"
-
-lemma treal_add: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal.
-   treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)"
+definition
+  treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal"  where
+  "treal_add == %(x1, y1) (x2, y2). (hreal_add x1 x2, hreal_add y1 y2)"
+
+lemma treal_add: "treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)"
   by (import realax treal_add)
 
-definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where 
+definition
+  treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal"  where
   "treal_mul ==
-%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
+%(x1, y1) (x2, y2).
    (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2),
     hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))"
 
-lemma treal_mul: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal.
-   treal_mul (x1, y1) (x2, y2) =
-   (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2),
-    hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))"
+lemma treal_mul: "treal_mul (x1, y1) (x2, y2) =
+(hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2),
+ hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))"
   by (import realax treal_mul)
 
-definition treal_lt :: "hreal * hreal => hreal * hreal => bool" where 
-  "treal_lt ==
-%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
-   hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
-
-lemma treal_lt: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal.
-   treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
+definition
+  treal_lt :: "hreal * hreal => hreal * hreal => bool"  where
+  "treal_lt == %(x1, y1) (x2, y2). hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
+
+lemma treal_lt: "treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
   by (import realax treal_lt)
 
-definition treal_inv :: "hreal * hreal => hreal * hreal" where 
+definition
+  treal_inv :: "hreal * hreal => hreal * hreal"  where
   "treal_inv ==
-%(x::hreal, y::hreal).
+%(x, y).
    if x = y then treal_0
    else if hreal_lt y x
         then (hreal_add (hreal_inv (hreal_sub x y)) hreal_1, hreal_1)
         else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1)"
 
-lemma treal_inv: "ALL (x::hreal) y::hreal.
-   treal_inv (x, y) =
-   (if x = y then treal_0
-    else if hreal_lt y x
-         then (hreal_add (hreal_inv (hreal_sub x y)) hreal_1, hreal_1)
-         else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1))"
+lemma treal_inv: "treal_inv (x, y) =
+(if x = y then treal_0
+ else if hreal_lt y x
+      then (hreal_add (hreal_inv (hreal_sub x y)) hreal_1, hreal_1)
+      else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1))"
   by (import realax treal_inv)
 
-definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where 
-  "treal_eq ==
-%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal).
-   hreal_add x1 y2 = hreal_add x2 y1"
-
-lemma treal_eq: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal.
-   treal_eq (x1, y1) (x2, y2) = (hreal_add x1 y2 = hreal_add x2 y1)"
+definition
+  treal_eq :: "hreal * hreal => hreal * hreal => bool"  where
+  "treal_eq == %(x1, y1) (x2, y2). hreal_add x1 y2 = hreal_add x2 y1"
+
+lemma treal_eq: "treal_eq (x1, y1) (x2, y2) = (hreal_add x1 y2 = hreal_add x2 y1)"
   by (import realax treal_eq)
 
-lemma TREAL_EQ_REFL: "ALL x::hreal * hreal. treal_eq x x"
+lemma TREAL_EQ_REFL: "treal_eq x x"
   by (import realax TREAL_EQ_REFL)
 
-lemma TREAL_EQ_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y = treal_eq y x"
+lemma TREAL_EQ_SYM: "treal_eq x y = treal_eq y x"
   by (import realax TREAL_EQ_SYM)
 
-lemma TREAL_EQ_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_eq x y & treal_eq y z --> treal_eq x z"
+lemma TREAL_EQ_TRANS: "treal_eq x y & treal_eq y z ==> treal_eq x z"
   by (import realax TREAL_EQ_TRANS)
 
-lemma TREAL_EQ_EQUIV: "ALL (p::hreal * hreal) q::hreal * hreal.
-   treal_eq p q = (treal_eq p = treal_eq q)"
+lemma TREAL_EQ_EQUIV: "treal_eq p q = (treal_eq p = treal_eq q)"
   by (import realax TREAL_EQ_EQUIV)
 
-lemma TREAL_EQ_AP: "ALL (p::hreal * hreal) q::hreal * hreal. p = q --> treal_eq p q"
+lemma TREAL_EQ_AP: "p = q ==> treal_eq p q"
   by (import realax TREAL_EQ_AP)
 
 lemma TREAL_10: "~ treal_eq treal_1 treal_0"
   by (import realax TREAL_10)
 
-lemma TREAL_ADD_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_add x y = treal_add y x"
+lemma TREAL_ADD_SYM: "treal_add x y = treal_add y x"
   by (import realax TREAL_ADD_SYM)
 
-lemma TREAL_MUL_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_mul x y = treal_mul y x"
+lemma TREAL_MUL_SYM: "treal_mul x y = treal_mul y x"
   by (import realax TREAL_MUL_SYM)
 
-lemma TREAL_ADD_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_add x (treal_add y z) = treal_add (treal_add x y) z"
+lemma TREAL_ADD_ASSOC: "treal_add x (treal_add y z) = treal_add (treal_add x y) z"
   by (import realax TREAL_ADD_ASSOC)
 
-lemma TREAL_MUL_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_mul x (treal_mul y z) = treal_mul (treal_mul x y) z"
+lemma TREAL_MUL_ASSOC: "treal_mul x (treal_mul y z) = treal_mul (treal_mul x y) z"
   by (import realax TREAL_MUL_ASSOC)
 
-lemma TREAL_LDISTRIB: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_mul x (treal_add y z) = treal_add (treal_mul x y) (treal_mul x z)"
+lemma TREAL_LDISTRIB: "treal_mul x (treal_add y z) = treal_add (treal_mul x y) (treal_mul x z)"
   by (import realax TREAL_LDISTRIB)
 
-lemma TREAL_ADD_LID: "ALL x::hreal * hreal. treal_eq (treal_add treal_0 x) x"
+lemma TREAL_ADD_LID: "treal_eq (treal_add treal_0 x) x"
   by (import realax TREAL_ADD_LID)
 
-lemma TREAL_MUL_LID: "ALL x::hreal * hreal. treal_eq (treal_mul treal_1 x) x"
+lemma TREAL_MUL_LID: "treal_eq (treal_mul treal_1 x) x"
   by (import realax TREAL_MUL_LID)
 
-lemma TREAL_ADD_LINV: "ALL x::hreal * hreal. treal_eq (treal_add (treal_neg x) x) treal_0"
+lemma TREAL_ADD_LINV: "treal_eq (treal_add (treal_neg x) x) treal_0"
   by (import realax TREAL_ADD_LINV)
 
 lemma TREAL_INV_0: "treal_eq (treal_inv treal_0) treal_0"
   by (import realax TREAL_INV_0)
 
-lemma TREAL_MUL_LINV: "ALL x::hreal * hreal.
-   ~ treal_eq x treal_0 --> treal_eq (treal_mul (treal_inv x) x) treal_1"
+lemma TREAL_MUL_LINV: "~ treal_eq x treal_0 ==> treal_eq (treal_mul (treal_inv x) x) treal_1"
   by (import realax TREAL_MUL_LINV)
 
-lemma TREAL_LT_TOTAL: "ALL (x::hreal * hreal) y::hreal * hreal.
-   treal_eq x y | treal_lt x y | treal_lt y x"
+lemma TREAL_LT_TOTAL: "treal_eq x y | treal_lt x y | treal_lt y x"
   by (import realax TREAL_LT_TOTAL)
 
-lemma TREAL_LT_REFL: "ALL x::hreal * hreal. ~ treal_lt x x"
+lemma TREAL_LT_REFL: "~ treal_lt x x"
   by (import realax TREAL_LT_REFL)
 
-lemma TREAL_LT_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_lt x y & treal_lt y z --> treal_lt x z"
+lemma TREAL_LT_TRANS: "treal_lt x y & treal_lt y z ==> treal_lt x z"
   by (import realax TREAL_LT_TRANS)
 
-lemma TREAL_LT_ADD: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal.
-   treal_lt y z --> treal_lt (treal_add x y) (treal_add x z)"
+lemma TREAL_LT_ADD: "treal_lt y z ==> treal_lt (treal_add x y) (treal_add x z)"
   by (import realax TREAL_LT_ADD)
 
-lemma TREAL_LT_MUL: "ALL (x::hreal * hreal) y::hreal * hreal.
-   treal_lt treal_0 x & treal_lt treal_0 y -->
-   treal_lt treal_0 (treal_mul x y)"
+lemma TREAL_LT_MUL: "treal_lt treal_0 x & treal_lt treal_0 y ==> treal_lt treal_0 (treal_mul x y)"
   by (import realax TREAL_LT_MUL)
 
-definition treal_of_hreal :: "hreal => hreal * hreal" where 
-  "treal_of_hreal == %x::hreal. (hreal_add x hreal_1, hreal_1)"
-
-lemma treal_of_hreal: "ALL x::hreal. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)"
+definition
+  treal_of_hreal :: "hreal => hreal * hreal"  where
+  "treal_of_hreal == %x. (hreal_add x hreal_1, hreal_1)"
+
+lemma treal_of_hreal: "treal_of_hreal x = (hreal_add x hreal_1, hreal_1)"
   by (import realax treal_of_hreal)
 
-definition hreal_of_treal :: "hreal * hreal => hreal" where 
-  "hreal_of_treal == %(x::hreal, y::hreal). SOME d::hreal. x = hreal_add y d"
-
-lemma hreal_of_treal: "ALL (x::hreal) y::hreal.
-   hreal_of_treal (x, y) = (SOME d::hreal. x = hreal_add y d)"
+definition
+  hreal_of_treal :: "hreal * hreal => hreal"  where
+  "hreal_of_treal == %(x, y). SOME d. x = hreal_add y d"
+
+lemma hreal_of_treal: "hreal_of_treal (x, y) = (SOME d. x = hreal_add y d)"
   by (import realax hreal_of_treal)
 
-lemma TREAL_BIJ: "(ALL h::hreal. hreal_of_treal (treal_of_hreal h) = h) &
-(ALL r::hreal * hreal.
-    treal_lt treal_0 r = treal_eq (treal_of_hreal (hreal_of_treal r)) r)"
+lemma TREAL_BIJ: "(ALL h. hreal_of_treal (treal_of_hreal h) = h) &
+(ALL r. treal_lt treal_0 r = treal_eq (treal_of_hreal (hreal_of_treal r)) r)"
   by (import realax TREAL_BIJ)
 
-lemma TREAL_ISO: "ALL (h::hreal) i::hreal.
-   hreal_lt h i --> treal_lt (treal_of_hreal h) (treal_of_hreal i)"
+lemma TREAL_ISO: "hreal_lt h i ==> treal_lt (treal_of_hreal h) (treal_of_hreal i)"
   by (import realax TREAL_ISO)
 
-lemma TREAL_BIJ_WELLDEF: "ALL (h::hreal * hreal) i::hreal * hreal.
-   treal_eq h i --> hreal_of_treal h = hreal_of_treal i"
+lemma TREAL_BIJ_WELLDEF: "treal_eq h i ==> hreal_of_treal h = hreal_of_treal i"
   by (import realax TREAL_BIJ_WELLDEF)
 
-lemma TREAL_NEG_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal.
-   treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)"
+lemma TREAL_NEG_WELLDEF: "treal_eq x1 x2 ==> treal_eq (treal_neg x1) (treal_neg x2)"
   by (import realax TREAL_NEG_WELLDEF)
 
-lemma TREAL_ADD_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal.
-   treal_eq x1 x2 --> treal_eq (treal_add x1 y) (treal_add x2 y)"
+lemma TREAL_ADD_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_add x1 y) (treal_add x2 y)"
   by (import realax TREAL_ADD_WELLDEFR)
 
-lemma TREAL_ADD_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
-   y2::hreal * hreal.
-   treal_eq x1 x2 & treal_eq y1 y2 -->
-   treal_eq (treal_add x1 y1) (treal_add x2 y2)"
+lemma TREAL_ADD_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
+==> treal_eq (treal_add x1 y1) (treal_add x2 y2)"
   by (import realax TREAL_ADD_WELLDEF)
 
-lemma TREAL_MUL_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal.
-   treal_eq x1 x2 --> treal_eq (treal_mul x1 y) (treal_mul x2 y)"
+lemma TREAL_MUL_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_mul x1 y) (treal_mul x2 y)"
   by (import realax TREAL_MUL_WELLDEFR)
 
-lemma TREAL_MUL_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
-   y2::hreal * hreal.
-   treal_eq x1 x2 & treal_eq y1 y2 -->
-   treal_eq (treal_mul x1 y1) (treal_mul x2 y2)"
+lemma TREAL_MUL_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
+==> treal_eq (treal_mul x1 y1) (treal_mul x2 y2)"
   by (import realax TREAL_MUL_WELLDEF)
 
-lemma TREAL_LT_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal.
-   treal_eq x1 x2 --> treal_lt x1 y = treal_lt x2 y"
+lemma TREAL_LT_WELLDEFR: "treal_eq x1 x2 ==> treal_lt x1 y = treal_lt x2 y"
   by (import realax TREAL_LT_WELLDEFR)
 
-lemma TREAL_LT_WELLDEFL: "ALL (x::hreal * hreal) (y1::hreal * hreal) y2::hreal * hreal.
-   treal_eq y1 y2 --> treal_lt x y1 = treal_lt x y2"
+lemma TREAL_LT_WELLDEFL: "treal_eq y1 y2 ==> treal_lt x y1 = treal_lt x y2"
   by (import realax TREAL_LT_WELLDEFL)
 
-lemma TREAL_LT_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal)
-   y2::hreal * hreal.
-   treal_eq x1 x2 & treal_eq y1 y2 --> treal_lt x1 y1 = treal_lt x2 y2"
+lemma TREAL_LT_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2 ==> treal_lt x1 y1 = treal_lt x2 y2"
   by (import realax TREAL_LT_WELLDEF)
 
-lemma TREAL_INV_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal.
-   treal_eq x1 x2 --> treal_eq (treal_inv x1) (treal_inv x2)"
+lemma TREAL_INV_WELLDEF: "treal_eq x1 x2 ==> treal_eq (treal_inv x1) (treal_inv x2)"
   by (import realax TREAL_INV_WELLDEF)
 
 ;end_setup
 
 ;setup_theory real
 
-lemma REAL_0: "(op =::real => real => bool) (0::real) (0::real)"
+lemma REAL_0: "(0::real) = (0::real)"
   by (import real REAL_0)
 
-lemma REAL_1: "(op =::real => real => bool) (1::real) (1::real)"
+lemma REAL_1: "(1::real) = (1::real)"
   by (import real REAL_1)
 
-lemma REAL_ADD_LID_UNIQ: "ALL (x::real) y::real. (x + y = y) = (x = 0)"
+lemma REAL_ADD_LID_UNIQ: "((x::real) + (y::real) = y) = (x = (0::real))"
   by (import real REAL_ADD_LID_UNIQ)
 
-lemma REAL_ADD_RID_UNIQ: "ALL (x::real) y::real. (x + y = x) = (y = 0)"
+lemma REAL_ADD_RID_UNIQ: "((x::real) + (y::real) = x) = (y = (0::real))"
   by (import real REAL_ADD_RID_UNIQ)
 
-lemma REAL_LNEG_UNIQ: "ALL (x::real) y::real. (x + y = 0) = (x = - y)"
-  by (import real REAL_LNEG_UNIQ)
-
-lemma REAL_LT_ANTISYM: "ALL (x::real) y::real. ~ (x < y & y < x)"
+lemma REAL_LT_ANTISYM: "~ ((x::real) < (y::real) & y < x)"
   by (import real REAL_LT_ANTISYM)
 
-lemma REAL_LTE_TOTAL: "ALL (x::real) y::real. x < y | y <= x"
+lemma REAL_LTE_TOTAL: "(x::real) < (y::real) | y <= x"
   by (import real REAL_LTE_TOTAL)
 
-lemma REAL_LET_ANTISYM: "ALL (x::real) y::real. ~ (x < y & y <= x)"
+lemma REAL_LET_ANTISYM: "~ ((x::real) < (y::real) & y <= x)"
   by (import real REAL_LET_ANTISYM)
 
-lemma REAL_LTE_ANTSYM: "ALL (x::real) y::real. ~ (x <= y & y < x)"
+lemma REAL_LTE_ANTSYM: "~ ((x::real) <= (y::real) & y < x)"
   by (import real REAL_LTE_ANTSYM)
 
-lemma REAL_LT_NEGTOTAL: "ALL x::real. x = 0 | 0 < x | 0 < - x"
+lemma REAL_LT_NEGTOTAL: "(x::real) = (0::real) | (0::real) < x | (0::real) < - x"
   by (import real REAL_LT_NEGTOTAL)
 
-lemma REAL_LE_NEGTOTAL: "ALL x::real. 0 <= x | 0 <= - x"
+lemma REAL_LE_NEGTOTAL: "(0::real) <= (x::real) | (0::real) <= - x"
   by (import real REAL_LE_NEGTOTAL)
 
-lemma REAL_LT_ADDNEG: "ALL (x::real) (y::real) z::real. (y < x + - z) = (y + z < x)"
+lemma REAL_LT_ADDNEG: "((y::real) < (x::real) + - (z::real)) = (y + z < x)"
   by (import real REAL_LT_ADDNEG)
 
-lemma REAL_LT_ADDNEG2: "ALL (x::real) (y::real) z::real. (x + - y < z) = (x < z + y)"
+lemma REAL_LT_ADDNEG2: "((x::real) + - (y::real) < (z::real)) = (x < z + y)"
   by (import real REAL_LT_ADDNEG2)
 
-lemma REAL_LT_ADD1: "ALL (x::real) y::real. x <= y --> x < y + 1"
+lemma REAL_LT_ADD1: "(x::real) <= (y::real) ==> x < y + (1::real)"
   by (import real REAL_LT_ADD1)
 
-lemma REAL_SUB_ADD2: "ALL (x::real) y::real. y + (x - y) = x"
+lemma REAL_SUB_ADD2: "(y::real) + ((x::real) - y) = x"
   by (import real REAL_SUB_ADD2)
 
-lemma REAL_SUB_LT: "ALL (x::real) y::real. (0 < x - y) = (y < x)"
+lemma REAL_SUB_LT: "((0::real) < (x::real) - (y::real)) = (y < x)"
   by (import real REAL_SUB_LT)
 
-lemma REAL_SUB_LE: "ALL (x::real) y::real. (0 <= x - y) = (y <= x)"
+lemma REAL_SUB_LE: "((0::real) <= (x::real) - (y::real)) = (y <= x)"
   by (import real REAL_SUB_LE)
 
-lemma REAL_ADD_SUB: "ALL (x::real) y::real. x + y - x = y"
+lemma REAL_ADD_SUB: "(x::real) + (y::real) - x = y"
   by (import real REAL_ADD_SUB)
 
-lemma REAL_NEG_EQ: "ALL (x::real) y::real. (- x = y) = (x = - y)"
+lemma REAL_NEG_EQ: "(- (x::real) = (y::real)) = (x = - y)"
   by (import real REAL_NEG_EQ)
 
-lemma REAL_NEG_MINUS1: "ALL x::real. - x = - 1 * x"
-  by (import real REAL_NEG_MINUS1)
-
-lemma REAL_LT_LMUL_0: "ALL (x::real) y::real. 0 < x --> (0 < x * y) = (0 < y)"
+lemma REAL_LT_LMUL_0: "(0::real) < (x::real) ==> ((0::real) < x * (y::real)) = ((0::real) < y)"
   by (import real REAL_LT_LMUL_0)
 
-lemma REAL_LT_RMUL_0: "ALL (x::real) y::real. 0 < y --> (0 < x * y) = (0 < x)"
+lemma REAL_LT_RMUL_0: "(0::real) < (y::real) ==> ((0::real) < (x::real) * y) = ((0::real) < x)"
   by (import real REAL_LT_RMUL_0)
 
-lemma REAL_LT_LMUL: "ALL (x::real) (y::real) z::real. 0 < x --> (x * y < x * z) = (y < z)"
-  by (import real REAL_LT_LMUL)
-
-lemma REAL_LINV_UNIQ: "ALL (x::real) y::real. x * y = 1 --> x = inverse y"
+lemma REAL_LINV_UNIQ: "(x::real) * (y::real) = (1::real) ==> x = inverse y"
   by (import real REAL_LINV_UNIQ)
 
-lemma REAL_LE_INV: "(All::(real => bool) => bool)
- (%x::real.
-     (op -->::bool => bool => bool)
-      ((op <=::real => real => bool) (0::real) x)
-      ((op <=::real => real => bool) (0::real) ((inverse::real => real) x)))"
+lemma REAL_LE_INV: "(0::real) <= (x::real) ==> (0::real) <= inverse x"
   by (import real REAL_LE_INV)
 
-lemma REAL_LE_ADDR: "ALL (x::real) y::real. (x <= x + y) = (0 <= y)"
+lemma REAL_LE_ADDR: "((x::real) <= x + (y::real)) = ((0::real) <= y)"
   by (import real REAL_LE_ADDR)
 
-lemma REAL_LE_ADDL: "ALL (x::real) y::real. (y <= x + y) = (0 <= x)"
+lemma REAL_LE_ADDL: "((y::real) <= (x::real) + y) = ((0::real) <= x)"
   by (import real REAL_LE_ADDL)
 
-lemma REAL_LT_ADDR: "ALL (x::real) y::real. (x < x + y) = (0 < y)"
+lemma REAL_LT_ADDR: "((x::real) < x + (y::real)) = ((0::real) < y)"
   by (import real REAL_LT_ADDR)
 
-lemma REAL_LT_ADDL: "ALL (x::real) y::real. (y < x + y) = (0 < x)"
+lemma REAL_LT_ADDL: "((y::real) < (x::real) + y) = ((0::real) < x)"
   by (import real REAL_LT_ADDL)
 
-lemma REAL_LT_NZ: "ALL n::nat. (real n ~= 0) = (0 < real n)"
+lemma REAL_LT_NZ: "(real (n::nat) ~= (0::real)) = ((0::real) < real n)"
   by (import real REAL_LT_NZ)
 
-lemma REAL_NZ_IMP_LT: "ALL n::nat. n ~= 0 --> 0 < real n"
+lemma REAL_NZ_IMP_LT: "(n::nat) ~= (0::nat) ==> (0::real) < real n"
   by (import real REAL_NZ_IMP_LT)
 
-lemma REAL_LT_RDIV_0: "ALL (y::real) z::real. 0 < z --> (0 < y / z) = (0 < y)"
+lemma REAL_LT_RDIV_0: "(0::real) < (z::real) ==> ((0::real) < (y::real) / z) = ((0::real) < y)"
   by (import real REAL_LT_RDIV_0)
 
-lemma REAL_LT_RDIV: "ALL (x::real) (y::real) z::real. 0 < z --> (x / z < y / z) = (x < y)"
+lemma REAL_LT_RDIV: "(0::real) < (z::real) ==> ((x::real) / z < (y::real) / z) = (x < y)"
   by (import real REAL_LT_RDIV)
 
-lemma REAL_LT_FRACTION_0: "ALL (n::nat) d::real. n ~= 0 --> (0 < d / real n) = (0 < d)"
+lemma REAL_LT_FRACTION_0: "(n::nat) ~= (0::nat) ==> ((0::real) < (d::real) / real n) = ((0::real) < d)"
   by (import real REAL_LT_FRACTION_0)
 
-lemma REAL_LT_MULTIPLE: "ALL (x::nat) xa::real. 1 < x --> (xa < real x * xa) = (0 < xa)"
+lemma REAL_LT_MULTIPLE: "(1::nat) < (x::nat) ==> ((xa::real) < real x * xa) = ((0::real) < xa)"
   by (import real REAL_LT_MULTIPLE)
 
-lemma REAL_LT_FRACTION: "ALL (n::nat) d::real. 1 < n --> (d / real n < d) = (0 < d)"
+lemma REAL_LT_FRACTION: "(1::nat) < (n::nat) ==> ((d::real) / real n < d) = ((0::real) < d)"
   by (import real REAL_LT_FRACTION)
 
-lemma REAL_LT_HALF2: "ALL d::real. (d / 2 < d) = (0 < d)"
+lemma REAL_LT_HALF2: "((d::real) / (2::real) < d) = ((0::real) < d)"
   by (import real REAL_LT_HALF2)
 
-lemma REAL_DIV_LMUL: "ALL (x::real) y::real. y ~= 0 --> y * (x / y) = x"
+lemma REAL_DIV_LMUL: "(y::real) ~= (0::real) ==> y * ((x::real) / y) = x"
   by (import real REAL_DIV_LMUL)
 
-lemma REAL_DIV_RMUL: "ALL (x::real) y::real. y ~= 0 --> x / y * y = x"
+lemma REAL_DIV_RMUL: "(y::real) ~= (0::real) ==> (x::real) / y * y = x"
   by (import real REAL_DIV_RMUL)
 
-lemma REAL_DOWN: "(All::(real => bool) => bool)
- (%x::real.
-     (op -->::bool => bool => bool)
-      ((op <::real => real => bool) (0::real) x)
-      ((Ex::(real => bool) => bool)
-        (%xa::real.
-            (op &::bool => bool => bool)
-             ((op <::real => real => bool) (0::real) xa)
-             ((op <::real => real => bool) xa x))))"
+lemma REAL_DOWN: "(0::real) < (x::real) ==> EX xa>0::real. xa < x"
   by (import real REAL_DOWN)
 
-lemma REAL_SUB_SUB: "ALL (x::real) y::real. x - y - x = - y"
+lemma REAL_SUB_SUB: "(x::real) - (y::real) - x = - y"
   by (import real REAL_SUB_SUB)
 
-lemma REAL_ADD2_SUB2: "ALL (a::real) (b::real) (c::real) d::real. a + b - (c + d) = a - c + (b - d)"
-  by (import real REAL_ADD2_SUB2)
-
-lemma REAL_SUB_LNEG: "ALL (x::real) y::real. - x - y = - (x + y)"
+lemma REAL_SUB_LNEG: "- (x::real) - (y::real) = - (x + y)"
   by (import real REAL_SUB_LNEG)
 
-lemma REAL_SUB_NEG2: "ALL (x::real) y::real. - x - - y = y - x"
+lemma REAL_SUB_NEG2: "- (x::real) - - (y::real) = y - x"
   by (import real REAL_SUB_NEG2)
 
-lemma REAL_SUB_TRIANGLE: "ALL (a::real) (b::real) c::real. a - b + (b - c) = a - c"
+lemma REAL_SUB_TRIANGLE: "(a::real) - (b::real) + (b - (c::real)) = a - c"
   by (import real REAL_SUB_TRIANGLE)
 
-lemma REAL_INV_MUL: "ALL (x::real) y::real.
-   x ~= 0 & y ~= 0 --> inverse (x * y) = inverse x * inverse y"
+lemma REAL_INV_MUL: "(x::real) ~= (0::real) & (y::real) ~= (0::real)
+==> inverse (x * y) = inverse x * inverse y"
   by (import real REAL_INV_MUL)
 
-lemma REAL_SUB_INV2: "ALL (x::real) y::real.
-   x ~= 0 & y ~= 0 --> inverse x - inverse y = (y - x) / (x * y)"
+lemma REAL_SUB_INV2: "(x::real) ~= (0::real) & (y::real) ~= (0::real)
+==> inverse x - inverse y = (y - x) / (x * y)"
   by (import real REAL_SUB_INV2)
 
-lemma REAL_SUB_SUB2: "ALL (x::real) y::real. x - (x - y) = y"
+lemma REAL_SUB_SUB2: "(x::real) - (x - (y::real)) = y"
   by (import real REAL_SUB_SUB2)
 
-lemma REAL_ADD_SUB2: "ALL (x::real) y::real. x - (x + y) = - y"
+lemma REAL_ADD_SUB2: "(x::real) - (x + (y::real)) = - y"
   by (import real REAL_ADD_SUB2)
 
-lemma REAL_LE_MUL2: "ALL (x1::real) (x2::real) (y1::real) y2::real.
-   0 <= x1 & 0 <= y1 & x1 <= x2 & y1 <= y2 --> x1 * y1 <= x2 * y2"
-  by (import real REAL_LE_MUL2)
-
-lemma REAL_LE_DIV: "ALL (x::real) xa::real. 0 <= x & 0 <= xa --> 0 <= x / xa"
+lemma REAL_LE_DIV: "(0::real) <= (x::real) & (0::real) <= (xa::real) ==> (0::real) <= x / xa"
   by (import real REAL_LE_DIV)
 
-lemma REAL_LT_1: "ALL (x::real) y::real. 0 <= x & x < y --> x / y < 1"
+lemma REAL_LT_1: "(0::real) <= (x::real) & x < (y::real) ==> x / y < (1::real)"
   by (import real REAL_LT_1)
 
-lemma REAL_POS_NZ: "(All::(real => bool) => bool)
- (%x::real.
-     (op -->::bool => bool => bool)
-      ((op <::real => real => bool) (0::real) x)
-      ((Not::bool => bool) ((op =::real => real => bool) x (0::real))))"
+lemma REAL_POS_NZ: "(0::real) < (x::real) ==> x ~= (0::real)"
   by (import real REAL_POS_NZ)
 
-lemma REAL_EQ_LMUL_IMP: "ALL (x::real) (xa::real) xb::real. x ~= 0 & x * xa = x * xb --> xa = xb"
+lemma REAL_EQ_RMUL_IMP: "(z::real) ~= (0::real) & (x::real) * z = (y::real) * z ==> x = y"
+  by (import real REAL_EQ_RMUL_IMP)
+
+lemma REAL_EQ_LMUL_IMP: "(x::real) ~= (0::real) & x * (xa::real) = x * (xb::real) ==> xa = xb"
   by (import real REAL_EQ_LMUL_IMP)
 
-lemma REAL_FACT_NZ: "ALL n::nat. real (FACT n) ~= 0"
+lemma REAL_FACT_NZ: "real (FACT n) ~= 0"
   by (import real REAL_FACT_NZ)
 
-lemma REAL_DIFFSQ: "ALL (x::real) y::real. (x + y) * (x - y) = x * x - y * y"
-  by (import real REAL_DIFFSQ)
-
-lemma REAL_POASQ: "ALL x::real. (0 < x * x) = (x ~= 0)"
+lemma REAL_POASQ: "((0::real) < (x::real) * x) = (x ~= (0::real))"
   by (import real REAL_POASQ)
 
-lemma REAL_SUMSQ: "ALL (x::real) y::real. (x * x + y * y = 0) = (x = 0 & y = 0)"
-  by (import real REAL_SUMSQ)
-
-lemma REAL_DIV_MUL2: "ALL (x::real) z::real.
-   x ~= 0 & z ~= 0 --> (ALL y::real. y / z = x * y / (x * z))"
+lemma REAL_DIV_MUL2: "(x::real) ~= (0::real) & (z::real) ~= (0::real)
+==> (y::real) / z = x * y / (x * z)"
   by (import real REAL_DIV_MUL2)
 
-lemma REAL_MIDDLE1: "ALL (a::real) b::real. a <= b --> a <= (a + b) / 2"
+lemma REAL_MIDDLE1: "(a::real) <= (b::real) ==> a <= (a + b) / (2::real)"
   by (import real REAL_MIDDLE1)
 
-lemma REAL_MIDDLE2: "ALL (a::real) b::real. a <= b --> (a + b) / 2 <= b"
+lemma REAL_MIDDLE2: "(a::real) <= (b::real) ==> (a + b) / (2::real) <= b"
   by (import real REAL_MIDDLE2)
 
-lemma ABS_LT_MUL2: "ALL (w::real) (x::real) (y::real) z::real.
-   abs w < y & abs x < z --> abs (w * x) < y * z"
+lemma ABS_LT_MUL2: "abs (w::real) < (y::real) & abs (x::real) < (z::real)
+==> abs (w * x) < y * z"
   by (import real ABS_LT_MUL2)
 
-lemma ABS_REFL: "ALL x::real. (abs x = x) = (0 <= x)"
+lemma ABS_REFL: "(abs (x::real) = x) = ((0::real) <= x)"
   by (import real ABS_REFL)
 
-lemma ABS_BETWEEN: "ALL (x::real) (y::real) d::real.
-   (0 < d & x - d < y & y < x + d) = (abs (y - x) < d)"
+lemma ABS_BETWEEN: "((0::real) < (d::real) & (x::real) - d < (y::real) & y < x + d) =
+(abs (y - x) < d)"
   by (import real ABS_BETWEEN)
 
-lemma ABS_BOUND: "ALL (x::real) (y::real) d::real. abs (x - y) < d --> y < x + d"
+lemma ABS_BOUND: "abs ((x::real) - (y::real)) < (d::real) ==> y < x + d"
   by (import real ABS_BOUND)
 
-lemma ABS_STILLNZ: "ALL (x::real) y::real. abs (x - y) < abs y --> x ~= 0"
+lemma ABS_STILLNZ: "abs ((x::real) - (y::real)) < abs y ==> x ~= (0::real)"
   by (import real ABS_STILLNZ)
 
-lemma ABS_CASES: "ALL x::real. x = 0 | 0 < abs x"
+lemma ABS_CASES: "(x::real) = (0::real) | (0::real) < abs x"
   by (import real ABS_CASES)
 
-lemma ABS_BETWEEN1: "ALL (x::real) (y::real) z::real. x < z & abs (y - x) < z - x --> y < z"
+lemma ABS_BETWEEN1: "(x::real) < (z::real) & abs ((y::real) - x) < z - x ==> y < z"
   by (import real ABS_BETWEEN1)
 
-lemma ABS_SIGN: "ALL (x::real) y::real. abs (x - y) < y --> 0 < x"
+lemma ABS_SIGN: "abs ((x::real) - (y::real)) < y ==> (0::real) < x"
   by (import real ABS_SIGN)
 
-lemma ABS_SIGN2: "ALL (x::real) y::real. abs (x - y) < - y --> x < 0"
+lemma ABS_SIGN2: "abs ((x::real) - (y::real)) < - y ==> x < (0::real)"
   by (import real ABS_SIGN2)
 
-lemma ABS_CIRCLE: "ALL (x::real) (y::real) h::real.
-   abs h < abs y - abs x --> abs (x + h) < abs y"
+lemma ABS_CIRCLE: "abs (h::real) < abs (y::real) - abs (x::real) ==> abs (x + h) < abs y"
   by (import real ABS_CIRCLE)
 
-lemma ABS_BETWEEN2: "ALL (x0::real) (x::real) (y0::real) y::real.
-   x0 < y0 & abs (x - x0) < (y0 - x0) / 2 & abs (y - y0) < (y0 - x0) / 2 -->
-   x < y"
+lemma ABS_BETWEEN2: "(x0::real) < (y0::real) &
+abs ((x::real) - x0) < (y0 - x0) / (2::real) &
+abs ((y::real) - y0) < (y0 - x0) / (2::real)
+==> x < y"
   by (import real ABS_BETWEEN2)
 
-lemma POW_PLUS1: "ALL e>0. ALL n::nat. 1 + real n * e <= (1 + e) ^ n"
+lemma POW_PLUS1: "0 < e ==> 1 + real n * e <= (1 + e) ^ n"
   by (import real POW_PLUS1)
 
-lemma POW_M1: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op =::real => real => bool)
-      ((abs::real => real)
-        ((op ^::real => nat => real) ((uminus::real => real) (1::real)) n))
-      (1::real))"
+lemma POW_M1: "abs ((- (1::real)) ^ (n::nat)) = (1::real)"
   by (import real POW_M1)
 
-lemma REAL_LE1_POW2: "(All::(real => bool) => bool)
- (%x::real.
-     (op -->::bool => bool => bool)
-      ((op <=::real => real => bool) (1::real) x)
-      ((op <=::real => real => bool) (1::real)
-        ((op ^::real => nat => real) x
-          ((number_of \<Colon> int => nat)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))))"
+lemma REAL_LE1_POW2: "(1::real) <= (x::real) ==> (1::real) <= x ^ (2::nat)"
   by (import real REAL_LE1_POW2)
 
-lemma REAL_LT1_POW2: "(All::(real => bool) => bool)
- (%x::real.
-     (op -->::bool => bool => bool)
-      ((op <::real => real => bool) (1::real) x)
-      ((op <::real => real => bool) (1::real)
-        ((op ^::real => nat => real) x
-          ((number_of \<Colon> int => nat)
-            ((Int.Bit0 \<Colon> int => int)
-              ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))))"
+lemma REAL_LT1_POW2: "(1::real) < (x::real) ==> (1::real) < x ^ (2::nat)"
   by (import real REAL_LT1_POW2)
 
-lemma POW_POS_LT: "ALL (x::real) n::nat. 0 < x --> 0 < x ^ Suc n"
+lemma POW_POS_LT: "(0::real) < (x::real) ==> (0::real) < x ^ Suc (n::nat)"
   by (import real POW_POS_LT)
 
-lemma POW_LT: "ALL (n::nat) (x::real) y::real. 0 <= x & x < y --> x ^ Suc n < y ^ Suc n"
+lemma POW_LT: "(0::real) <= (x::real) & x < (y::real) ==> x ^ Suc (n::nat) < y ^ Suc n"
   by (import real POW_LT)
 
-lemma POW_ZERO_EQ: "ALL (n::nat) x::real. (x ^ Suc n = 0) = (x = 0)"
+lemma POW_ZERO: "(x::real) ^ (n::nat) = (0::real) ==> x = (0::real)"
+  by (import real POW_ZERO)
+
+lemma POW_ZERO_EQ: "((x::real) ^ Suc (n::nat) = (0::real)) = (x = (0::real))"
   by (import real POW_ZERO_EQ)
 
-lemma REAL_POW_LT2: "ALL (n::nat) (x::real) y::real. n ~= 0 & 0 <= x & x < y --> x ^ n < y ^ n"
+lemma REAL_POW_LT2: "(n::nat) ~= (0::nat) & (0::real) <= (x::real) & x < (y::real)
+==> x ^ n < y ^ n"
   by (import real REAL_POW_LT2)
 
-lemma REAL_POW_MONO_LT: "ALL (m::nat) (n::nat) x::real. 1 < x & m < n --> x ^ m < x ^ n"
+lemma REAL_POW_MONO_LT: "(1::real) < (x::real) & (m::nat) < (n::nat) ==> x ^ m < x ^ n"
   by (import real REAL_POW_MONO_LT)
 
-lemma REAL_SUP_SOMEPOS: "ALL P::real => bool.
-   (EX x::real. P x & 0 < x) & (EX z::real. ALL x::real. P x --> x < z) -->
-   (EX s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s))"
+lemma REAL_SUP_SOMEPOS: "(EX x::real. (P::real => bool) x & (0::real) < x) &
+(EX z::real. ALL x::real. P x --> x < z)
+==> EX s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)"
   by (import real REAL_SUP_SOMEPOS)
 
-lemma SUP_LEMMA1: "ALL (P::real => bool) (s::real) d::real.
-   (ALL y::real. (EX x::real. P (x + d) & y < x) = (y < s)) -->
-   (ALL y::real. (EX x::real. P x & y < x) = (y < s + d))"
+lemma SUP_LEMMA1: "(!!y::real.
+    (EX x::real. (P::real => bool) (x + (d::real)) & y < x) =
+    (y < (s::real)))
+==> (EX x::real. P x & (y::real) < x) = (y < s + d)"
   by (import real SUP_LEMMA1)
 
-lemma SUP_LEMMA2: "ALL P::real => bool. Ex P --> (EX (d::real) x::real. P (x + d) & 0 < x)"
+lemma SUP_LEMMA2: "Ex (P::real => bool) ==> EX (d::real) x::real. P (x + d) & (0::real) < x"
   by (import real SUP_LEMMA2)
 
-lemma SUP_LEMMA3: "ALL d::real.
-   (EX z::real. ALL x::real. (P::real => bool) x --> x < z) -->
-   (EX x::real. ALL xa::real. P (xa + d) --> xa < x)"
+lemma SUP_LEMMA3: "EX z::real. ALL x::real. (P::real => bool) x --> x < z
+==> EX x::real. ALL xa::real. P (xa + (d::real)) --> xa < x"
   by (import real SUP_LEMMA3)
 
-lemma REAL_SUP_EXISTS: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x < z) -->
-   (EX x::real. ALL y::real. (EX x::real. P x & y < x) = (y < x))"
+lemma REAL_SUP_EXISTS: "Ex (P::real => bool) & (EX z::real. ALL x::real. P x --> x < z)
+==> EX x::real. ALL y::real. (EX x::real. P x & y < x) = (y < x)"
   by (import real REAL_SUP_EXISTS)
 
-definition sup :: "(real => bool) => real" where 
-  "sup ==
-%P::real => bool.
-   SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)"
-
-lemma sup: "ALL P::real => bool.
-   sup P = (SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s))"
+consts
+  sup :: "(real => bool) => real" 
+
+defs
+  sup_def: "real.sup == %P. SOME s. ALL y. (EX x. P x & y < x) = (y < s)"
+
+lemma sup: "real.sup P = (SOME s. ALL y. (EX x. P x & y < x) = (y < s))"
   by (import real sup)
 
-lemma REAL_SUP: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x < z) -->
-   (ALL y::real. (EX x::real. P x & y < x) = (y < sup P))"
+lemma REAL_SUP: "Ex P & (EX z. ALL x. P x --> x < z)
+==> (EX x. P x & y < x) = (y < real.sup P)"
   by (import real REAL_SUP)
 
-lemma REAL_SUP_UBOUND: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x < z) -->
-   (ALL y::real. P y --> y <= sup P)"
+lemma REAL_SUP_UBOUND: "[| Ex P & (EX z. ALL x. P x --> x < z); P y |] ==> y <= real.sup P"
   by (import real REAL_SUP_UBOUND)
 
-lemma SETOK_LE_LT: "ALL P::real => bool.
-   (Ex P & (EX z::real. ALL x::real. P x --> x <= z)) =
-   (Ex P & (EX z::real. ALL x::real. P x --> x < z))"
+lemma SETOK_LE_LT: "(Ex (P::real => bool) & (EX z::real. ALL x::real. P x --> x <= z)) =
+(Ex P & (EX z::real. ALL x::real. P x --> x < z))"
   by (import real SETOK_LE_LT)
 
-lemma REAL_SUP_LE: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x <= z) -->
-   (ALL y::real. (EX x::real. P x & y < x) = (y < sup P))"
+lemma REAL_SUP_LE: "Ex P & (EX z. ALL x. P x --> x <= z)
+==> (EX x. P x & y < x) = (y < real.sup P)"
   by (import real REAL_SUP_LE)
 
-lemma REAL_SUP_UBOUND_LE: "ALL P::real => bool.
-   Ex P & (EX z::real. ALL x::real. P x --> x <= z) -->
-   (ALL y::real. P y --> y <= sup P)"
+lemma REAL_SUP_UBOUND_LE: "[| Ex P & (EX z. ALL x. P x --> x <= z); P y |] ==> y <= real.sup P"
   by (import real REAL_SUP_UBOUND_LE)
 
-lemma REAL_ARCH_LEAST: "ALL y>0. ALL x>=0. EX n::nat. real n * y <= x & x < real (Suc n) * y"
-  by (import real REAL_ARCH_LEAST)
-
 consts
   sumc :: "nat => nat => (nat => real) => real" 
 
-specification (sumc) sumc: "(ALL (n::nat) f::nat => real. sumc n 0 f = 0) &
-(ALL (n::nat) (m::nat) f::nat => real.
-    sumc n (Suc m) f = sumc n m f + f (n + m))"
+specification (sumc) sumc: "(ALL n f. sumc n 0 f = 0) &
+(ALL n m f. sumc n (Suc m) f = sumc n m f + f (n + m))"
   by (import real sumc)
 
 consts
   sum :: "nat * nat => (nat => real) => real" 
 
 defs
-  sum_def: "(op ==::(nat * nat => (nat => real) => real)
-        => (nat * nat => (nat => real) => real) => prop)
- (real.sum::nat * nat => (nat => real) => real)
- ((split::(nat => nat => (nat => real) => real)
-          => nat * nat => (nat => real) => real)
-   (sumc::nat => nat => (nat => real) => real))"
-
-lemma SUM_DEF: "ALL (m::nat) (n::nat) f::nat => real. real.sum (m, n) f = sumc m n f"
+  sum_def: "real.sum == %(x, y). sumc x y"
+
+lemma SUM_DEF: "real.sum (m, n) f = sumc m n f"
   by (import real SUM_DEF)
 
-lemma sum: "ALL (x::nat => real) (xa::nat) xb::nat.
-   real.sum (xa, 0) x = 0 &
-   real.sum (xa, Suc xb) x = real.sum (xa, xb) x + x (xa + xb)"
+lemma sum: "real.sum (xa, 0) x = 0 &
+real.sum (xa, Suc xb) x = real.sum (xa, xb) x + x (xa + xb)"
   by (import real sum)
 
-lemma SUM_TWO: "ALL (f::nat => real) (n::nat) p::nat.
-   real.sum (0, n) f + real.sum (n, p) f = real.sum (0, n + p) f"
+lemma SUM_TWO: "real.sum (0, n) f + real.sum (n, p) f = real.sum (0, n + p) f"
   by (import real SUM_TWO)
 
-lemma SUM_DIFF: "ALL (f::nat => real) (m::nat) n::nat.
-   real.sum (m, n) f = real.sum (0, m + n) f - real.sum (0, m) f"
+lemma SUM_DIFF: "real.sum (m, n) f = real.sum (0, m + n) f - real.sum (0, m) f"
   by (import real SUM_DIFF)
 
-lemma ABS_SUM: "ALL (f::nat => real) (m::nat) n::nat.
-   abs (real.sum (m, n) f) <= real.sum (m, n) (%n::nat. abs (f n))"
+lemma ABS_SUM: "abs (real.sum (m, n) f) <= real.sum (m, n) (%n. abs (f n))"
   by (import real ABS_SUM)
 
-lemma SUM_LE: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat.
-   (ALL r::nat. m <= r & r < n + m --> f r <= g r) -->
-   real.sum (m, n) f <= real.sum (m, n) g"
+lemma SUM_LE: "(!!r. m <= r & r < n + m ==> f r <= g r)
+==> real.sum (m, n) f <= real.sum (m, n) g"
   by (import real SUM_LE)
 
-lemma SUM_EQ: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat.
-   (ALL r::nat. m <= r & r < n + m --> f r = g r) -->
-   real.sum (m, n) f = real.sum (m, n) g"
+lemma SUM_EQ: "(!!r. m <= r & r < n + m ==> f r = g r)
+==> real.sum (m, n) f = real.sum (m, n) g"
   by (import real SUM_EQ)
 
-lemma SUM_POS: "ALL f::nat => real.
-   (ALL n::nat. 0 <= f n) --> (ALL (m::nat) n::nat. 0 <= real.sum (m, n) f)"
+lemma SUM_POS: "(!!n. 0 <= f n) ==> 0 <= real.sum (m, n) f"
   by (import real SUM_POS)
 
-lemma SUM_POS_GEN: "ALL (f::nat => real) m::nat.
-   (ALL n::nat. m <= n --> 0 <= f n) -->
-   (ALL n::nat. 0 <= real.sum (m, n) f)"
+lemma SUM_POS_GEN: "(!!n. m <= n ==> 0 <= f n) ==> 0 <= real.sum (m, n) f"
   by (import real SUM_POS_GEN)
 
-lemma SUM_ABS: "ALL (f::nat => real) (m::nat) x::nat.
-   abs (real.sum (m, x) (%m::nat. abs (f m))) =
-   real.sum (m, x) (%m::nat. abs (f m))"
+lemma SUM_ABS: "abs (real.sum (m, x) (%m. abs (f m))) = real.sum (m, x) (%m. abs (f m))"
   by (import real SUM_ABS)
 
-lemma SUM_ABS_LE: "ALL (f::nat => real) (m::nat) n::nat.
-   abs (real.sum (m, n) f) <= real.sum (m, n) (%n::nat. abs (f n))"
+lemma SUM_ABS_LE: "abs (real.sum (m, n) f) <= real.sum (m, n) (%n. abs (f n))"
   by (import real SUM_ABS_LE)
 
-lemma SUM_ZERO: "ALL (f::nat => real) N::nat.
-   (ALL n::nat. N <= n --> f n = 0) -->
-   (ALL (m::nat) n::nat. N <= m --> real.sum (m, n) f = 0)"
+lemma SUM_ZERO: "[| !!n. N <= n ==> f n = 0; N <= m |] ==> real.sum (m, n) f = 0"
   by (import real SUM_ZERO)
 
-lemma SUM_ADD: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat.
-   real.sum (m, n) (%n::nat. f n + g n) =
-   real.sum (m, n) f + real.sum (m, n) g"
+lemma SUM_ADD: "real.sum (m, n) (%n. f n + g n) = real.sum (m, n) f + real.sum (m, n) g"
   by (import real SUM_ADD)
 
-lemma SUM_CMUL: "ALL (f::nat => real) (c::real) (m::nat) n::nat.
-   real.sum (m, n) (%n::nat. c * f n) = c * real.sum (m, n) f"
+lemma SUM_CMUL: "real.sum (m, n) (%n. c * f n) = c * real.sum (m, n) f"
   by (import real SUM_CMUL)
 
-lemma SUM_NEG: "ALL (f::nat => real) (n::nat) d::nat.
-   real.sum (n, d) (%n::nat. - f n) = - real.sum (n, d) f"
+lemma SUM_NEG: "real.sum (n, d) (%n. - f n) = - real.sum (n, d) f"
   by (import real SUM_NEG)
 
-lemma SUM_SUB: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat.
-   real.sum (m, n) (%x::nat. f x - g x) =
-   real.sum (m, n) f - real.sum (m, n) g"
+lemma SUM_SUB: "real.sum (m, n) (%x. f x - g x) = real.sum (m, n) f - real.sum (m, n) g"
   by (import real SUM_SUB)
 
-lemma SUM_SUBST: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat.
-   (ALL p::nat. m <= p & p < m + n --> f p = g p) -->
-   real.sum (m, n) f = real.sum (m, n) g"
+lemma SUM_SUBST: "(!!p. m <= p & p < m + n ==> f p = g p)
+==> real.sum (m, n) f = real.sum (m, n) g"
   by (import real SUM_SUBST)
 
-lemma SUM_NSUB: "ALL (n::nat) (f::nat => real) c::real.
-   real.sum (0, n) f - real n * c = real.sum (0, n) (%p::nat. f p - c)"
+lemma SUM_NSUB: "real.sum (0, n) f - real n * c = real.sum (0, n) (%p. f p - c)"
   by (import real SUM_NSUB)
 
-lemma SUM_BOUND: "ALL (f::nat => real) (k::real) (m::nat) n::nat.
-   (ALL p::nat. m <= p & p < m + n --> f p <= k) -->
-   real.sum (m, n) f <= real n * k"
+lemma SUM_BOUND: "(!!p. m <= p & p < m + n ==> f p <= k) ==> real.sum (m, n) f <= real n * k"
   by (import real SUM_BOUND)
 
-lemma SUM_GROUP: "ALL (n::nat) (k::nat) f::nat => real.
-   real.sum (0, n) (%m::nat. real.sum (m * k, k) f) = real.sum (0, n * k) f"
+lemma SUM_GROUP: "real.sum (0, n) (%m. real.sum (m * k, k) f) = real.sum (0, n * k) f"
   by (import real SUM_GROUP)
 
-lemma SUM_1: "ALL (f::nat => real) n::nat. real.sum (n, 1) f = f n"
+lemma SUM_1: "real.sum (n, 1) f = f n"
   by (import real SUM_1)
 
-lemma SUM_2: "ALL (f::nat => real) n::nat. real.sum (n, 2) f = f n + f (n + 1)"
+lemma SUM_2: "real.sum (n, 2) f = f n + f (n + 1)"
   by (import real SUM_2)
 
-lemma SUM_OFFSET: "ALL (f::nat => real) (n::nat) k::nat.
-   real.sum (0, n) (%m::nat. f (m + k)) =
-   real.sum (0, n + k) f - real.sum (0, k) f"
+lemma SUM_OFFSET: "real.sum (0, n) (%m. f (m + k)) = real.sum (0, n + k) f - real.sum (0, k) f"
   by (import real SUM_OFFSET)
 
-lemma SUM_REINDEX: "ALL (f::nat => real) (m::nat) (k::nat) n::nat.
-   real.sum (m + k, n) f = real.sum (m, n) (%r::nat. f (r + k))"
+lemma SUM_REINDEX: "real.sum (m + k, n) f = real.sum (m, n) (%r. f (r + k))"
   by (import real SUM_REINDEX)
 
-lemma SUM_0: "ALL (m::nat) n::nat. real.sum (m, n) (%r::nat. 0) = 0"
+lemma SUM_0: "real.sum (m, n) (%r. 0) = 0"
   by (import real SUM_0)
 
-lemma SUM_PERMUTE_0: "(All::(nat => bool) => bool)
- (%n::nat.
-     (All::((nat => nat) => bool) => bool)
-      (%p::nat => nat.
-          (op -->::bool => bool => bool)
-           ((All::(nat => bool) => bool)
-             (%y::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <::nat => nat => bool) y n)
-                  ((Ex1::(nat => bool) => bool)
-                    (%x::nat.
-                        (op &::bool => bool => bool)
-                         ((op <::nat => nat => bool) x n)
-                         ((op =::nat => nat => bool) (p x) y)))))
-           ((All::((nat => real) => bool) => bool)
-             (%f::nat => real.
-                 (op =::real => real => bool)
-                  ((real.sum::nat * nat => (nat => real) => real)
-                    ((Pair::nat => nat => nat * nat) (0::nat) n)
-                    (%n::nat. f (p n)))
-                  ((real.sum::nat * nat => (nat => real) => real)
-                    ((Pair::nat => nat => nat * nat) (0::nat) n) f)))))"
+lemma SUM_PERMUTE_0: "(!!y. y < n ==> EX! x. x < n & p x = y)
+==> real.sum (0, n) (%n. f (p n)) = real.sum (0, n) f"
   by (import real SUM_PERMUTE_0)
 
-lemma SUM_CANCEL: "ALL (f::nat => real) (n::nat) d::nat.
-   real.sum (n, d) (%n::nat. f (Suc n) - f n) = f (n + d) - f n"
+lemma SUM_CANCEL: "real.sum (n, d) (%n. f (Suc n) - f n) = f (n + d) - f n"
   by (import real SUM_CANCEL)
 
-lemma REAL_EQ_RDIV_EQ: "ALL (x::real) (xa::real) xb::real. 0 < xb --> (x = xa / xb) = (x * xb = xa)"
+lemma REAL_LE_RNEG: "((x::real) <= - (y::real)) = (x + y <= (0::real))"
+  by (import real REAL_LE_RNEG)
+
+lemma REAL_EQ_RDIV_EQ: "(0::real) < (xb::real) ==> ((x::real) = (xa::real) / xb) = (x * xb = xa)"
   by (import real REAL_EQ_RDIV_EQ)
 
-lemma REAL_EQ_LDIV_EQ: "ALL (x::real) (xa::real) xb::real. 0 < xb --> (x / xb = xa) = (x = xa * xb)"
+lemma REAL_EQ_LDIV_EQ: "(0::real) < (xb::real) ==> ((x::real) / xb = (xa::real)) = (x = xa * xb)"
   by (import real REAL_EQ_LDIV_EQ)
 
 ;end_setup
 
 ;setup_theory topology
 
-definition re_Union :: "(('a => bool) => bool) => 'a => bool" where 
-  "re_Union ==
-%(P::('a::type => bool) => bool) x::'a::type.
-   EX s::'a::type => bool. P s & s x"
-
-lemma re_Union: "ALL P::('a::type => bool) => bool.
-   re_Union P = (%x::'a::type. EX s::'a::type => bool. P s & s x)"
+definition
+  re_Union :: "(('a => bool) => bool) => 'a => bool"  where
+  "re_Union == %P x. EX s. P s & s x"
+
+lemma re_Union: "re_Union P = (%x. EX s. P s & s x)"
   by (import topology re_Union)
 
-definition re_union :: "('a => bool) => ('a => bool) => 'a => bool" where 
-  "re_union ==
-%(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x | Q x"
-
-lemma re_union: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   re_union P Q = (%x::'a::type. P x | Q x)"
+definition
+  re_union :: "('a => bool) => ('a => bool) => 'a => bool"  where
+  "re_union == %P Q x. P x | Q x"
+
+lemma re_union: "re_union P Q = (%x. P x | Q x)"
   by (import topology re_union)
 
-definition re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" where 
-  "re_intersect ==
-%(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x & Q x"
-
-lemma re_intersect: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   re_intersect P Q = (%x::'a::type. P x & Q x)"
+definition
+  re_intersect :: "('a => bool) => ('a => bool) => 'a => bool"  where
+  "re_intersect == %P Q x. P x & Q x"
+
+lemma re_intersect: "re_intersect P Q = (%x. P x & Q x)"
   by (import topology re_intersect)
 
-definition re_null :: "'a => bool" where 
-  "re_null == %x::'a::type. False"
-
-lemma re_null: "re_null = (%x::'a::type. False)"
+definition
+  re_null :: "'a => bool"  where
+  "re_null == %x. False"
+
+lemma re_null: "re_null = (%x. False)"
   by (import topology re_null)
 
-definition re_universe :: "'a => bool" where 
-  "re_universe == %x::'a::type. True"
-
-lemma re_universe: "re_universe = (%x::'a::type. True)"
+definition
+  re_universe :: "'a => bool"  where
+  "re_universe == %x. True"
+
+lemma re_universe: "re_universe = (%x. True)"
   by (import topology re_universe)
 
-definition re_subset :: "('a => bool) => ('a => bool) => bool" where 
-  "re_subset ==
-%(P::'a::type => bool) Q::'a::type => bool. ALL x::'a::type. P x --> Q x"
-
-lemma re_subset: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   re_subset P Q = (ALL x::'a::type. P x --> Q x)"
+definition
+  re_subset :: "('a => bool) => ('a => bool) => bool"  where
+  "re_subset == %P Q. ALL x. P x --> Q x"
+
+lemma re_subset: "re_subset P Q = (ALL x. P x --> Q x)"
   by (import topology re_subset)
 
-definition re_compl :: "('a => bool) => 'a => bool" where 
-  "re_compl == %(P::'a::type => bool) x::'a::type. ~ P x"
-
-lemma re_compl: "ALL P::'a::type => bool. re_compl P = (%x::'a::type. ~ P x)"
+definition
+  re_compl :: "('a => bool) => 'a => bool"  where
+  "re_compl == %P x. ~ P x"
+
+lemma re_compl: "re_compl P = (%x. ~ P x)"
   by (import topology re_compl)
 
-lemma SUBSET_REFL: "ALL P::'a::type => bool. re_subset P P"
+lemma SUBSET_REFL: "re_subset P P"
   by (import topology SUBSET_REFL)
 
-lemma COMPL_MEM: "ALL (P::'a::type => bool) x::'a::type. P x = (~ re_compl P x)"
+lemma COMPL_MEM: "P x = (~ re_compl P x)"
   by (import topology COMPL_MEM)
 
-lemma SUBSET_ANTISYM: "ALL (P::'a::type => bool) Q::'a::type => bool.
-   (re_subset P Q & re_subset Q P) = (P = Q)"
+lemma SUBSET_ANTISYM: "(re_subset P Q & re_subset Q P) = (P = Q)"
   by (import topology SUBSET_ANTISYM)
 
-lemma SUBSET_TRANS: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool.
-   re_subset P Q & re_subset Q R --> re_subset P R"
+lemma SUBSET_TRANS: "re_subset P Q & re_subset Q R ==> re_subset P R"
   by (import topology SUBSET_TRANS)
 
-definition istopology :: "(('a => bool) => bool) => bool" where 
+definition
+  istopology :: "(('a => bool) => bool) => bool"  where
   "istopology ==
-%L::('a::type => bool) => bool.
-   L re_null &
-   L re_universe &
-   (ALL (a::'a::type => bool) b::'a::type => bool.
-       L a & L b --> L (re_intersect a b)) &
-   (ALL P::('a::type => bool) => bool. re_subset P L --> L (re_Union P))"
-
-lemma istopology: "ALL L::('a::type => bool) => bool.
-   istopology L =
-   (L re_null &
+%L. L re_null &
     L re_universe &
-    (ALL (a::'a::type => bool) b::'a::type => bool.
-        L a & L b --> L (re_intersect a b)) &
-    (ALL P::('a::type => bool) => bool. re_subset P L --> L (re_Union P)))"
+    (ALL a b. L a & L b --> L (re_intersect a b)) &
+    (ALL P. re_subset P L --> L (re_Union P))"
+
+lemma istopology: "istopology L =
+(L re_null &
+ L re_universe &
+ (ALL a b. L a & L b --> L (re_intersect a b)) &
+ (ALL P. re_subset P L --> L (re_Union P)))"
   by (import topology istopology)
 
-typedef (open) ('a) topology = "(Collect::((('a::type => bool) => bool) => bool)
-          => (('a::type => bool) => bool) set)
- (istopology::(('a::type => bool) => bool) => bool)" 
+typedef (open) ('a) topology = "Collect istopology::(('a::type => bool) => bool) set"
   by (rule typedef_helper,import topology topology_TY_DEF)
 
 lemmas topology_TY_DEF = typedef_hol2hol4 [OF type_definition_topology]
@@ -864,94 +708,84 @@
   topology :: "(('a => bool) => bool) => 'a topology" 
   "open" :: "'a topology => ('a => bool) => bool" 
 
-specification ("open" topology) topology_tybij: "(ALL a::'a::type topology. topology (open a) = a) &
-(ALL r::('a::type => bool) => bool. istopology r = (open (topology r) = r))"
+specification ("open" topology) topology_tybij: "(ALL a::'a topology. topology (topology.open a) = a) &
+(ALL r::('a => bool) => bool.
+    istopology r = (topology.open (topology r) = r))"
   by (import topology topology_tybij)
 
-lemma TOPOLOGY: "ALL L::'a::type topology.
-   open L re_null &
-   open L re_universe &
-   (ALL (a::'a::type => bool) b::'a::type => bool.
-       open L a & open L b --> open L (re_intersect a b)) &
-   (ALL P::('a::type => bool) => bool.
-       re_subset P (open L) --> open L (re_Union P))"
+lemma TOPOLOGY: "topology.open L re_null &
+topology.open L re_universe &
+(ALL a b.
+    topology.open L a & topology.open L b -->
+    topology.open L (re_intersect a b)) &
+(ALL P. re_subset P (topology.open L) --> topology.open L (re_Union P))"
   by (import topology TOPOLOGY)
 
-lemma TOPOLOGY_UNION: "ALL (x::'a::type topology) xa::('a::type => bool) => bool.
-   re_subset xa (open x) --> open x (re_Union xa)"
+lemma TOPOLOGY_UNION: "re_subset xa (topology.open x) ==> topology.open x (re_Union xa)"
   by (import topology TOPOLOGY_UNION)
 
-definition neigh :: "'a topology => ('a => bool) * 'a => bool" where 
-  "neigh ==
-%(top::'a::type topology) (N::'a::type => bool, x::'a::type).
-   EX P::'a::type => bool. open top P & re_subset P N & P x"
-
-lemma neigh: "ALL (top::'a::type topology) (N::'a::type => bool) x::'a::type.
-   neigh top (N, x) =
-   (EX P::'a::type => bool. open top P & re_subset P N & P x)"
+definition
+  neigh :: "'a topology => ('a => bool) * 'a => bool"  where
+  "neigh == %tp (N, x). EX P. topology.open tp P & re_subset P N & P x"
+
+lemma neigh: "neigh (tp::'a::type topology) (N::'a::type => bool, x::'a::type) =
+(EX P::'a::type => bool. topology.open tp P & re_subset P N & P x)"
   by (import topology neigh)
 
-lemma OPEN_OWN_NEIGH: "ALL (S'::'a::type => bool) (top::'a::type topology) x::'a::type.
-   open top S' & S' x --> neigh top (S', x)"
+lemma OPEN_OWN_NEIGH: "topology.open (tp::'a::type topology) (S'::'a::type => bool) &
+S' (x::'a::type)
+==> neigh tp (S', x)"
   by (import topology OPEN_OWN_NEIGH)
 
-lemma OPEN_UNOPEN: "ALL (S'::'a::type => bool) top::'a::type topology.
-   open top S' =
-   (re_Union (%P::'a::type => bool. open top P & re_subset P S') = S')"
+lemma OPEN_UNOPEN: "topology.open (tp::'a::type topology) (S'::'a::type => bool) =
+(re_Union (%P::'a::type => bool. topology.open tp P & re_subset P S') = S')"
   by (import topology OPEN_UNOPEN)
 
-lemma OPEN_SUBOPEN: "ALL (S'::'a::type => bool) top::'a::type topology.
-   open top S' =
-   (ALL x::'a::type.
-       S' x --> (EX P::'a::type => bool. P x & open top P & re_subset P S'))"
+lemma OPEN_SUBOPEN: "topology.open (tp::'a::type topology) (S'::'a::type => bool) =
+(ALL x::'a::type.
+    S' x -->
+    (EX P::'a::type => bool. P x & topology.open tp P & re_subset P S'))"
   by (import topology OPEN_SUBOPEN)
 
-lemma OPEN_NEIGH: "ALL (S'::'a::type => bool) top::'a::type topology.
-   open top S' =
-   (ALL x::'a::type.
-       S' x --> (EX N::'a::type => bool. neigh top (N, x) & re_subset N S'))"
+lemma OPEN_NEIGH: "topology.open (tp::'a::type topology) (S'::'a::type => bool) =
+(ALL x::'a::type.
+    S' x --> (EX N::'a::type => bool. neigh tp (N, x) & re_subset N S'))"
   by (import topology OPEN_NEIGH)
 
-definition closed :: "'a topology => ('a => bool) => bool" where 
-  "closed == %(L::'a::type topology) S'::'a::type => bool. open L (re_compl S')"
-
-lemma closed: "ALL (L::'a::type topology) S'::'a::type => bool.
-   closed L S' = open L (re_compl S')"
+consts
+  closed :: "'a topology => ('a => bool) => bool" 
+
+defs
+  closed_def: "topology.closed == %L S'. topology.open L (re_compl S')"
+
+lemma closed: "topology.closed L S' = topology.open L (re_compl S')"
   by (import topology closed)
 
-definition limpt :: "'a topology => 'a => ('a => bool) => bool" where 
-  "limpt ==
-%(top::'a::type topology) (x::'a::type) S'::'a::type => bool.
-   ALL N::'a::type => bool.
-      neigh top (N, x) --> (EX y::'a::type. x ~= y & S' y & N y)"
-
-lemma limpt: "ALL (top::'a::type topology) (x::'a::type) S'::'a::type => bool.
-   limpt top x S' =
-   (ALL N::'a::type => bool.
-       neigh top (N, x) --> (EX y::'a::type. x ~= y & S' y & N y))"
+definition
+  limpt :: "'a topology => 'a => ('a => bool) => bool"  where
+  "limpt == %tp x S'. ALL N. neigh tp (N, x) --> (EX y. x ~= y & S' y & N y)"
+
+lemma limpt: "limpt (tp::'a::type topology) (x::'a::type) (S'::'a::type => bool) =
+(ALL N::'a::type => bool.
+    neigh tp (N, x) --> (EX y::'a::type. x ~= y & S' y & N y))"
   by (import topology limpt)
 
-lemma CLOSED_LIMPT: "ALL (top::'a::type topology) S'::'a::type => bool.
-   closed top S' = (ALL x::'a::type. limpt top x S' --> S' x)"
+lemma CLOSED_LIMPT: "topology.closed (tp::'a::type topology) (S'::'a::type => bool) =
+(ALL x::'a::type. limpt tp x S' --> S' x)"
   by (import topology CLOSED_LIMPT)
 
-definition ismet :: "('a * 'a => real) => bool" where 
+definition
+  ismet :: "('a * 'a => real) => bool"  where
   "ismet ==
-%m::'a::type * 'a::type => real.
-   (ALL (x::'a::type) y::'a::type. (m (x, y) = 0) = (x = y)) &
-   (ALL (x::'a::type) (y::'a::type) z::'a::type.
-       m (y, z) <= m (x, y) + m (x, z))"
-
-lemma ismet: "ALL m::'a::type * 'a::type => real.
-   ismet m =
-   ((ALL (x::'a::type) y::'a::type. (m (x, y) = 0) = (x = y)) &
-    (ALL (x::'a::type) (y::'a::type) z::'a::type.
-        m (y, z) <= m (x, y) + m (x, z)))"
+%m. (ALL x y. (m (x, y) = 0) = (x = y)) &
+    (ALL x y z. m (y, z) <= m (x, y) + m (x, z))"
+
+lemma ismet: "ismet m =
+((ALL x y. (m (x, y) = 0) = (x = y)) &
+ (ALL x y z. m (y, z) <= m (x, y) + m (x, z)))"
   by (import topology ismet)
 
-typedef (open) ('a) metric = "(Collect::(('a::type * 'a::type => real) => bool)
-          => ('a::type * 'a::type => real) set)
- (ismet::('a::type * 'a::type => real) => bool)" 
+typedef (open) ('a) metric = "Collect ismet :: ('a::type * 'a::type => real) set"
   by (rule typedef_helper,import topology metric_TY_DEF)
 
 lemmas metric_TY_DEF = typedef_hol2hol4 [OF type_definition_metric]
@@ -960,330 +794,265 @@
   metric :: "('a * 'a => real) => 'a metric" 
   dist :: "'a metric => 'a * 'a => real" 
 
-specification (dist metric) metric_tybij: "(ALL a::'a::type metric. metric (dist a) = a) &
-(ALL r::'a::type * 'a::type => real. ismet r = (dist (metric r) = r))"
+specification (dist metric) metric_tybij: "(ALL a::'a metric. metric (topology.dist a) = a) &
+(ALL r::'a * 'a => real. ismet r = (topology.dist (metric r) = r))"
   by (import topology metric_tybij)
 
-lemma METRIC_ISMET: "ALL m::'a::type metric. ismet (dist m)"
+lemma METRIC_ISMET: "ismet (topology.dist m)"
   by (import topology METRIC_ISMET)
 
-lemma METRIC_ZERO: "ALL (m::'a::type metric) (x::'a::type) y::'a::type.
-   (dist m (x, y) = 0) = (x = y)"
+lemma METRIC_ZERO: "(topology.dist m (x, y) = 0) = (x = y)"
   by (import topology METRIC_ZERO)
 
-lemma METRIC_SAME: "ALL (m::'a::type metric) x::'a::type. dist m (x, x) = 0"
+lemma METRIC_SAME: "topology.dist m (x, x) = 0"
   by (import topology METRIC_SAME)
 
-lemma METRIC_POS: "ALL (m::'a::type metric) (x::'a::type) y::'a::type. 0 <= dist m (x, y)"
+lemma METRIC_POS: "0 <= topology.dist m (x, y)"
   by (import topology METRIC_POS)
 
-lemma METRIC_SYM: "ALL (m::'a::type metric) (x::'a::type) y::'a::type.
-   dist m (x, y) = dist m (y, x)"
+lemma METRIC_SYM: "topology.dist m (x, y) = topology.dist m (y, x)"
   by (import topology METRIC_SYM)
 
-lemma METRIC_TRIANGLE: "ALL (m::'a::type metric) (x::'a::type) (y::'a::type) z::'a::type.
-   dist m (x, z) <= dist m (x, y) + dist m (y, z)"
+lemma METRIC_TRIANGLE: "topology.dist m (x, z) <= topology.dist m (x, y) + topology.dist m (y, z)"
   by (import topology METRIC_TRIANGLE)
 
-lemma METRIC_NZ: "ALL (m::'a::type metric) (x::'a::type) y::'a::type.
-   x ~= y --> 0 < dist m (x, y)"
+lemma METRIC_NZ: "x ~= y ==> 0 < topology.dist m (x, y)"
   by (import topology METRIC_NZ)
 
-definition mtop :: "'a metric => 'a topology" where 
+definition
+  mtop :: "'a metric => 'a topology"  where
   "mtop ==
-%m::'a::type metric.
-   topology
-    (%S'::'a::type => bool.
-        ALL x::'a::type.
-           S' x --> (EX e>0. ALL y::'a::type. dist m (x, y) < e --> S' y))"
-
-lemma mtop: "ALL m::'a::type metric.
-   mtop m =
-   topology
-    (%S'::'a::type => bool.
-        ALL x::'a::type.
-           S' x --> (EX e>0. ALL y::'a::type. dist m (x, y) < e --> S' y))"
+%m. topology
+     (%S'. ALL x.
+              S' x --> (EX e>0. ALL y. topology.dist m (x, y) < e --> S' y))"
+
+lemma mtop: "mtop m =
+topology
+ (%S'. ALL x. S' x --> (EX e>0. ALL y. topology.dist m (x, y) < e --> S' y))"
   by (import topology mtop)
 
-lemma mtop_istopology: "ALL m::'a::type metric.
-   istopology
-    (%S'::'a::type => bool.
-        ALL x::'a::type.
-           S' x --> (EX e>0. ALL y::'a::type. dist m (x, y) < e --> S' y))"
+lemma mtop_istopology: "istopology
+ (%S'. ALL x. S' x --> (EX e>0. ALL y. topology.dist m (x, y) < e --> S' y))"
   by (import topology mtop_istopology)
 
-lemma MTOP_OPEN: "ALL (S'::'a::type => bool) x::'a::type metric.
-   open (mtop x) S' =
-   (ALL xa::'a::type.
-       S' xa --> (EX e>0. ALL y::'a::type. dist x (xa, y) < e --> S' y))"
+lemma MTOP_OPEN: "topology.open (mtop x) S' =
+(ALL xa. S' xa --> (EX e>0. ALL y. topology.dist x (xa, y) < e --> S' y))"
   by (import topology MTOP_OPEN)
 
-definition B :: "'a metric => 'a * real => 'a => bool" where 
-  "B ==
-%(m::'a::type metric) (x::'a::type, e::real) y::'a::type. dist m (x, y) < e"
-
-lemma ball: "ALL (m::'a::type metric) (x::'a::type) e::real.
-   B m (x, e) = (%y::'a::type. dist m (x, y) < e)"
+definition
+  B :: "'a metric => 'a * real => 'a => bool"  where
+  "B == %m (x, e) y. topology.dist m (x, y) < e"
+
+lemma ball: "B m (x, e) = (%y. topology.dist m (x, y) < e)"
   by (import topology ball)
 
-lemma BALL_OPEN: "ALL (m::'a::type metric) (x::'a::type) e::real.
-   0 < e --> open (mtop m) (B m (x, e))"
+lemma BALL_OPEN: "0 < e ==> topology.open (mtop m) (B m (x, e))"
   by (import topology BALL_OPEN)
 
-lemma BALL_NEIGH: "ALL (m::'a::type metric) (x::'a::type) e::real.
-   0 < e --> neigh (mtop m) (B m (x, e), x)"
+lemma BALL_NEIGH: "0 < e ==> neigh (mtop m) (B m (x, e), x)"
   by (import topology BALL_NEIGH)
 
-lemma MTOP_LIMPT: "ALL (m::'a::type metric) (x::'a::type) S'::'a::type => bool.
-   limpt (mtop m) x S' =
-   (ALL e>0. EX y::'a::type. x ~= y & S' y & dist m (x, y) < e)"
+lemma MTOP_LIMPT: "limpt (mtop m) x S' =
+(ALL e>0. EX y. x ~= y & S' y & topology.dist m (x, y) < e)"
   by (import topology MTOP_LIMPT)
 
-lemma ISMET_R1: "ismet (%(x::real, y::real). abs (y - x))"
+lemma ISMET_R1: "ismet (%(x, y). abs (y - x))"
   by (import topology ISMET_R1)
 
-definition mr1 :: "real metric" where 
-  "mr1 == metric (%(x::real, y::real). abs (y - x))"
-
-lemma mr1: "mr1 = metric (%(x::real, y::real). abs (y - x))"
+definition
+  mr1 :: "real metric"  where
+  "mr1 == metric (%(x, y). abs (y - x))"
+
+lemma mr1: "mr1 = metric (%(x, y). abs (y - x))"
   by (import topology mr1)
 
-lemma MR1_DEF: "ALL (x::real) y::real. dist mr1 (x, y) = abs (y - x)"
+lemma MR1_DEF: "topology.dist mr1 (x, y) = abs (y - x)"
   by (import topology MR1_DEF)
 
-lemma MR1_ADD: "ALL (x::real) d::real. dist mr1 (x, x + d) = abs d"
+lemma MR1_ADD: "topology.dist mr1 (x, x + d) = abs d"
   by (import topology MR1_ADD)
 
-lemma MR1_SUB: "ALL (x::real) d::real. dist mr1 (x, x - d) = abs d"
+lemma MR1_SUB: "topology.dist mr1 (x, x - d) = abs d"
   by (import topology MR1_SUB)
 
-lemma MR1_ADD_POS: "ALL (x::real) d::real. 0 <= d --> dist mr1 (x, x + d) = d"
+lemma MR1_ADD_POS: "0 <= d ==> topology.dist mr1 (x, x + d) = d"
   by (import topology MR1_ADD_POS)
 
-lemma MR1_SUB_LE: "ALL (x::real) d::real. 0 <= d --> dist mr1 (x, x - d) = d"
+lemma MR1_SUB_LE: "0 <= d ==> topology.dist mr1 (x, x - d) = d"
   by (import topology MR1_SUB_LE)
 
-lemma MR1_ADD_LT: "ALL (x::real) d::real. 0 < d --> dist mr1 (x, x + d) = d"
+lemma MR1_ADD_LT: "0 < d ==> topology.dist mr1 (x, x + d) = d"
   by (import topology MR1_ADD_LT)
 
-lemma MR1_SUB_LT: "ALL (x::real) d::real. 0 < d --> dist mr1 (x, x - d) = d"
+lemma MR1_SUB_LT: "0 < d ==> topology.dist mr1 (x, x - d) = d"
   by (import topology MR1_SUB_LT)
 
-lemma MR1_BETWEEN1: "ALL (x::real) (y::real) z::real. x < z & dist mr1 (x, y) < z - x --> y < z"
+lemma MR1_BETWEEN1: "x < z & topology.dist mr1 (x, y) < z - x ==> y < z"
   by (import topology MR1_BETWEEN1)
 
-lemma MR1_LIMPT: "ALL x::real. limpt (mtop mr1) x re_universe"
+lemma MR1_LIMPT: "limpt (mtop mr1) x re_universe"
   by (import topology MR1_LIMPT)
 
 ;end_setup
 
 ;setup_theory nets
 
-definition dorder :: "('a => 'a => bool) => bool" where 
+definition
+  dorder :: "('a => 'a => bool) => bool"  where
   "dorder ==
-%g::'a::type => 'a::type => bool.
-   ALL (x::'a::type) y::'a::type.
-      g x x & g y y -->
-      (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y))"
-
-lemma dorder: "ALL g::'a::type => 'a::type => bool.
-   dorder g =
-   (ALL (x::'a::type) y::'a::type.
-       g x x & g y y -->
-       (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y)))"
+%g. ALL x y.
+       g x x & g y y --> (EX z. g z z & (ALL w. g w z --> g w x & g w y))"
+
+lemma dorder: "dorder g =
+(ALL x y.
+    g x x & g y y --> (EX z. g z z & (ALL w. g w z --> g w x & g w y)))"
   by (import nets dorder)
 
-definition tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" where 
+definition
+  tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool"  where
   "tends ==
-%(s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology,
-   g::'b::type => 'b::type => bool).
-   ALL N::'a::type => bool.
-      neigh top (N, l) -->
-      (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m)))"
-
-lemma tends: "ALL (s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology)
-   g::'b::type => 'b::type => bool.
-   tends s l (top, g) =
-   (ALL N::'a::type => bool.
-       neigh top (N, l) -->
-       (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m))))"
+%(s::'b => 'a) (l::'a) (tp::'a topology, g::'b => 'b => bool).
+   ALL N::'a => bool.
+      neigh tp (N, l) -->
+      (EX n::'b. g n n & (ALL m::'b. g m n --> N (s m)))"
+
+lemma tends: "tends (s::'b::type => 'a::type) (l::'a::type)
+ (tp::'a::type topology, g::'b::type => 'b::type => bool) =
+(ALL N::'a::type => bool.
+    neigh tp (N, l) -->
+    (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m))))"
   by (import nets tends)
 
-definition bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" where 
+definition
+  bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool"  where
   "bounded ==
-%(m::'a::type metric, g::'b::type => 'b::type => bool)
-   f::'b::type => 'a::type.
-   EX (k::real) (x::'a::type) N::'b::type.
-      g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k)"
-
-lemma bounded: "ALL (m::'a::type metric) (g::'b::type => 'b::type => bool)
-   f::'b::type => 'a::type.
-   bounded (m, g) f =
-   (EX (k::real) (x::'a::type) N::'b::type.
-       g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k))"
+%(m, g) f. EX k x N. g N N & (ALL n. g n N --> topology.dist m (f n, x) < k)"
+
+lemma bounded: "bounded (m, g) f =
+(EX k x N. g N N & (ALL n. g n N --> topology.dist m (f n, x) < k))"
   by (import nets bounded)
 
-definition tendsto :: "'a metric * 'a => 'a => 'a => bool" where 
-  "tendsto ==
-%(m::'a::type metric, x::'a::type) (y::'a::type) z::'a::type.
-   0 < dist m (x, y) & dist m (x, y) <= dist m (x, z)"
-
-lemma tendsto: "ALL (m::'a::type metric) (x::'a::type) (y::'a::type) z::'a::type.
-   tendsto (m, x) y z = (0 < dist m (x, y) & dist m (x, y) <= dist m (x, z))"
+consts
+  tendsto :: "'a metric * 'a => 'a => 'a => bool" 
+
+defs
+  tendsto_def: "nets.tendsto ==
+%(m, x) y z.
+   0 < topology.dist m (x, y) &
+   topology.dist m (x, y) <= topology.dist m (x, z)"
+
+lemma tendsto: "nets.tendsto (m, x) y z =
+(0 < topology.dist m (x, y) &
+ topology.dist m (x, y) <= topology.dist m (x, z))"
   by (import nets tendsto)
 
-lemma DORDER_LEMMA: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (P::'a::type => bool) Q::'a::type => bool.
-       (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> P m)) &
-       (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> Q m)) -->
-       (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> P m & Q m)))"
+lemma DORDER_LEMMA: "[| dorder g;
+   (EX n. g n n & (ALL m. g m n --> P m)) &
+   (EX n. g n n & (ALL m. g m n --> Q m)) |]
+==> EX n. g n n & (ALL m. g m n --> P m & Q m)"
   by (import nets DORDER_LEMMA)
 
 lemma DORDER_NGE: "dorder nat_ge"
   by (import nets DORDER_NGE)
 
-lemma DORDER_TENDSTO: "ALL (m::'a::type metric) x::'a::type. dorder (tendsto (m, x))"
+lemma DORDER_TENDSTO: "dorder (nets.tendsto (m, x))"
   by (import nets DORDER_TENDSTO)
 
-lemma MTOP_TENDS: "ALL (d::'a::type metric) (g::'b::type => 'b::type => bool)
-   (x::'b::type => 'a::type) x0::'a::type.
-   tends x x0 (mtop d, g) =
-   (ALL e>0.
-       EX n::'b::type.
-          g n n & (ALL m::'b::type. g m n --> dist d (x m, x0) < e))"
+lemma MTOP_TENDS: "tends (x::'b => 'a) (x0::'a) (mtop (d::'a metric), g::'b => 'b => bool) =
+(ALL e>0::real.
+    EX n::'b. g n n & (ALL m::'b. g m n --> topology.dist d (x m, x0) < e))"
   by (import nets MTOP_TENDS)
 
-lemma MTOP_TENDS_UNIQ: "ALL (g::'b::type => 'b::type => bool) d::'a::type metric.
-   dorder g -->
-   tends (x::'b::type => 'a::type) (x0::'a::type) (mtop d, g) &
-   tends x (x1::'a::type) (mtop d, g) -->
-   x0 = x1"
+lemma MTOP_TENDS_UNIQ: "[| dorder (g::'b => 'b => bool);
+   tends (x::'b => 'a) (x0::'a) (mtop (d::'a metric), g) &
+   tends x (x1::'a) (mtop d, g) |]
+==> x0 = x1"
   by (import nets MTOP_TENDS_UNIQ)
 
-lemma SEQ_TENDS: "ALL (d::'a::type metric) (x::nat => 'a::type) x0::'a::type.
-   tends x x0 (mtop d, nat_ge) =
-   (ALL xa>0. EX xb::nat. ALL xc::nat. xb <= xc --> dist d (x xc, x0) < xa)"
+lemma SEQ_TENDS: "tends x x0 (mtop d, nat_ge) =
+(ALL xa>0. EX xb. ALL xc>=xb. topology.dist d (x xc, x0) < xa)"
   by (import nets SEQ_TENDS)
 
-lemma LIM_TENDS: "ALL (m1::'a::type metric) (m2::'b::type metric) (f::'a::type => 'b::type)
-   (x0::'a::type) y0::'b::type.
-   limpt (mtop m1) x0 re_universe -->
-   tends f y0 (mtop m2, tendsto (m1, x0)) =
-   (ALL e>0.
-       EX d>0.
-          ALL x::'a::type.
-             0 < dist m1 (x, x0) & dist m1 (x, x0) <= d -->
-             dist m2 (f x, y0) < e)"
+lemma LIM_TENDS: "limpt (mtop m1) x0 re_universe
+==> tends f y0 (mtop m2, nets.tendsto (m1, x0)) =
+    (ALL e>0.
+        EX d>0.
+           ALL x.
+              0 < topology.dist m1 (x, x0) &
+              topology.dist m1 (x, x0) <= d -->
+              topology.dist m2 (f x, y0) < e)"
   by (import nets LIM_TENDS)
 
-lemma LIM_TENDS2: "ALL (m1::'a::type metric) (m2::'b::type metric) (f::'a::type => 'b::type)
-   (x0::'a::type) y0::'b::type.
-   limpt (mtop m1) x0 re_universe -->
-   tends f y0 (mtop m2, tendsto (m1, x0)) =
-   (ALL e>0.
-       EX d>0.
-          ALL x::'a::type.
-             0 < dist m1 (x, x0) & dist m1 (x, x0) < d -->
-             dist m2 (f x, y0) < e)"
+lemma LIM_TENDS2: "limpt (mtop m1) x0 re_universe
+==> tends f y0 (mtop m2, nets.tendsto (m1, x0)) =
+    (ALL e>0.
+        EX d>0.
+           ALL x.
+              0 < topology.dist m1 (x, x0) &
+              topology.dist m1 (x, x0) < d -->
+              topology.dist m2 (f x, y0) < e)"
   by (import nets LIM_TENDS2)
 
-lemma MR1_BOUNDED: "ALL (g::'a::type => 'a::type => bool) f::'a::type => real.
-   bounded (mr1, g) f =
-   (EX (k::real) N::'a::type.
-       g N N & (ALL n::'a::type. g n N --> abs (f n) < k))"
+lemma MR1_BOUNDED: "bounded (mr1, g) f = (EX k N. g N N & (ALL n. g n N --> abs (f n) < k))"
   by (import nets MR1_BOUNDED)
 
-lemma NET_NULL: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real.
-   tends x x0 (mtop mr1, g) = tends (%n::'a::type. x n - x0) 0 (mtop mr1, g)"
+lemma NET_NULL: "tends x x0 (mtop mr1, g) = tends (%n. x n - x0) 0 (mtop mr1, g)"
   by (import nets NET_NULL)
 
-lemma NET_CONV_BOUNDED: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real.
-   tends x x0 (mtop mr1, g) --> bounded (mr1, g) x"
+lemma NET_CONV_BOUNDED: "tends x x0 (mtop mr1, g) ==> bounded (mr1, g) x"
   by (import nets NET_CONV_BOUNDED)
 
-lemma NET_CONV_NZ: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real.
-   tends x x0 (mtop mr1, g) & x0 ~= 0 -->
-   (EX N::'a::type. g N N & (ALL n::'a::type. g n N --> x n ~= 0))"
+lemma NET_CONV_NZ: "tends x x0 (mtop mr1, g) & x0 ~= 0
+==> EX N. g N N & (ALL n. g n N --> x n ~= 0)"
   by (import nets NET_CONV_NZ)
 
-lemma NET_CONV_IBOUNDED: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real.
-   tends x x0 (mtop mr1, g) & x0 ~= 0 -->
-   bounded (mr1, g) (%n::'a::type. inverse (x n))"
+lemma NET_CONV_IBOUNDED: "tends x x0 (mtop mr1, g) & x0 ~= 0 ==> bounded (mr1, g) (%n. inverse (x n))"
   by (import nets NET_CONV_IBOUNDED)
 
-lemma NET_NULL_ADD: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) y::'a::type => real.
-       tends x 0 (mtop mr1, g) & tends y 0 (mtop mr1, g) -->
-       tends (%n::'a::type. x n + y n) 0 (mtop mr1, g))"
+lemma NET_NULL_ADD: "[| dorder g; tends x 0 (mtop mr1, g) & tends y 0 (mtop mr1, g) |]
+==> tends (%n. x n + y n) 0 (mtop mr1, g)"
   by (import nets NET_NULL_ADD)
 
-lemma NET_NULL_MUL: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) y::'a::type => real.
-       bounded (mr1, g) x & tends y 0 (mtop mr1, g) -->
-       tends (%n::'a::type. x n * y n) 0 (mtop mr1, g))"
+lemma NET_NULL_MUL: "[| dorder g; bounded (mr1, g) x & tends y 0 (mtop mr1, g) |]
+==> tends (%n. x n * y n) 0 (mtop mr1, g)"
   by (import nets NET_NULL_MUL)
 
-lemma NET_NULL_CMUL: "ALL (g::'a::type => 'a::type => bool) (k::real) x::'a::type => real.
-   tends x 0 (mtop mr1, g) --> tends (%n::'a::type. k * x n) 0 (mtop mr1, g)"
+lemma NET_NULL_CMUL: "tends x 0 (mtop mr1, g) ==> tends (%n. k * x n) 0 (mtop mr1, g)"
   by (import nets NET_NULL_CMUL)
 
-lemma NET_ADD: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real.
-       tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) -->
-       tends (%n::'a::type. x n + y n) (x0 + y0) (mtop mr1, g))"
+lemma NET_ADD: "[| dorder g; tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) |]
+==> tends (%n. x n + y n) (x0 + y0) (mtop mr1, g)"
   by (import nets NET_ADD)
 
-lemma NET_NEG: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) x0::real.
-       tends x x0 (mtop mr1, g) =
-       tends (%n::'a::type. - x n) (- x0) (mtop mr1, g))"
+lemma NET_NEG: "dorder g
+==> tends x x0 (mtop mr1, g) = tends (%n. - x n) (- x0) (mtop mr1, g)"
   by (import nets NET_NEG)
 
-lemma NET_SUB: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real.
-       tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) -->
-       tends (%xa::'a::type. x xa - y xa) (x0 - y0) (mtop mr1, g))"
+lemma NET_SUB: "[| dorder g; tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) |]
+==> tends (%xa. x xa - y xa) (x0 - y0) (mtop mr1, g)"
   by (import nets NET_SUB)
 
-lemma NET_MUL: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) (y::'a::type => real) (x0::real) y0::real.
-       tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) -->
-       tends (%n::'a::type. x n * y n) (x0 * y0) (mtop mr1, g))"
+lemma NET_MUL: "[| dorder g; tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) |]
+==> tends (%n. x n * y n) (x0 * y0) (mtop mr1, g)"
   by (import nets NET_MUL)
 
-lemma NET_INV: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) x0::real.
-       tends x x0 (mtop mr1, g) & x0 ~= 0 -->
-       tends (%n::'a::type. inverse (x n)) (inverse x0) (mtop mr1, g))"
+lemma NET_INV: "[| dorder g; tends x x0 (mtop mr1, g) & x0 ~= 0 |]
+==> tends (%n. inverse (x n)) (inverse x0) (mtop mr1, g)"
   by (import nets NET_INV)
 
-lemma NET_DIV: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real.
-       tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) & y0 ~= 0 -->
-       tends (%xa::'a::type. x xa / y xa) (x0 / y0) (mtop mr1, g))"
+lemma NET_DIV: "[| dorder g;
+   tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) & y0 ~= 0 |]
+==> tends (%xa. x xa / y xa) (x0 / y0) (mtop mr1, g)"
   by (import nets NET_DIV)
 
-lemma NET_ABS: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real.
-   tends x x0 (mtop mr1, g) -->
-   tends (%n::'a::type. abs (x n)) (abs x0) (mtop mr1, g)"
+lemma NET_ABS: "tends x x0 (mtop mr1, g) ==> tends (%n. abs (x n)) (abs x0) (mtop mr1, g)"
   by (import nets NET_ABS)
 
-lemma NET_LE: "ALL g::'a::type => 'a::type => bool.
-   dorder g -->
-   (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real.
-       tends x x0 (mtop mr1, g) &
-       tends y y0 (mtop mr1, g) &
-       (EX N::'a::type. g N N & (ALL n::'a::type. g n N --> x n <= y n)) -->
-       x0 <= y0)"
+lemma NET_LE: "[| dorder g;
+   tends x x0 (mtop mr1, g) &
+   tends y y0 (mtop mr1, g) &
+   (EX N. g N N & (ALL n. g n N --> x n <= y n)) |]
+==> x0 <= y0"
   by (import nets NET_LE)
 
 ;end_setup
@@ -1294,84 +1063,77 @@
   "hol4-->" :: "(nat => real) => real => bool" ("hol4-->")
 
 defs
-  "hol4-->_def": "hol4--> == %(x::nat => real) x0::real. tends x x0 (mtop mr1, nat_ge)"
-
-lemma tends_num_real: "ALL (x::nat => real) x0::real. hol4--> x x0 = tends x x0 (mtop mr1, nat_ge)"
+  "hol4-->_def": "hol4--> == %x x0. tends x x0 (mtop mr1, nat_ge)"
+
+lemma tends_num_real: "hol4--> x x0 = tends x x0 (mtop mr1, nat_ge)"
   by (import seq tends_num_real)
 
-lemma SEQ: "ALL (x::nat => real) x0::real.
-   hol4--> x x0 =
-   (ALL e>0. EX N::nat. ALL n::nat. N <= n --> abs (x n - x0) < e)"
+lemma SEQ: "hol4--> x x0 = (ALL e>0. EX N. ALL n>=N. abs (x n - x0) < e)"
   by (import seq SEQ)
 
-lemma SEQ_CONST: "ALL k::real. hol4--> (%x::nat. k) k"
+lemma SEQ_CONST: "hol4--> (%x. k) k"
   by (import seq SEQ_CONST)
 
-lemma SEQ_ADD: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   hol4--> x x0 & hol4--> y y0 --> hol4--> (%n::nat. x n + y n) (x0 + y0)"
+lemma SEQ_ADD: "hol4--> x x0 & hol4--> y y0 ==> hol4--> (%n. x n + y n) (x0 + y0)"
   by (import seq SEQ_ADD)
 
-lemma SEQ_MUL: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   hol4--> x x0 & hol4--> y y0 --> hol4--> (%n::nat. x n * y n) (x0 * y0)"
+lemma SEQ_MUL: "hol4--> x x0 & hol4--> y y0 ==> hol4--> (%n. x n * y n) (x0 * y0)"
   by (import seq SEQ_MUL)
 
-lemma SEQ_NEG: "ALL (x::nat => real) x0::real.
-   hol4--> x x0 = hol4--> (%n::nat. - x n) (- x0)"
+lemma SEQ_NEG: "hol4--> x x0 = hol4--> (%n. - x n) (- x0)"
   by (import seq SEQ_NEG)
 
-lemma SEQ_INV: "ALL (x::nat => real) x0::real.
-   hol4--> x x0 & x0 ~= 0 --> hol4--> (%n::nat. inverse (x n)) (inverse x0)"
+lemma SEQ_INV: "hol4--> x x0 & x0 ~= 0 ==> hol4--> (%n. inverse (x n)) (inverse x0)"
   by (import seq SEQ_INV)
 
-lemma SEQ_SUB: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   hol4--> x x0 & hol4--> y y0 --> hol4--> (%n::nat. x n - y n) (x0 - y0)"
+lemma SEQ_SUB: "hol4--> x x0 & hol4--> y y0 ==> hol4--> (%n. x n - y n) (x0 - y0)"
   by (import seq SEQ_SUB)
 
-lemma SEQ_DIV: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   hol4--> x x0 & hol4--> y y0 & y0 ~= 0 -->
-   hol4--> (%n::nat. x n / y n) (x0 / y0)"
+lemma SEQ_DIV: "hol4--> x x0 & hol4--> y y0 & y0 ~= 0 ==> hol4--> (%n. x n / y n) (x0 / y0)"
   by (import seq SEQ_DIV)
 
-lemma SEQ_UNIQ: "ALL (x::nat => real) (x1::real) x2::real.
-   hol4--> x x1 & hol4--> x x2 --> x1 = x2"
+lemma SEQ_UNIQ: "hol4--> x x1 & hol4--> x x2 ==> x1 = x2"
   by (import seq SEQ_UNIQ)
 
-definition convergent :: "(nat => real) => bool" where 
-  "convergent == %f::nat => real. Ex (hol4--> f)"
-
-lemma convergent: "ALL f::nat => real. convergent f = Ex (hol4--> f)"
+consts
+  convergent :: "(nat => real) => bool" 
+
+defs
+  convergent_def: "seq.convergent == %f. Ex (hol4--> f)"
+
+lemma convergent: "seq.convergent f = Ex (hol4--> f)"
   by (import seq convergent)
 
-definition cauchy :: "(nat => real) => bool" where 
+definition
+  cauchy :: "(nat => real) => bool"  where
   "cauchy ==
-%f::nat => real.
-   ALL e>0.
-      EX N::nat.
-         ALL (m::nat) n::nat. N <= m & N <= n --> abs (f m - f n) < e"
-
-lemma cauchy: "ALL f::nat => real.
-   cauchy f =
-   (ALL e>0.
-       EX N::nat.
-          ALL (m::nat) n::nat. N <= m & N <= n --> abs (f m - f n) < e)"
+%f. ALL e>0. EX N. ALL m n. N <= m & N <= n --> abs (f m - f n) < e"
+
+lemma cauchy: "cauchy f = (ALL e>0. EX N. ALL m n. N <= m & N <= n --> abs (f m - f n) < e)"
   by (import seq cauchy)
 
-definition lim :: "(nat => real) => real" where 
-  "lim == %f::nat => real. Eps (hol4--> f)"
-
-lemma lim: "ALL f::nat => real. lim f = Eps (hol4--> f)"
+consts
+  lim :: "(nat => real) => real" 
+
+defs
+  lim_def: "seq.lim == %f. Eps (hol4--> f)"
+
+lemma lim: "seq.lim f = Eps (hol4--> f)"
   by (import seq lim)
 
-lemma SEQ_LIM: "ALL f::nat => real. convergent f = hol4--> f (lim f)"
+lemma SEQ_LIM: "seq.convergent f = hol4--> f (seq.lim f)"
   by (import seq SEQ_LIM)
 
-definition subseq :: "(nat => nat) => bool" where 
-  "subseq == %f::nat => nat. ALL (m::nat) n::nat. m < n --> f m < f n"
-
-lemma subseq: "ALL f::nat => nat. subseq f = (ALL (m::nat) n::nat. m < n --> f m < f n)"
+consts
+  subseq :: "(nat => nat) => bool" 
+
+defs
+  subseq_def: "seq.subseq == %f. ALL m n. m < n --> f m < f n"
+
+lemma subseq: "seq.subseq f = (ALL m n. m < n --> f m < f n)"
   by (import seq subseq)
 
-lemma SUBSEQ_SUC: "ALL f::nat => nat. subseq f = (ALL n::nat. f n < f (Suc n))"
+lemma SUBSEQ_SUC: "seq.subseq f = (ALL n. f n < f (Suc n))"
   by (import seq SUBSEQ_SUC)
 
 consts
@@ -1379,1818 +1141,1414 @@
 
 defs
   mono_def: "seq.mono ==
-%f::nat => real.
-   (ALL (m::nat) n::nat. m <= n --> f m <= f n) |
-   (ALL (m::nat) n::nat. m <= n --> f n <= f m)"
-
-lemma mono: "ALL f::nat => real.
-   seq.mono f =
-   ((ALL (m::nat) n::nat. m <= n --> f m <= f n) |
-    (ALL (m::nat) n::nat. m <= n --> f n <= f m))"
+%f. (ALL m n. m <= n --> f m <= f n) | (ALL m n. m <= n --> f n <= f m)"
+
+lemma mono: "seq.mono f =
+((ALL m n. m <= n --> f m <= f n) | (ALL m n. m <= n --> f n <= f m))"
   by (import seq mono)
 
-lemma MONO_SUC: "ALL f::nat => real.
-   seq.mono f =
-   ((ALL x::nat. f x <= f (Suc x)) | (ALL n::nat. f (Suc n) <= f n))"
+lemma MONO_SUC: "seq.mono f = ((ALL x. f x <= f (Suc x)) | (ALL n. f (Suc n) <= f n))"
   by (import seq MONO_SUC)
 
-lemma MAX_LEMMA: "(All::((nat => real) => bool) => bool)
- (%s::nat => real.
-     (All::(nat => bool) => bool)
-      (%N::nat.
-          (Ex::(real => bool) => bool)
-           (%k::real.
-               (All::(nat => bool) => bool)
-                (%n::nat.
-                    (op -->::bool => bool => bool)
-                     ((op <::nat => nat => bool) n N)
-                     ((op <::real => real => bool)
-                       ((abs::real => real) (s n)) k)))))"
+lemma MAX_LEMMA: "EX k::real. ALL n<N::nat. abs ((s::nat => real) n) < k"
   by (import seq MAX_LEMMA)
 
-lemma SEQ_BOUNDED: "ALL s::nat => real.
-   bounded (mr1, nat_ge) s = (EX k::real. ALL n::nat. abs (s n) < k)"
+lemma SEQ_BOUNDED: "bounded (mr1, nat_ge) s = (EX k. ALL n. abs (s n) < k)"
   by (import seq SEQ_BOUNDED)
 
-lemma SEQ_BOUNDED_2: "ALL (f::nat => real) (k::real) k'::real.
-   (ALL n::nat. k <= f n & f n <= k') --> bounded (mr1, nat_ge) f"
+lemma SEQ_BOUNDED_2: "(!!n. k <= f n & f n <= k') ==> bounded (mr1, nat_ge) f"
   by (import seq SEQ_BOUNDED_2)
 
-lemma SEQ_CBOUNDED: "ALL f::nat => real. cauchy f --> bounded (mr1, nat_ge) f"
+lemma SEQ_CBOUNDED: "cauchy f ==> bounded (mr1, nat_ge) f"
   by (import seq SEQ_CBOUNDED)
 
-lemma SEQ_ICONV: "ALL f::nat => real.
-   bounded (mr1, nat_ge) f &
-   (ALL (m::nat) n::nat. n <= m --> f n <= f m) -->
-   convergent f"
+lemma SEQ_ICONV: "bounded (mr1, nat_ge) f & (ALL m n. n <= m --> f n <= f m)
+==> seq.convergent f"
   by (import seq SEQ_ICONV)
 
-lemma SEQ_NEG_CONV: "ALL f::nat => real. convergent f = convergent (%n::nat. - f n)"
+lemma SEQ_NEG_CONV: "seq.convergent f = seq.convergent (%n. - f n)"
   by (import seq SEQ_NEG_CONV)
 
-lemma SEQ_NEG_BOUNDED: "ALL f::nat => real.
-   bounded (mr1, nat_ge) (%n::nat. - f n) = bounded (mr1, nat_ge) f"
+lemma SEQ_NEG_BOUNDED: "bounded (mr1, nat_ge) (%n. - f n) = bounded (mr1, nat_ge) f"
   by (import seq SEQ_NEG_BOUNDED)
 
-lemma SEQ_BCONV: "ALL f::nat => real. bounded (mr1, nat_ge) f & seq.mono f --> convergent f"
+lemma SEQ_BCONV: "bounded (mr1, nat_ge) f & seq.mono f ==> seq.convergent f"
   by (import seq SEQ_BCONV)
 
-lemma SEQ_MONOSUB: "ALL s::nat => real. EX f::nat => nat. subseq f & seq.mono (%n::nat. s (f n))"
+lemma SEQ_MONOSUB: "EX f. seq.subseq f & seq.mono (%n. s (f n))"
   by (import seq SEQ_MONOSUB)
 
-lemma SEQ_SBOUNDED: "ALL (s::nat => real) f::nat => nat.
-   bounded (mr1, nat_ge) s --> bounded (mr1, nat_ge) (%n::nat. s (f n))"
+lemma SEQ_SBOUNDED: "bounded (mr1, nat_ge) s ==> bounded (mr1, nat_ge) (%n. s (f n))"
   by (import seq SEQ_SBOUNDED)
 
-lemma SEQ_SUBLE: "ALL f::nat => nat. subseq f --> (ALL n::nat. n <= f n)"
+lemma SEQ_SUBLE: "seq.subseq f ==> n <= f n"
   by (import seq SEQ_SUBLE)
 
-lemma SEQ_DIRECT: "ALL f::nat => nat.
-   subseq f --> (ALL (N1::nat) N2::nat. EX x::nat. N1 <= x & N2 <= f x)"
+lemma SEQ_DIRECT: "seq.subseq f ==> EX x>=N1. N2 <= f x"
   by (import seq SEQ_DIRECT)
 
-lemma SEQ_CAUCHY: "ALL f::nat => real. cauchy f = convergent f"
+lemma SEQ_CAUCHY: "cauchy f = seq.convergent f"
   by (import seq SEQ_CAUCHY)
 
-lemma SEQ_LE: "ALL (f::nat => real) (g::nat => real) (l::real) m::real.
-   hol4--> f l &
-   hol4--> g m & (EX x::nat. ALL xa::nat. x <= xa --> f xa <= g xa) -->
-   l <= m"
+lemma SEQ_LE: "hol4--> f l & hol4--> g m & (EX x. ALL xa>=x. f xa <= g xa) ==> l <= m"
   by (import seq SEQ_LE)
 
-lemma SEQ_SUC: "ALL (f::nat => real) l::real. hol4--> f l = hol4--> (%n::nat. f (Suc n)) l"
+lemma SEQ_SUC: "hol4--> f l = hol4--> (%n. f (Suc n)) l"
   by (import seq SEQ_SUC)
 
-lemma SEQ_ABS: "ALL f::nat => real. hol4--> (%n::nat. abs (f n)) 0 = hol4--> f 0"
+lemma SEQ_ABS: "hol4--> (%n. abs (f n)) 0 = hol4--> f 0"
   by (import seq SEQ_ABS)
 
-lemma SEQ_ABS_IMP: "ALL (f::nat => real) l::real.
-   hol4--> f l --> hol4--> (%n::nat. abs (f n)) (abs l)"
+lemma SEQ_ABS_IMP: "hol4--> f l ==> hol4--> (%n. abs (f n)) (abs l)"
   by (import seq SEQ_ABS_IMP)
 
-lemma SEQ_INV0: "ALL f::nat => real.
-   (ALL y::real. EX N::nat. ALL n::nat. N <= n --> y < f n) -->
-   hol4--> (%n::nat. inverse (f n)) 0"
+lemma SEQ_INV0: "(!!y. EX N. ALL n>=N. y < f n) ==> hol4--> (%n. inverse (f n)) 0"
   by (import seq SEQ_INV0)
 
-lemma SEQ_POWER_ABS: "ALL c::real. abs c < 1 --> hol4--> (op ^ (abs c)) 0"
+lemma SEQ_POWER_ABS: "abs c < 1 ==> hol4--> (op ^ (abs c)) 0"
   by (import seq SEQ_POWER_ABS)
 
-lemma SEQ_POWER: "ALL c::real. abs c < 1 --> hol4--> (op ^ c) 0"
+lemma SEQ_POWER: "abs c < 1 ==> hol4--> (op ^ c) 0"
   by (import seq SEQ_POWER)
 
-lemma NEST_LEMMA: "ALL (f::nat => real) g::nat => real.
-   (ALL n::nat. f n <= f (Suc n)) &
-   (ALL n::nat. g (Suc n) <= g n) & (ALL n::nat. f n <= g n) -->
-   (EX (l::real) m::real.
+lemma NEST_LEMMA: "(ALL n. f n <= f (Suc n)) & (ALL n. g (Suc n) <= g n) & (ALL n. f n <= g n)
+==> EX l m.
        l <= m &
-       ((ALL n::nat. f n <= l) & hol4--> f l) &
-       (ALL n::nat. m <= g n) & hol4--> g m)"
+       ((ALL n. f n <= l) & hol4--> f l) & (ALL n. m <= g n) & hol4--> g m"
   by (import seq NEST_LEMMA)
 
-lemma NEST_LEMMA_UNIQ: "ALL (f::nat => real) g::nat => real.
-   (ALL n::nat. f n <= f (Suc n)) &
-   (ALL n::nat. g (Suc n) <= g n) &
-   (ALL n::nat. f n <= g n) & hol4--> (%n::nat. f n - g n) 0 -->
-   (EX x::real.
-       ((ALL n::nat. f n <= x) & hol4--> f x) &
-       (ALL n::nat. x <= g n) & hol4--> g x)"
+lemma NEST_LEMMA_UNIQ: "(ALL n. f n <= f (Suc n)) &
+(ALL n. g (Suc n) <= g n) & (ALL n. f n <= g n) & hol4--> (%n. f n - g n) 0
+==> EX x. ((ALL n. f n <= x) & hol4--> f x) &
+          (ALL n. x <= g n) & hol4--> g x"
   by (import seq NEST_LEMMA_UNIQ)
 
-lemma BOLZANO_LEMMA: "ALL P::real * real => bool.
-   (ALL (a::real) (b::real) c::real.
-       a <= b & b <= c & P (a, b) & P (b, c) --> P (a, c)) &
-   (ALL x::real.
-       EX d>0.
-          ALL (a::real) b::real.
-             a <= x & x <= b & b - a < d --> P (a, b)) -->
-   (ALL (a::real) b::real. a <= b --> P (a, b))"
-  by (import seq BOLZANO_LEMMA)
-
-definition sums :: "(nat => real) => real => bool" where 
-  "sums == %f::nat => real. hol4--> (%n::nat. real.sum (0, n) f)"
-
-lemma sums: "ALL (f::nat => real) s::real.
-   sums f s = hol4--> (%n::nat. real.sum (0, n) f) s"
+consts
+  sums :: "(nat => real) => real => bool" 
+
+defs
+  sums_def: "seq.sums == %f. hol4--> (%n. real.sum (0, n) f)"
+
+lemma sums: "seq.sums f s = hol4--> (%n. real.sum (0, n) f) s"
   by (import seq sums)
 
-definition summable :: "(nat => real) => bool" where 
-  "summable == %f::nat => real. Ex (sums f)"
-
-lemma summable: "ALL f::nat => real. summable f = Ex (sums f)"
+consts
+  summable :: "(nat => real) => bool" 
+
+defs
+  summable_def: "seq.summable == %f. Ex (seq.sums f)"
+
+lemma summable: "seq.summable f = Ex (seq.sums f)"
   by (import seq summable)
 
-definition suminf :: "(nat => real) => real" where 
-  "suminf == %f::nat => real. Eps (sums f)"
-
-lemma suminf: "ALL f::nat => real. suminf f = Eps (sums f)"
+consts
+  suminf :: "(nat => real) => real" 
+
+defs
+  suminf_def: "seq.suminf == %f. Eps (seq.sums f)"
+
+lemma suminf: "seq.suminf f = Eps (seq.sums f)"
   by (import seq suminf)
 
-lemma SUM_SUMMABLE: "ALL (f::nat => real) l::real. sums f l --> summable f"
+lemma SUM_SUMMABLE: "seq.sums f l ==> seq.summable f"
   by (import seq SUM_SUMMABLE)
 
-lemma SUMMABLE_SUM: "ALL f::nat => real. summable f --> sums f (suminf f)"
+lemma SUMMABLE_SUM: "seq.summable f ==> seq.sums f (seq.suminf f)"
   by (import seq SUMMABLE_SUM)
 
-lemma SUM_UNIQ: "ALL (f::nat => real) x::real. sums f x --> x = suminf f"
+lemma SUM_UNIQ: "seq.sums f x ==> x = seq.suminf f"
   by (import seq SUM_UNIQ)
 
-lemma SER_0: "ALL (f::nat => real) n::nat.
-   (ALL m::nat. n <= m --> f m = 0) --> sums f (real.sum (0, n) f)"
+lemma SER_0: "(!!m. n <= m ==> f m = 0) ==> seq.sums f (real.sum (0, n) f)"
   by (import seq SER_0)
 
-lemma SER_POS_LE: "ALL (f::nat => real) n::nat.
-   summable f & (ALL m::nat. n <= m --> 0 <= f m) -->
-   real.sum (0, n) f <= suminf f"
+lemma SER_POS_LE: "seq.summable f & (ALL m>=n. 0 <= f m) ==> real.sum (0, n) f <= seq.suminf f"
   by (import seq SER_POS_LE)
 
-lemma SER_POS_LT: "ALL (f::nat => real) n::nat.
-   summable f & (ALL m::nat. n <= m --> 0 < f m) -->
-   real.sum (0, n) f < suminf f"
+lemma SER_POS_LT: "seq.summable f & (ALL m>=n. 0 < f m) ==> real.sum (0, n) f < seq.suminf f"
   by (import seq SER_POS_LT)
 
-lemma SER_GROUP: "ALL (f::nat => real) k::nat.
-   summable f & 0 < k --> sums (%n::nat. real.sum (n * k, k) f) (suminf f)"
+lemma SER_GROUP: "seq.summable f & 0 < k
+==> seq.sums (%n. real.sum (n * k, k) f) (seq.suminf f)"
   by (import seq SER_GROUP)
 
-lemma SER_PAIR: "ALL f::nat => real.
-   summable f --> sums (%n::nat. real.sum (2 * n, 2) f) (suminf f)"
+lemma SER_PAIR: "seq.summable f ==> seq.sums (%n. real.sum (2 * n, 2) f) (seq.suminf f)"
   by (import seq SER_PAIR)
 
-lemma SER_OFFSET: "ALL f::nat => real.
-   summable f -->
-   (ALL k::nat. sums (%n::nat. f (n + k)) (suminf f - real.sum (0, k) f))"
+lemma SER_OFFSET: "seq.summable f
+==> seq.sums (%n. f (n + k)) (seq.suminf f - real.sum (0, k) f)"
   by (import seq SER_OFFSET)
 
-lemma SER_POS_LT_PAIR: "ALL (f::nat => real) n::nat.
-   summable f & (ALL d::nat. 0 < f (n + 2 * d) + f (n + (2 * d + 1))) -->
-   real.sum (0, n) f < suminf f"
+lemma SER_POS_LT_PAIR: "seq.summable f & (ALL d. 0 < f (n + 2 * d) + f (n + (2 * d + 1)))
+==> real.sum (0, n) f < seq.suminf f"
   by (import seq SER_POS_LT_PAIR)
 
-lemma SER_ADD: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   sums x x0 & sums y y0 --> sums (%n::nat. x n + y n) (x0 + y0)"
+lemma SER_ADD: "seq.sums x x0 & seq.sums y y0 ==> seq.sums (%n. x n + y n) (x0 + y0)"
   by (import seq SER_ADD)
 
-lemma SER_CMUL: "ALL (x::nat => real) (x0::real) c::real.
-   sums x x0 --> sums (%n::nat. c * x n) (c * x0)"
+lemma SER_CMUL: "seq.sums x x0 ==> seq.sums (%n. c * x n) (c * x0)"
   by (import seq SER_CMUL)
 
-lemma SER_NEG: "ALL (x::nat => real) x0::real. sums x x0 --> sums (%xa::nat. - x xa) (- x0)"
+lemma SER_NEG: "seq.sums x x0 ==> seq.sums (%xa. - x xa) (- x0)"
   by (import seq SER_NEG)
 
-lemma SER_SUB: "ALL (x::nat => real) (x0::real) (y::nat => real) y0::real.
-   sums x x0 & sums y y0 --> sums (%xa::nat. x xa - y xa) (x0 - y0)"
+lemma SER_SUB: "seq.sums x x0 & seq.sums y y0 ==> seq.sums (%xa. x xa - y xa) (x0 - y0)"
   by (import seq SER_SUB)
 
-lemma SER_CDIV: "ALL (x::nat => real) (x0::real) c::real.
-   sums x x0 --> sums (%xa::nat. x xa / c) (x0 / c)"
+lemma SER_CDIV: "seq.sums x x0 ==> seq.sums (%xa. x xa / c) (x0 / c)"
   by (import seq SER_CDIV)
 
-lemma SER_CAUCHY: "ALL f::nat => real.
-   summable f =
-   (ALL e>0.
-       EX N::nat.
-          ALL (m::nat) n::nat. N <= m --> abs (real.sum (m, n) f) < e)"
+lemma SER_CAUCHY: "seq.summable f =
+(ALL e>0. EX N. ALL m n. N <= m --> abs (real.sum (m, n) f) < e)"
   by (import seq SER_CAUCHY)
 
-lemma SER_ZERO: "ALL f::nat => real. summable f --> hol4--> f 0"
+lemma SER_ZERO: "seq.summable f ==> hol4--> f 0"
   by (import seq SER_ZERO)
 
-lemma SER_COMPAR: "ALL (f::nat => real) g::nat => real.
-   (EX x::nat. ALL xa::nat. x <= xa --> abs (f xa) <= g xa) & summable g -->
-   summable f"
+lemma SER_COMPAR: "(EX x. ALL xa>=x. abs (f xa) <= g xa) & seq.summable g ==> seq.summable f"
   by (import seq SER_COMPAR)
 
-lemma SER_COMPARA: "ALL (f::nat => real) g::nat => real.
-   (EX x::nat. ALL xa::nat. x <= xa --> abs (f xa) <= g xa) & summable g -->
-   summable (%k::nat. abs (f k))"
+lemma SER_COMPARA: "(EX x. ALL xa>=x. abs (f xa) <= g xa) & seq.summable g
+==> seq.summable (%k. abs (f k))"
   by (import seq SER_COMPARA)
 
-lemma SER_LE: "ALL (f::nat => real) g::nat => real.
-   (ALL n::nat. f n <= g n) & summable f & summable g -->
-   suminf f <= suminf g"
+lemma SER_LE: "(ALL n. f n <= g n) & seq.summable f & seq.summable g
+==> seq.suminf f <= seq.suminf g"
   by (import seq SER_LE)
 
-lemma SER_LE2: "ALL (f::nat => real) g::nat => real.
-   (ALL n::nat. abs (f n) <= g n) & summable g -->
-   summable f & suminf f <= suminf g"
+lemma SER_LE2: "(ALL n. abs (f n) <= g n) & seq.summable g
+==> seq.summable f & seq.suminf f <= seq.suminf g"
   by (import seq SER_LE2)
 
-lemma SER_ACONV: "ALL f::nat => real. summable (%n::nat. abs (f n)) --> summable f"
+lemma SER_ACONV: "seq.summable (%n. abs (f n)) ==> seq.summable f"
   by (import seq SER_ACONV)
 
-lemma SER_ABS: "ALL f::nat => real.
-   summable (%n::nat. abs (f n)) -->
-   abs (suminf f) <= suminf (%n::nat. abs (f n))"
+lemma SER_ABS: "seq.summable (%n. abs (f n))
+==> abs (seq.suminf f) <= seq.suminf (%n. abs (f n))"
   by (import seq SER_ABS)
 
-lemma GP_FINITE: "ALL x::real.
-   x ~= 1 --> (ALL n::nat. real.sum (0, n) (op ^ x) = (x ^ n - 1) / (x - 1))"
+lemma GP_FINITE: "x ~= 1 ==> real.sum (0, n) (op ^ x) = (x ^ n - 1) / (x - 1)"
   by (import seq GP_FINITE)
 
-lemma GP: "ALL x::real. abs x < 1 --> sums (op ^ x) (inverse (1 - x))"
+lemma GP: "abs x < 1 ==> seq.sums (op ^ x) (inverse (1 - x))"
   by (import seq GP)
 
-lemma ABS_NEG_LEMMA: "(All::(real => bool) => bool)
- (%c::real.
-     (op -->::bool => bool => bool)
-      ((op <=::real => real => bool) c (0::real))
-      ((All::(real => bool) => bool)
-        (%x::real.
-            (All::(real => bool) => bool)
-             (%y::real.
-                 (op -->::bool => bool => bool)
-                  ((op <=::real => real => bool) ((abs::real => real) x)
-                    ((op *::real => real => real) c
-                      ((abs::real => real) y)))
-                  ((op =::real => real => bool) x (0::real))))))"
-  by (import seq ABS_NEG_LEMMA)
-
-lemma SER_RATIO: "ALL (f::nat => real) (c::real) N::nat.
-   c < 1 & (ALL n::nat. N <= n --> abs (f (Suc n)) <= c * abs (f n)) -->
-   summable f"
+lemma SER_RATIO: "c < 1 & (ALL n>=N. abs (f (Suc n)) <= c * abs (f n)) ==> seq.summable f"
   by (import seq SER_RATIO)
 
 ;end_setup
 
 ;setup_theory lim
 
-definition tends_real_real :: "(real => real) => real => real => bool" where 
-  "tends_real_real ==
-%(f::real => real) (l::real) x0::real.
-   tends f l (mtop mr1, tendsto (mr1, x0))"
-
-lemma tends_real_real: "ALL (f::real => real) (l::real) x0::real.
-   tends_real_real f l x0 = tends f l (mtop mr1, tendsto (mr1, x0))"
+definition
+  tends_real_real :: "(real => real) => real => real => bool"  where
+  "tends_real_real == %f l x0. tends f l (mtop mr1, nets.tendsto (mr1, x0))"
+
+lemma tends_real_real: "tends_real_real f l x0 = tends f l (mtop mr1, nets.tendsto (mr1, x0))"
   by (import lim tends_real_real)
 
-lemma LIM: "ALL (f::real => real) (y0::real) x0::real.
-   tends_real_real f y0 x0 =
-   (ALL e>0.
-       EX d>0.
-          ALL x::real.
-             0 < abs (x - x0) & abs (x - x0) < d --> abs (f x - y0) < e)"
+lemma LIM: "tends_real_real f y0 x0 =
+(ALL e>0.
+    EX d>0.
+       ALL x. 0 < abs (x - x0) & abs (x - x0) < d --> abs (f x - y0) < e)"
   by (import lim LIM)
 
-lemma LIM_CONST: "ALL k::real. All (tends_real_real (%x::real. k) k)"
+lemma LIM_CONST: "tends_real_real (%x. k) k x"
   by (import lim LIM_CONST)
 
-lemma LIM_ADD: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   tends_real_real f l x & tends_real_real g m x -->
-   tends_real_real (%x::real. f x + g x) (l + m) x"
+lemma LIM_ADD: "tends_real_real f l x & tends_real_real g m x
+==> tends_real_real (%x. f x + g x) (l + m) x"
   by (import lim LIM_ADD)
 
-lemma LIM_MUL: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   tends_real_real f l x & tends_real_real g m x -->
-   tends_real_real (%x::real. f x * g x) (l * m) x"
+lemma LIM_MUL: "tends_real_real f l x & tends_real_real g m x
+==> tends_real_real (%x. f x * g x) (l * m) x"
   by (import lim LIM_MUL)
 
-lemma LIM_NEG: "ALL (f::real => real) (l::real) x::real.
-   tends_real_real f l x = tends_real_real (%x::real. - f x) (- l) x"
+lemma LIM_NEG: "tends_real_real f l x = tends_real_real (%x. - f x) (- l) x"
   by (import lim LIM_NEG)
 
-lemma LIM_INV: "ALL (f::real => real) (l::real) x::real.
-   tends_real_real f l x & l ~= 0 -->
-   tends_real_real (%x::real. inverse (f x)) (inverse l) x"
+lemma LIM_INV: "tends_real_real f l x & l ~= 0
+==> tends_real_real (%x. inverse (f x)) (inverse l) x"
   by (import lim LIM_INV)
 
-lemma LIM_SUB: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   tends_real_real f l x & tends_real_real g m x -->
-   tends_real_real (%x::real. f x - g x) (l - m) x"
+lemma LIM_SUB: "tends_real_real f l x & tends_real_real g m x
+==> tends_real_real (%x. f x - g x) (l - m) x"
   by (import lim LIM_SUB)
 
-lemma LIM_DIV: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   tends_real_real f l x & tends_real_real g m x & m ~= 0 -->
-   tends_real_real (%x::real. f x / g x) (l / m) x"
+lemma LIM_DIV: "tends_real_real f l x & tends_real_real g m x & m ~= 0
+==> tends_real_real (%x. f x / g x) (l / m) x"
   by (import lim LIM_DIV)
 
-lemma LIM_NULL: "ALL (f::real => real) (l::real) x::real.
-   tends_real_real f l x = tends_real_real (%x::real. f x - l) 0 x"
+lemma LIM_NULL: "tends_real_real f l x = tends_real_real (%x. f x - l) 0 x"
   by (import lim LIM_NULL)
 
-lemma LIM_X: "ALL x0::real. tends_real_real (%x::real. x) x0 x0"
+lemma LIM_X: "tends_real_real (%x. x) x0 x0"
   by (import lim LIM_X)
 
-lemma LIM_UNIQ: "ALL (f::real => real) (l::real) (m::real) x::real.
-   tends_real_real f l x & tends_real_real f m x --> l = m"
+lemma LIM_UNIQ: "tends_real_real f l x & tends_real_real f m x ==> l = m"
   by (import lim LIM_UNIQ)
 
-lemma LIM_EQUAL: "ALL (f::real => real) (g::real => real) (l::real) x0::real.
-   (ALL x::real. x ~= x0 --> f x = g x) -->
-   tends_real_real f l x0 = tends_real_real g l x0"
+lemma LIM_EQUAL: "(!!x. x ~= x0 ==> f x = g x)
+==> tends_real_real f l x0 = tends_real_real g l x0"
   by (import lim LIM_EQUAL)
 
-lemma LIM_TRANSFORM: "ALL (f::real => real) (g::real => real) (x0::real) l::real.
-   tends_real_real (%x::real. f x - g x) 0 x0 & tends_real_real g l x0 -->
-   tends_real_real f l x0"
+lemma LIM_TRANSFORM: "tends_real_real (%x. f x - g x) 0 x0 & tends_real_real g l x0
+==> tends_real_real f l x0"
   by (import lim LIM_TRANSFORM)
 
-definition diffl :: "(real => real) => real => real => bool" where 
-  "diffl ==
-%(f::real => real) (l::real) x::real.
-   tends_real_real (%h::real. (f (x + h) - f x) / h) l 0"
-
-lemma diffl: "ALL (f::real => real) (l::real) x::real.
-   diffl f l x = tends_real_real (%h::real. (f (x + h) - f x) / h) l 0"
+definition
+  diffl :: "(real => real) => real => real => bool"  where
+  "diffl == %f l x. tends_real_real (%h. (f (x + h) - f x) / h) l 0"
+
+lemma diffl: "diffl f l x = tends_real_real (%h. (f (x + h) - f x) / h) l 0"
   by (import lim diffl)
 
-definition contl :: "(real => real) => real => bool" where 
-  "contl ==
-%(f::real => real) x::real. tends_real_real (%h::real. f (x + h)) (f x) 0"
-
-lemma contl: "ALL (f::real => real) x::real.
-   contl f x = tends_real_real (%h::real. f (x + h)) (f x) 0"
+definition
+  contl :: "(real => real) => real => bool"  where
+  "contl == %f x. tends_real_real (%h. f (x + h)) (f x) 0"
+
+lemma contl: "contl f x = tends_real_real (%h. f (x + h)) (f x) 0"
   by (import lim contl)
 
-definition differentiable :: "(real => real) => real => bool" where 
-  "differentiable == %(f::real => real) x::real. EX l::real. diffl f l x"
-
-lemma differentiable: "ALL (f::real => real) x::real.
-   differentiable f x = (EX l::real. diffl f l x)"
+consts
+  differentiable :: "(real => real) => real => bool" 
+
+defs
+  differentiable_def: "lim.differentiable == %f x. EX l. diffl f l x"
+
+lemma differentiable: "lim.differentiable f x = (EX l. diffl f l x)"
   by (import lim differentiable)
 
-lemma DIFF_UNIQ: "ALL (f::real => real) (l::real) (m::real) x::real.
-   diffl f l x & diffl f m x --> l = m"
+lemma DIFF_UNIQ: "diffl f l x & diffl f m x ==> l = m"
   by (import lim DIFF_UNIQ)
 
-lemma DIFF_CONT: "ALL (f::real => real) (l::real) x::real. diffl f l x --> contl f x"
+lemma DIFF_CONT: "diffl f l x ==> contl f x"
   by (import lim DIFF_CONT)
 
-lemma CONTL_LIM: "ALL (f::real => real) x::real. contl f x = tends_real_real f (f x) x"
+lemma CONTL_LIM: "contl f x = tends_real_real f (f x) x"
   by (import lim CONTL_LIM)
 
-lemma DIFF_CARAT: "ALL (f::real => real) (l::real) x::real.
-   diffl f l x =
-   (EX g::real => real.
-       (ALL z::real. f z - f x = g z * (z - x)) & contl g x & g x = l)"
+lemma DIFF_CARAT: "diffl f l x =
+(EX g. (ALL z. f z - f x = g z * (z - x)) & contl g x & g x = l)"
   by (import lim DIFF_CARAT)
 
-lemma CONT_CONST: "ALL k::real. All (contl (%x::real. k))"
+lemma CONT_CONST: "contl (%x. k) x"
   by (import lim CONT_CONST)
 
-lemma CONT_ADD: "ALL (f::real => real) (g::real => real) x::real.
-   contl f x & contl g x --> contl (%x::real. f x + g x) x"
+lemma CONT_ADD: "contl f x & contl g x ==> contl (%x. f x + g x) x"
   by (import lim CONT_ADD)
 
-lemma CONT_MUL: "ALL (f::real => real) (g::real => real) x::real.
-   contl f x & contl g x --> contl (%x::real. f x * g x) x"
+lemma CONT_MUL: "contl f x & contl g x ==> contl (%x. f x * g x) x"
   by (import lim CONT_MUL)
 
-lemma CONT_NEG: "ALL (f::real => real) x::real. contl f x --> contl (%x::real. - f x) x"
+lemma CONT_NEG: "contl f x ==> contl (%x. - f x) x"
   by (import lim CONT_NEG)
 
-lemma CONT_INV: "ALL (f::real => real) x::real.
-   contl f x & f x ~= 0 --> contl (%x::real. inverse (f x)) x"
+lemma CONT_INV: "contl f x & f x ~= 0 ==> contl (%x. inverse (f x)) x"
   by (import lim CONT_INV)
 
-lemma CONT_SUB: "ALL (f::real => real) (g::real => real) x::real.
-   contl f x & contl g x --> contl (%x::real. f x - g x) x"
+lemma CONT_SUB: "contl f x & contl g x ==> contl (%x. f x - g x) x"
   by (import lim CONT_SUB)
 
-lemma CONT_DIV: "ALL (f::real => real) (g::real => real) x::real.
-   contl f x & contl g x & g x ~= 0 --> contl (%x::real. f x / g x) x"
+lemma CONT_DIV: "contl f x & contl g x & g x ~= 0 ==> contl (%x. f x / g x) x"
   by (import lim CONT_DIV)
 
-lemma CONT_COMPOSE: "ALL (f::real => real) (g::real => real) x::real.
-   contl f x & contl g (f x) --> contl (%x::real. g (f x)) x"
+lemma CONT_COMPOSE: "contl f x & contl g (f x) ==> contl (%x. g (f x)) x"
   by (import lim CONT_COMPOSE)
 
-lemma IVT: "ALL (f::real => real) (a::real) (b::real) y::real.
-   a <= b &
-   (f a <= y & y <= f b) & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX x::real. a <= x & x <= b & f x = y)"
+lemma IVT: "a <= b & (f a <= y & y <= f b) & (ALL x. a <= x & x <= b --> contl f x)
+==> EX x>=a. x <= b & f x = y"
   by (import lim IVT)
 
-lemma IVT2: "ALL (f::real => real) (a::real) (b::real) y::real.
-   a <= b &
-   (f b <= y & y <= f a) & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX x::real. a <= x & x <= b & f x = y)"
+lemma IVT2: "a <= b & (f b <= y & y <= f a) & (ALL x. a <= x & x <= b --> contl f x)
+==> EX x>=a. x <= b & f x = y"
   by (import lim IVT2)
 
-lemma DIFF_CONST: "ALL k::real. All (diffl (%x::real. k) 0)"
+lemma DIFF_CONST: "diffl (%x. k) 0 x"
   by (import lim DIFF_CONST)
 
-lemma DIFF_ADD: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   diffl f l x & diffl g m x --> diffl (%x::real. f x + g x) (l + m) x"
+lemma DIFF_ADD: "diffl f l x & diffl g m x ==> diffl (%x. f x + g x) (l + m) x"
   by (import lim DIFF_ADD)
 
-lemma DIFF_MUL: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   diffl f l x & diffl g m x -->
-   diffl (%x::real. f x * g x) (l * g x + m * f x) x"
+lemma DIFF_MUL: "diffl f l x & diffl g m x ==> diffl (%x. f x * g x) (l * g x + m * f x) x"
   by (import lim DIFF_MUL)
 
-lemma DIFF_CMUL: "ALL (f::real => real) (c::real) (l::real) x::real.
-   diffl f l x --> diffl (%x::real. c * f x) (c * l) x"
+lemma DIFF_CMUL: "diffl f l x ==> diffl (%x. c * f x) (c * l) x"
   by (import lim DIFF_CMUL)
 
-lemma DIFF_NEG: "ALL (f::real => real) (l::real) x::real.
-   diffl f l x --> diffl (%x::real. - f x) (- l) x"
+lemma DIFF_NEG: "diffl f l x ==> diffl (%x. - f x) (- l) x"
   by (import lim DIFF_NEG)
 
-lemma DIFF_SUB: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   diffl f l x & diffl g m x --> diffl (%x::real. f x - g x) (l - m) x"
+lemma DIFF_SUB: "diffl f l x & diffl g m x ==> diffl (%x. f x - g x) (l - m) x"
   by (import lim DIFF_SUB)
 
-lemma DIFF_CHAIN: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   diffl f l (g x) & diffl g m x --> diffl (%x::real. f (g x)) (l * m) x"
+lemma DIFF_CHAIN: "diffl f l (g x) & diffl g m x ==> diffl (%x. f (g x)) (l * m) x"
   by (import lim DIFF_CHAIN)
 
-lemma DIFF_X: "All (diffl (%x::real. x) 1)"
+lemma DIFF_X: "diffl (%x. x) 1 x"
   by (import lim DIFF_X)
 
-lemma DIFF_POW: "ALL (n::nat) x::real. diffl (%x::real. x ^ n) (real n * x ^ (n - 1)) x"
+lemma DIFF_POW: "diffl (%x. x ^ n) (real n * x ^ (n - 1)) x"
   by (import lim DIFF_POW)
 
-lemma DIFF_XM1: "ALL x::real. x ~= 0 --> diffl inverse (- (inverse x ^ 2)) x"
+lemma DIFF_XM1: "x ~= 0 ==> diffl inverse (- (inverse x ^ 2)) x"
   by (import lim DIFF_XM1)
 
-lemma DIFF_INV: "ALL (f::real => real) (l::real) x::real.
-   diffl f l x & f x ~= 0 -->
-   diffl (%x::real. inverse (f x)) (- (l / f x ^ 2)) x"
+lemma DIFF_INV: "diffl f l x & f x ~= 0 ==> diffl (%x. inverse (f x)) (- (l / f x ^ 2)) x"
   by (import lim DIFF_INV)
 
-lemma DIFF_DIV: "ALL (f::real => real) (g::real => real) (l::real) (m::real) x::real.
-   diffl f l x & diffl g m x & g x ~= 0 -->
-   diffl (%x::real. f x / g x) ((l * g x - m * f x) / g x ^ 2) x"
+lemma DIFF_DIV: "diffl f l x & diffl g m x & g x ~= 0
+==> diffl (%x. f x / g x) ((l * g x - m * f x) / g x ^ 2) x"
   by (import lim DIFF_DIV)
 
-lemma DIFF_SUM: "ALL (f::nat => real => real) (f'::nat => real => real) (m::nat) (n::nat)
-   x::real.
-   (ALL r::nat. m <= r & r < m + n --> diffl (f r) (f' r x) x) -->
-   diffl (%x::real. real.sum (m, n) (%n::nat. f n x))
-    (real.sum (m, n) (%r::nat. f' r x)) x"
+lemma DIFF_SUM: "(!!r. m <= r & r < m + n ==> diffl (f r) (f' r x) x)
+==> diffl (%x. real.sum (m, n) (%n. f n x)) (real.sum (m, n) (%r. f' r x)) x"
   by (import lim DIFF_SUM)
 
-lemma CONT_BOUNDED: "ALL (f::real => real) (a::real) b::real.
-   a <= b & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX M::real. ALL x::real. a <= x & x <= b --> f x <= M)"
+lemma CONT_BOUNDED: "a <= b & (ALL x. a <= x & x <= b --> contl f x)
+==> EX M. ALL x. a <= x & x <= b --> f x <= M"
   by (import lim CONT_BOUNDED)
 
-lemma CONT_HASSUP: "(All::((real => real) => bool) => bool)
- (%f::real => real.
-     (All::(real => bool) => bool)
-      (%a::real.
-          (All::(real => bool) => bool)
-           (%b::real.
-               (op -->::bool => bool => bool)
-                ((op &::bool => bool => bool)
-                  ((op <=::real => real => bool) a b)
-                  ((All::(real => bool) => bool)
-                    (%x::real.
-                        (op -->::bool => bool => bool)
-                         ((op &::bool => bool => bool)
-                           ((op <=::real => real => bool) a x)
-                           ((op <=::real => real => bool) x b))
-                         ((contl::(real => real) => real => bool) f x))))
-                ((Ex::(real => bool) => bool)
-                  (%M::real.
-                      (op &::bool => bool => bool)
-                       ((All::(real => bool) => bool)
-                         (%x::real.
-                             (op -->::bool => bool => bool)
-                              ((op &::bool => bool => bool)
-                                ((op <=::real => real => bool) a x)
-                                ((op <=::real => real => bool) x b))
-                              ((op <=::real => real => bool) (f x) M)))
-                       ((All::(real => bool) => bool)
-                         (%N::real.
-                             (op -->::bool => bool => bool)
-                              ((op <::real => real => bool) N M)
-                              ((Ex::(real => bool) => bool)
-                                (%x::real.
-                                    (op &::bool => bool => bool)
-                                     ((op <=::real => real => bool) a x)
-                                     ((op &::bool => bool => bool)
- ((op <=::real => real => bool) x b)
- ((op <::real => real => bool) N (f x))))))))))))"
+lemma CONT_HASSUP: "a <= b & (ALL x. a <= x & x <= b --> contl f x)
+==> EX M. (ALL x. a <= x & x <= b --> f x <= M) &
+          (ALL N<M. EX x>=a. x <= b & N < f x)"
   by (import lim CONT_HASSUP)
 
-lemma CONT_ATTAINS: "ALL (f::real => real) (a::real) b::real.
-   a <= b & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX x::real.
-       (ALL xa::real. a <= xa & xa <= b --> f xa <= x) &
-       (EX xa::real. a <= xa & xa <= b & f xa = x))"
+lemma CONT_ATTAINS: "a <= b & (ALL x. a <= x & x <= b --> contl f x)
+==> EX x. (ALL xa. a <= xa & xa <= b --> f xa <= x) &
+          (EX xa>=a. xa <= b & f xa = x)"
   by (import lim CONT_ATTAINS)
 
-lemma CONT_ATTAINS2: "ALL (f::real => real) (a::real) b::real.
-   a <= b & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX x::real.
-       (ALL xa::real. a <= xa & xa <= b --> x <= f xa) &
-       (EX xa::real. a <= xa & xa <= b & f xa = x))"
+lemma CONT_ATTAINS2: "a <= b & (ALL x. a <= x & x <= b --> contl f x)
+==> EX x. (ALL xa. a <= xa & xa <= b --> x <= f xa) &
+          (EX xa>=a. xa <= b & f xa = x)"
   by (import lim CONT_ATTAINS2)
 
-lemma CONT_ATTAINS_ALL: "ALL (f::real => real) (a::real) b::real.
-   a <= b & (ALL x::real. a <= x & x <= b --> contl f x) -->
-   (EX (x::real) M::real.
+lemma CONT_ATTAINS_ALL: "a <= b & (ALL x. a <= x & x <= b --> contl f x)
+==> EX x M.
        x <= M &
-       (ALL y::real.
-           x <= y & y <= M --> (EX x::real. a <= x & x <= b & f x = y)) &
-       (ALL xa::real. a <= xa & xa <= b --> x <= f xa & f xa <= M))"
+       (ALL y. x <= y & y <= M --> (EX x>=a. x <= b & f x = y)) &
+       (ALL xa. a <= xa & xa <= b --> x <= f xa & f xa <= M)"
   by (import lim CONT_ATTAINS_ALL)
 
-lemma DIFF_LINC: "ALL (f::real => real) (x::real) l::real.
-   diffl f l x & 0 < l -->
-   (EX d>0. ALL h::real. 0 < h & h < d --> f x < f (x + h))"
+lemma DIFF_LINC: "diffl f l x & 0 < l ==> EX d>0. ALL h. 0 < h & h < d --> f x < f (x + h)"
   by (import lim DIFF_LINC)
 
-lemma DIFF_LDEC: "ALL (f::real => real) (x::real) l::real.
-   diffl f l x & l < 0 -->
-   (EX d>0. ALL h::real. 0 < h & h < d --> f x < f (x - h))"
+lemma DIFF_LDEC: "diffl f l x & l < 0 ==> EX d>0. ALL h. 0 < h & h < d --> f x < f (x - h)"
   by (import lim DIFF_LDEC)
 
-lemma DIFF_LMAX: "ALL (f::real => real) (x::real) l::real.
-   diffl f l x & (EX d>0. ALL y::real. abs (x - y) < d --> f y <= f x) -->
-   l = 0"
+lemma DIFF_LMAX: "diffl f l x & (EX d>0. ALL y. abs (x - y) < d --> f y <= f x) ==> l = 0"
   by (import lim DIFF_LMAX)
 
-lemma DIFF_LMIN: "ALL (f::real => real) (x::real) l::real.
-   diffl f l x & (EX d>0. ALL y::real. abs (x - y) < d --> f x <= f y) -->
-   l = 0"
+lemma DIFF_LMIN: "diffl f l x & (EX d>0. ALL y. abs (x - y) < d --> f x <= f y) ==> l = 0"
   by (import lim DIFF_LMIN)
 
-lemma DIFF_LCONST: "ALL (f::real => real) (x::real) l::real.
-   diffl f l x & (EX d>0. ALL y::real. abs (x - y) < d --> f y = f x) -->
-   l = 0"
+lemma DIFF_LCONST: "diffl f l x & (EX d>0. ALL y. abs (x - y) < d --> f y = f x) ==> l = 0"
   by (import lim DIFF_LCONST)
 
-lemma INTERVAL_LEMMA: "ALL (a::real) (b::real) x::real.
-   a < x & x < b -->
-   (EX d>0. ALL y::real. abs (x - y) < d --> a <= y & y <= b)"
-  by (import lim INTERVAL_LEMMA)
-
-lemma ROLLE: "ALL (f::real => real) (a::real) b::real.
-   a < b &
-   f a = f b &
-   (ALL x::real. a <= x & x <= b --> contl f x) &
-   (ALL x::real. a < x & x < b --> differentiable f x) -->
-   (EX z::real. a < z & z < b & diffl f 0 z)"
+lemma ROLLE: "a < b &
+f a = f b &
+(ALL x. a <= x & x <= b --> contl f x) &
+(ALL x. a < x & x < b --> lim.differentiable f x)
+==> EX z>a. z < b & diffl f 0 z"
   by (import lim ROLLE)
 
-lemma MVT_LEMMA: "ALL (f::real => real) (a::real) b::real.
-   f a - (f b - f a) / (b - a) * a = f b - (f b - f a) / (b - a) * b"
-  by (import lim MVT_LEMMA)
-
-lemma MVT: "ALL (f::real => real) (a::real) b::real.
-   a < b &
-   (ALL x::real. a <= x & x <= b --> contl f x) &
-   (ALL x::real. a < x & x < b --> differentiable f x) -->
-   (EX (l::real) z::real.
-       a < z & z < b & diffl f l z & f b - f a = (b - a) * l)"
+lemma MVT: "a < b &
+(ALL x. a <= x & x <= b --> contl f x) &
+(ALL x. a < x & x < b --> lim.differentiable f x)
+==> EX l z. a < z & z < b & diffl f l z & f b - f a = (b - a) * l"
   by (import lim MVT)
 
-lemma DIFF_ISCONST_END: "ALL (f::real => real) (a::real) b::real.
-   a < b &
-   (ALL x::real. a <= x & x <= b --> contl f x) &
-   (ALL x::real. a < x & x < b --> diffl f 0 x) -->
-   f b = f a"
+lemma DIFF_ISCONST_END: "a < b &
+(ALL x. a <= x & x <= b --> contl f x) &
+(ALL x. a < x & x < b --> diffl f 0 x)
+==> f b = f a"
   by (import lim DIFF_ISCONST_END)
 
-lemma DIFF_ISCONST: "ALL (f::real => real) (a::real) b::real.
-   a < b &
-   (ALL x::real. a <= x & x <= b --> contl f x) &
-   (ALL x::real. a < x & x < b --> diffl f 0 x) -->
-   (ALL x::real. a <= x & x <= b --> f x = f a)"
+lemma DIFF_ISCONST: "[| a < b &
+   (ALL x. a <= x & x <= b --> contl f x) &
+   (ALL x. a < x & x < b --> diffl f 0 x);
+   a <= x & x <= b |]
+==> f x = f a"
   by (import lim DIFF_ISCONST)
 
-lemma DIFF_ISCONST_ALL: "ALL f::real => real. All (diffl f 0) --> (ALL (x::real) y::real. f x = f y)"
+lemma DIFF_ISCONST_ALL: "(!!x. diffl f 0 x) ==> f x = f y"
   by (import lim DIFF_ISCONST_ALL)
 
-lemma INTERVAL_ABS: "ALL (x::real) (z::real) d::real.
-   (x - d <= z & z <= x + d) = (abs (z - x) <= d)"
+lemma INTERVAL_ABS: "((x::real) - (d::real) <= (z::real) & z <= x + d) = (abs (z - x) <= d)"
   by (import lim INTERVAL_ABS)
 
-lemma CONT_INJ_LEMMA: "ALL (f::real => real) (g::real => real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) <= d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) <= d --> contl f z) -->
-   ~ (ALL z::real. abs (z - x) <= d --> f z <= f x)"
+lemma CONT_INJ_LEMMA: "0 < d &
+(ALL z. abs (z - x) <= d --> g (f z) = z) &
+(ALL z. abs (z - x) <= d --> contl f z)
+==> ~ (ALL z. abs (z - x) <= d --> f z <= f x)"
   by (import lim CONT_INJ_LEMMA)
 
-lemma CONT_INJ_LEMMA2: "ALL (f::real => real) (g::real => real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) <= d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) <= d --> contl f z) -->
-   ~ (ALL z::real. abs (z - x) <= d --> f x <= f z)"
+lemma CONT_INJ_LEMMA2: "0 < d &
+(ALL z. abs (z - x) <= d --> g (f z) = z) &
+(ALL z. abs (z - x) <= d --> contl f z)
+==> ~ (ALL z. abs (z - x) <= d --> f x <= f z)"
   by (import lim CONT_INJ_LEMMA2)
 
-lemma CONT_INJ_RANGE: "ALL (f::real => real) (g::real => real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) <= d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) <= d --> contl f z) -->
-   (EX e>0.
-       ALL y::real.
-          abs (y - f x) <= e --> (EX z::real. abs (z - x) <= d & f z = y))"
+lemma CONT_INJ_RANGE: "0 < d &
+(ALL z. abs (z - x) <= d --> g (f z) = z) &
+(ALL z. abs (z - x) <= d --> contl f z)
+==> EX e>0. ALL y. abs (y - f x) <= e --> (EX z. abs (z - x) <= d & f z = y)"
   by (import lim CONT_INJ_RANGE)
 
-lemma CONT_INVERSE: "ALL (f::real => real) (g::real => real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) <= d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) <= d --> contl f z) -->
-   contl g (f x)"
+lemma CONT_INVERSE: "0 < d &
+(ALL z. abs (z - x) <= d --> g (f z) = z) &
+(ALL z. abs (z - x) <= d --> contl f z)
+==> contl g (f x)"
   by (import lim CONT_INVERSE)
 
-lemma DIFF_INVERSE: "ALL (f::real => real) (g::real => real) (l::real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) <= d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) <= d --> contl f z) & diffl f l x & l ~= 0 -->
-   diffl g (inverse l) (f x)"
+lemma DIFF_INVERSE: "0 < d &
+(ALL z. abs (z - x) <= d --> g (f z) = z) &
+(ALL z. abs (z - x) <= d --> contl f z) & diffl f l x & l ~= 0
+==> diffl g (inverse l) (f x)"
   by (import lim DIFF_INVERSE)
 
-lemma DIFF_INVERSE_LT: "ALL (f::real => real) (g::real => real) (l::real) (x::real) d::real.
-   0 < d &
-   (ALL z::real. abs (z - x) < d --> g (f z) = z) &
-   (ALL z::real. abs (z - x) < d --> contl f z) & diffl f l x & l ~= 0 -->
-   diffl g (inverse l) (f x)"
+lemma DIFF_INVERSE_LT: "0 < d &
+(ALL z. abs (z - x) < d --> g (f z) = z) &
+(ALL z. abs (z - x) < d --> contl f z) & diffl f l x & l ~= 0
+==> diffl g (inverse l) (f x)"
   by (import lim DIFF_INVERSE_LT)
 
-lemma INTERVAL_CLEMMA: "ALL (a::real) (b::real) x::real.
-   a < x & x < b -->
-   (EX d>0. ALL y::real. abs (y - x) <= d --> a < y & y < b)"
+lemma INTERVAL_CLEMMA: "(a::real) < (x::real) & x < (b::real)
+==> EX d>0::real. ALL y::real. abs (y - x) <= d --> a < y & y < b"
   by (import lim INTERVAL_CLEMMA)
 
-lemma DIFF_INVERSE_OPEN: "ALL (f::real => real) (g::real => real) (l::real) (a::real) (x::real)
-   b::real.
-   a < x &
-   x < b &
-   (ALL z::real. a < z & z < b --> g (f z) = z & contl f z) &
-   diffl f l x & l ~= 0 -->
-   diffl g (inverse l) (f x)"
+lemma DIFF_INVERSE_OPEN: "a < x &
+x < b &
+(ALL z. a < z & z < b --> g (f z) = z & contl f z) & diffl f l x & l ~= 0
+==> diffl g (inverse l) (f x)"
   by (import lim DIFF_INVERSE_OPEN)
 
 ;end_setup
 
 ;setup_theory powser
 
-lemma POWDIFF_LEMMA: "ALL (n::nat) (x::real) y::real.
-   real.sum (0, Suc n) (%p::nat. x ^ p * y ^ (Suc n - p)) =
-   y * real.sum (0, Suc n) (%p::nat. x ^ p * y ^ (n - p))"
+lemma POWDIFF_LEMMA: "real.sum (0, Suc n) (%p. x ^ p * y ^ (Suc n - p)) =
+y * real.sum (0, Suc n) (%p. x ^ p * y ^ (n - p))"
   by (import powser POWDIFF_LEMMA)
 
-lemma POWDIFF: "ALL (n::nat) (x::real) y::real.
-   x ^ Suc n - y ^ Suc n =
-   (x - y) * real.sum (0, Suc n) (%p::nat. x ^ p * y ^ (n - p))"
+lemma POWDIFF: "x ^ Suc n - y ^ Suc n =
+(x - y) * real.sum (0, Suc n) (%p. x ^ p * y ^ (n - p))"
   by (import powser POWDIFF)
 
-lemma POWREV: "ALL (n::nat) (x::real) y::real.
-   real.sum (0, Suc n) (%xa::nat. x ^ xa * y ^ (n - xa)) =
-   real.sum (0, Suc n) (%xa::nat. x ^ (n - xa) * y ^ xa)"
+lemma POWREV: "real.sum (0, Suc n) (%xa. x ^ xa * y ^ (n - xa)) =
+real.sum (0, Suc n) (%xa. x ^ (n - xa) * y ^ xa)"
   by (import powser POWREV)
 
-lemma POWSER_INSIDEA: "ALL (f::nat => real) (x::real) z::real.
-   summable (%n::nat. f n * x ^ n) & abs z < abs x -->
-   summable (%n::nat. abs (f n) * z ^ n)"
+lemma POWSER_INSIDEA: "seq.summable (%n. f n * x ^ n) & abs z < abs x
+==> seq.summable (%n. abs (f n) * z ^ n)"
   by (import powser POWSER_INSIDEA)
 
-lemma POWSER_INSIDE: "ALL (f::nat => real) (x::real) z::real.
-   summable (%n::nat. f n * x ^ n) & abs z < abs x -->
-   summable (%n::nat. f n * z ^ n)"
+lemma POWSER_INSIDE: "seq.summable (%n. f n * x ^ n) & abs z < abs x
+==> seq.summable (%n. f n * z ^ n)"
   by (import powser POWSER_INSIDE)
 
-definition diffs :: "(nat => real) => nat => real" where 
-  "diffs == %(c::nat => real) n::nat. real (Suc n) * c (Suc n)"
-
-lemma diffs: "ALL c::nat => real. diffs c = (%n::nat. real (Suc n) * c (Suc n))"
+consts
+  diffs :: "(nat => real) => nat => real" 
+
+defs
+  diffs_def: "powser.diffs == %c n. real (Suc n) * c (Suc n)"
+
+lemma diffs: "powser.diffs c = (%n. real (Suc n) * c (Suc n))"
   by (import powser diffs)
 
-lemma DIFFS_NEG: "ALL c::nat => real. diffs (%n::nat. - c n) = (%x::nat. - diffs c x)"
+lemma DIFFS_NEG: "powser.diffs (%n. - c n) = (%x. - powser.diffs c x)"
   by (import powser DIFFS_NEG)
 
-lemma DIFFS_LEMMA: "ALL (n::nat) (c::nat => real) x::real.
-   real.sum (0, n) (%n::nat. diffs c n * x ^ n) =
-   real.sum (0, n) (%n::nat. real n * (c n * x ^ (n - 1))) +
-   real n * (c n * x ^ (n - 1))"
+lemma DIFFS_LEMMA: "real.sum (0, n) (%n. powser.diffs c n * x ^ n) =
+real.sum (0, n) (%n. real n * (c n * x ^ (n - 1))) +
+real n * (c n * x ^ (n - 1))"
   by (import powser DIFFS_LEMMA)
 
-lemma DIFFS_LEMMA2: "ALL (n::nat) (c::nat => real) x::real.
-   real.sum (0, n) (%n::nat. real n * (c n * x ^ (n - 1))) =
-   real.sum (0, n) (%n::nat. diffs c n * x ^ n) -
-   real n * (c n * x ^ (n - 1))"
+lemma DIFFS_LEMMA2: "real.sum (0, n) (%n. real n * (c n * x ^ (n - 1))) =
+real.sum (0, n) (%n. powser.diffs c n * x ^ n) -
+real n * (c n * x ^ (n - 1))"
   by (import powser DIFFS_LEMMA2)
 
-lemma DIFFS_EQUIV: "ALL (c::nat => real) x::real.
-   summable (%n::nat. diffs c n * x ^ n) -->
-   sums (%n::nat. real n * (c n * x ^ (n - 1)))
-    (suminf (%n::nat. diffs c n * x ^ n))"
+lemma DIFFS_EQUIV: "seq.summable (%n. powser.diffs c n * x ^ n)
+==> seq.sums (%n. real n * (c n * x ^ (n - 1)))
+     (seq.suminf (%n. powser.diffs c n * x ^ n))"
   by (import powser DIFFS_EQUIV)
 
-lemma TERMDIFF_LEMMA1: "ALL (m::nat) (z::real) h::real.
-   real.sum (0, m) (%p::nat. (z + h) ^ (m - p) * z ^ p - z ^ m) =
-   real.sum (0, m) (%p::nat. z ^ p * ((z + h) ^ (m - p) - z ^ (m - p)))"
+lemma TERMDIFF_LEMMA1: "real.sum (0, m) (%p. (z + h) ^ (m - p) * z ^ p - z ^ m) =
+real.sum (0, m) (%p. z ^ p * ((z + h) ^ (m - p) - z ^ (m - p)))"
   by (import powser TERMDIFF_LEMMA1)
 
-lemma TERMDIFF_LEMMA2: "ALL (z::real) (h::real) n::nat.
-   h ~= 0 -->
-   ((z + h) ^ n - z ^ n) / h - real n * z ^ (n - 1) =
-   h *
-   real.sum (0, n - 1)
-    (%p::nat.
-        z ^ p *
-        real.sum (0, n - 1 - p)
-         (%q::nat. (z + h) ^ q * z ^ (n - 2 - p - q)))"
+lemma TERMDIFF_LEMMA2: "h ~= 0
+==> ((z + h) ^ n - z ^ n) / h - real n * z ^ (n - 1) =
+    h *
+    real.sum (0, n - 1)
+     (%p. z ^ p *
+          real.sum (0, n - 1 - p) (%q. (z + h) ^ q * z ^ (n - 2 - p - q)))"
   by (import powser TERMDIFF_LEMMA2)
 
-lemma TERMDIFF_LEMMA3: "ALL (z::real) (h::real) (n::nat) k'::real.
-   h ~= 0 & abs z <= k' & abs (z + h) <= k' -->
-   abs (((z + h) ^ n - z ^ n) / h - real n * z ^ (n - 1))
-   <= real n * (real (n - 1) * (k' ^ (n - 2) * abs h))"
+lemma TERMDIFF_LEMMA3: "h ~= 0 & abs z <= k' & abs (z + h) <= k'
+==> abs (((z + h) ^ n - z ^ n) / h - real n * z ^ (n - 1))
+    <= real n * (real (n - 1) * (k' ^ (n - 2) * abs h))"
   by (import powser TERMDIFF_LEMMA3)
 
-lemma TERMDIFF_LEMMA4: "ALL (f::real => real) (k'::real) k::real.
-   0 < k &
-   (ALL h::real. 0 < abs h & abs h < k --> abs (f h) <= k' * abs h) -->
-   tends_real_real f 0 0"
+lemma TERMDIFF_LEMMA4: "0 < k & (ALL h. 0 < abs h & abs h < k --> abs (f h) <= k' * abs h)
+==> tends_real_real f 0 0"
   by (import powser TERMDIFF_LEMMA4)
 
-lemma TERMDIFF_LEMMA5: "ALL (f::nat => real) (g::real => nat => real) k::real.
-   0 < k &
-   summable f &
-   (ALL h::real.
-       0 < abs h & abs h < k -->
-       (ALL n::nat. abs (g h n) <= f n * abs h)) -->
-   tends_real_real (%h::real. suminf (g h)) 0 0"
+lemma TERMDIFF_LEMMA5: "0 < k &
+seq.summable f &
+(ALL h. 0 < abs h & abs h < k --> (ALL n. abs (g h n) <= f n * abs h))
+==> tends_real_real (%h. seq.suminf (g h)) 0 0"
   by (import powser TERMDIFF_LEMMA5)
 
-lemma TERMDIFF: "ALL (c::nat => real) (k'::real) x::real.
-   summable (%n::nat. c n * k' ^ n) &
-   summable (%n::nat. diffs c n * k' ^ n) &
-   summable (%n::nat. diffs (diffs c) n * k' ^ n) & abs x < abs k' -->
-   diffl (%x::real. suminf (%n::nat. c n * x ^ n))
-    (suminf (%n::nat. diffs c n * x ^ n)) x"
+lemma TERMDIFF: "seq.summable (%n. c n * k' ^ n) &
+seq.summable (%n. powser.diffs c n * k' ^ n) &
+seq.summable (%n. powser.diffs (powser.diffs c) n * k' ^ n) & abs x < abs k'
+==> diffl (%x. seq.suminf (%n. c n * x ^ n))
+     (seq.suminf (%n. powser.diffs c n * x ^ n)) x"
   by (import powser TERMDIFF)
 
 ;end_setup
 
 ;setup_theory transc
 
-definition exp :: "real => real" where 
-  "exp == %x::real. suminf (%n::nat. inverse (real (FACT n)) * x ^ n)"
-
-lemma exp: "ALL x::real. exp x = suminf (%n::nat. inverse (real (FACT n)) * x ^ n)"
+consts
+  exp :: "real => real" 
+
+defs
+  exp_def: "transc.exp == %x. seq.suminf (%n. inverse (real (FACT n)) * x ^ n)"
+
+lemma exp: "transc.exp x = seq.suminf (%n. inverse (real (FACT n)) * x ^ n)"
   by (import transc exp)
 
-definition cos :: "real => real" where 
-  "cos ==
-%x::real.
-   suminf
-    (%n::nat.
-        (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
-
-lemma cos: "ALL x::real.
-   cos x =
-   suminf
-    (%n::nat.
-        (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
+consts
+  cos :: "real => real" 
+
+defs
+  cos_def: "transc.cos ==
+%x. seq.suminf
+     (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
+
+lemma cos: "transc.cos x =
+seq.suminf
+ (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
   by (import transc cos)
 
-definition sin :: "real => real" where 
-  "sin ==
-%x::real.
-   suminf
-    (%n::nat.
-        (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
-        x ^ n)"
-
-lemma sin: "ALL x::real.
-   sin x =
-   suminf
-    (%n::nat.
-        (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
-        x ^ n)"
+consts
+  sin :: "real => real" 
+
+defs
+  sin_def: "transc.sin ==
+%x. seq.suminf
+     (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+          x ^ n)"
+
+lemma sin: "transc.sin x =
+seq.suminf
+ (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+      x ^ n)"
   by (import transc sin)
 
-lemma EXP_CONVERGES: "ALL x::real. sums (%n::nat. inverse (real (FACT n)) * x ^ n) (exp x)"
+lemma EXP_CONVERGES: "seq.sums (%n. inverse (real (FACT n)) * x ^ n) (transc.exp x)"
   by (import transc EXP_CONVERGES)
 
-lemma SIN_CONVERGES: "ALL x::real.
-   sums
-    (%n::nat.
-        (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
-        x ^ n)
-    (sin x)"
+lemma SIN_CONVERGES: "seq.sums
+ (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+      x ^ n)
+ (transc.sin x)"
   by (import transc SIN_CONVERGES)
 
-lemma COS_CONVERGES: "ALL x::real.
-   sums
-    (%n::nat.
-        (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)
-    (cos x)"
+lemma COS_CONVERGES: "seq.sums
+ (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)
+ (transc.cos x)"
   by (import transc COS_CONVERGES)
 
-lemma EXP_FDIFF: "diffs (%n::nat. inverse (real (FACT n))) =
-(%n::nat. inverse (real (FACT n)))"
+lemma EXP_FDIFF: "powser.diffs (%n. inverse (real (FACT n))) = (%n. inverse (real (FACT n)))"
   by (import transc EXP_FDIFF)
 
-lemma SIN_FDIFF: "diffs
- (%n::nat. if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) =
-(%n::nat. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0)"
+lemma SIN_FDIFF: "powser.diffs
+ (%n. if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) =
+(%n. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0)"
   by (import transc SIN_FDIFF)
 
-lemma COS_FDIFF: "diffs (%n::nat. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) =
-(%n::nat. - (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)))"
+lemma COS_FDIFF: "powser.diffs (%n. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) =
+(%n. - (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)))"
   by (import transc COS_FDIFF)
 
-lemma SIN_NEGLEMMA: "ALL x::real.
-   - sin x =
-   suminf
-    (%n::nat.
-        - ((if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
-           x ^ n))"
+lemma SIN_NEGLEMMA: "- transc.sin x =
+seq.suminf
+ (%n. - ((if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+         x ^ n))"
   by (import transc SIN_NEGLEMMA)
 
-lemma DIFF_EXP: "ALL x::real. diffl exp (exp x) x"
+lemma DIFF_EXP: "diffl transc.exp (transc.exp x) x"
   by (import transc DIFF_EXP)
 
-lemma DIFF_SIN: "ALL x::real. diffl sin (cos x) x"
+lemma DIFF_SIN: "diffl transc.sin (transc.cos x) x"
   by (import transc DIFF_SIN)
 
-lemma DIFF_COS: "ALL x::real. diffl cos (- sin x) x"
+lemma DIFF_COS: "diffl transc.cos (- transc.sin x) x"
   by (import transc DIFF_COS)
 
-lemma DIFF_COMPOSITE: "(diffl (f::real => real) (l::real) (x::real) & f x ~= 0 -->
- diffl (%x::real. inverse (f x)) (- (l / f x ^ 2)) x) &
-(diffl f l x & diffl (g::real => real) (m::real) x & g x ~= 0 -->
- diffl (%x::real. f x / g x) ((l * g x - m * f x) / g x ^ 2) x) &
-(diffl f l x & diffl g m x --> diffl (%x::real. f x + g x) (l + m) x) &
+lemma DIFF_COMPOSITE: "(diffl f l x & f x ~= 0 --> diffl (%x. inverse (f x)) (- (l / f x ^ 2)) x) &
+(diffl f l x & diffl g m x & g x ~= 0 -->
+ diffl (%x. f x / g x) ((l * g x - m * f x) / g x ^ 2) x) &
+(diffl f l x & diffl g m x --> diffl (%x. f x + g x) (l + m) x) &
 (diffl f l x & diffl g m x -->
- diffl (%x::real. f x * g x) (l * g x + m * f x) x) &
-(diffl f l x & diffl g m x --> diffl (%x::real. f x - g x) (l - m) x) &
-(diffl f l x --> diffl (%x::real. - f x) (- l) x) &
-(diffl g m x -->
- diffl (%x::real. g x ^ (n::nat)) (real n * g x ^ (n - 1) * m) x) &
-(diffl g m x --> diffl (%x::real. exp (g x)) (exp (g x) * m) x) &
-(diffl g m x --> diffl (%x::real. sin (g x)) (cos (g x) * m) x) &
-(diffl g m x --> diffl (%x::real. cos (g x)) (- sin (g x) * m) x)"
+ diffl (%x. f x * g x) (l * g x + m * f x) x) &
+(diffl f l x & diffl g m x --> diffl (%x. f x - g x) (l - m) x) &
+(diffl f l x --> diffl (%x. - f x) (- l) x) &
+(diffl g m x --> diffl (%x. g x ^ n) (real n * g x ^ (n - 1) * m) x) &
+(diffl g m x --> diffl (%x. transc.exp (g x)) (transc.exp (g x) * m) x) &
+(diffl g m x --> diffl (%x. transc.sin (g x)) (transc.cos (g x) * m) x) &
+(diffl g m x --> diffl (%x. transc.cos (g x)) (- transc.sin (g x) * m) x)"
   by (import transc DIFF_COMPOSITE)
 
-lemma EXP_0: "exp 0 = 1"
+lemma EXP_0: "transc.exp 0 = 1"
   by (import transc EXP_0)
 
-lemma EXP_LE_X: "ALL x>=0. 1 + x <= exp x"
+lemma EXP_LE_X: "0 <= x ==> 1 + x <= transc.exp x"
   by (import transc EXP_LE_X)
 
-lemma EXP_LT_1: "ALL x>0. 1 < exp x"
+lemma EXP_LT_1: "0 < x ==> 1 < transc.exp x"
   by (import transc EXP_LT_1)
 
-lemma EXP_ADD_MUL: "ALL (x::real) y::real. exp (x + y) * exp (- x) = exp y"
+lemma EXP_ADD_MUL: "transc.exp (x + y) * transc.exp (- x) = transc.exp y"
   by (import transc EXP_ADD_MUL)
 
-lemma EXP_NEG_MUL: "ALL x::real. exp x * exp (- x) = 1"
+lemma EXP_NEG_MUL: "transc.exp x * transc.exp (- x) = 1"
   by (import transc EXP_NEG_MUL)
 
-lemma EXP_NEG_MUL2: "ALL x::real. exp (- x) * exp x = 1"
+lemma EXP_NEG_MUL2: "transc.exp (- x) * transc.exp x = 1"
   by (import transc EXP_NEG_MUL2)
 
-lemma EXP_NEG: "ALL x::real. exp (- x) = inverse (exp x)"
+lemma EXP_NEG: "transc.exp (- x) = inverse (transc.exp x)"
   by (import transc EXP_NEG)
 
-lemma EXP_ADD: "ALL (x::real) y::real. exp (x + y) = exp x * exp y"
+lemma EXP_ADD: "transc.exp (x + y) = transc.exp x * transc.exp y"
   by (import transc EXP_ADD)
 
-lemma EXP_POS_LE: "ALL x::real. 0 <= exp x"
+lemma EXP_POS_LE: "0 <= transc.exp x"
   by (import transc EXP_POS_LE)
 
-lemma EXP_NZ: "ALL x::real. exp x ~= 0"
+lemma EXP_NZ: "transc.exp x ~= 0"
   by (import transc EXP_NZ)
 
-lemma EXP_POS_LT: "ALL x::real. 0 < exp x"
+lemma EXP_POS_LT: "0 < transc.exp x"
   by (import transc EXP_POS_LT)
 
-lemma EXP_N: "ALL (n::nat) x::real. exp (real n * x) = exp x ^ n"
+lemma EXP_N: "transc.exp (real n * x) = transc.exp x ^ n"
   by (import transc EXP_N)
 
-lemma EXP_SUB: "ALL (x::real) y::real. exp (x - y) = exp x / exp y"
+lemma EXP_SUB: "transc.exp (x - y) = transc.exp x / transc.exp y"
   by (import transc EXP_SUB)
 
-lemma EXP_MONO_IMP: "ALL (x::real) y::real. x < y --> exp x < exp y"
+lemma EXP_MONO_IMP: "x < y ==> transc.exp x < transc.exp y"
   by (import transc EXP_MONO_IMP)
 
-lemma EXP_MONO_LT: "ALL (x::real) y::real. (exp x < exp y) = (x < y)"
+lemma EXP_MONO_LT: "(transc.exp x < transc.exp y) = (x < y)"
   by (import transc EXP_MONO_LT)
 
-lemma EXP_MONO_LE: "ALL (x::real) y::real. (exp x <= exp y) = (x <= y)"
+lemma EXP_MONO_LE: "(transc.exp x <= transc.exp y) = (x <= y)"
   by (import transc EXP_MONO_LE)
 
-lemma EXP_INJ: "ALL (x::real) y::real. (exp x = exp y) = (x = y)"
+lemma EXP_INJ: "(transc.exp x = transc.exp y) = (x = y)"
   by (import transc EXP_INJ)
 
-lemma EXP_TOTAL_LEMMA: "ALL y>=1. EX x>=0. x <= y - 1 & exp x = y"
+lemma EXP_TOTAL_LEMMA: "1 <= y ==> EX x>=0. x <= y - 1 & transc.exp x = y"
   by (import transc EXP_TOTAL_LEMMA)
 
-lemma EXP_TOTAL: "ALL y>0. EX x::real. exp x = y"
+lemma EXP_TOTAL: "0 < y ==> EX x. transc.exp x = y"
   by (import transc EXP_TOTAL)
 
-definition ln :: "real => real" where 
-  "ln == %x::real. SOME u::real. exp u = x"
-
-lemma ln: "ALL x::real. ln x = (SOME u::real. exp u = x)"
+consts
+  ln :: "real => real" 
+
+defs
+  ln_def: "transc.ln == %x. SOME u. transc.exp u = x"
+
+lemma ln: "transc.ln x = (SOME u. transc.exp u = x)"
   by (import transc ln)
 
-lemma LN_EXP: "ALL x::real. ln (exp x) = x"
+lemma LN_EXP: "transc.ln (transc.exp x) = x"
   by (import transc LN_EXP)
 
-lemma EXP_LN: "ALL x::real. (exp (ln x) = x) = (0 < x)"
+lemma EXP_LN: "(transc.exp (transc.ln x) = x) = (0 < x)"
   by (import transc EXP_LN)
 
-lemma LN_MUL: "ALL (x::real) y::real. 0 < x & 0 < y --> ln (x * y) = ln x + ln y"
+lemma LN_MUL: "0 < x & 0 < y ==> transc.ln (x * y) = transc.ln x + transc.ln y"
   by (import transc LN_MUL)
 
-lemma LN_INJ: "ALL (x::real) y::real. 0 < x & 0 < y --> (ln x = ln y) = (x = y)"
+lemma LN_INJ: "0 < x & 0 < y ==> (transc.ln x = transc.ln y) = (x = y)"
   by (import transc LN_INJ)
 
-lemma LN_1: "ln 1 = 0"
+lemma LN_1: "transc.ln 1 = 0"
   by (import transc LN_1)
 
-lemma LN_INV: "ALL x>0. ln (inverse x) = - ln x"
+lemma LN_INV: "0 < x ==> transc.ln (inverse x) = - transc.ln x"
   by (import transc LN_INV)
 
-lemma LN_DIV: "ALL (x::real) y::real. 0 < x & 0 < y --> ln (x / y) = ln x - ln y"
+lemma LN_DIV: "0 < x & 0 < y ==> transc.ln (x / y) = transc.ln x - transc.ln y"
   by (import transc LN_DIV)
 
-lemma LN_MONO_LT: "ALL (x::real) y::real. 0 < x & 0 < y --> (ln x < ln y) = (x < y)"
+lemma LN_MONO_LT: "0 < x & 0 < y ==> (transc.ln x < transc.ln y) = (x < y)"
   by (import transc LN_MONO_LT)
 
-lemma LN_MONO_LE: "ALL (x::real) y::real. 0 < x & 0 < y --> (ln x <= ln y) = (x <= y)"
+lemma LN_MONO_LE: "0 < x & 0 < y ==> (transc.ln x <= transc.ln y) = (x <= y)"
   by (import transc LN_MONO_LE)
 
-lemma LN_POW: "ALL (n::nat) x::real. 0 < x --> ln (x ^ n) = real n * ln x"
+lemma LN_POW: "0 < x ==> transc.ln (x ^ n) = real n * transc.ln x"
   by (import transc LN_POW)
 
-lemma LN_LE: "ALL x>=0. ln (1 + x) <= x"
+lemma LN_LE: "0 <= x ==> transc.ln (1 + x) <= x"
   by (import transc LN_LE)
 
-lemma LN_LT_X: "ALL x>0. ln x < x"
+lemma LN_LT_X: "0 < x ==> transc.ln x < x"
   by (import transc LN_LT_X)
 
-lemma LN_POS: "ALL x>=1. 0 <= ln x"
+lemma LN_POS: "1 <= x ==> 0 <= transc.ln x"
   by (import transc LN_POS)
 
-definition root :: "nat => real => real" where 
-  "root == %(n::nat) x::real. SOME u::real. (0 < x --> 0 < u) & u ^ n = x"
-
-lemma root: "ALL (n::nat) x::real.
-   root n x = (SOME u::real. (0 < x --> 0 < u) & u ^ n = x)"
+consts
+  root :: "nat => real => real" 
+
+defs
+  root_def: "transc.root == %n x. SOME u. (0 < x --> 0 < u) & u ^ n = x"
+
+lemma root: "transc.root n x = (SOME u. (0 < x --> 0 < u) & u ^ n = x)"
   by (import transc root)
 
-definition sqrt :: "real => real" where 
-  "sqrt == root 2"
-
-lemma sqrt: "ALL x::real. sqrt x = root 2 x"
+consts
+  sqrt :: "real => real" 
+
+defs
+  sqrt_def: "transc.sqrt == transc.root 2"
+
+lemma sqrt: "transc.sqrt x = transc.root 2 x"
   by (import transc sqrt)
 
-lemma ROOT_LT_LEMMA: "ALL (n::nat) x::real. 0 < x --> exp (ln x / real (Suc n)) ^ Suc n = x"
+lemma ROOT_LT_LEMMA: "0 < x ==> transc.exp (transc.ln x / real (Suc n)) ^ Suc n = x"
   by (import transc ROOT_LT_LEMMA)
 
-lemma ROOT_LN: "ALL (n::nat) x::real. 0 < x --> root (Suc n) x = exp (ln x / real (Suc n))"
+lemma ROOT_LN: "0 < x ==> transc.root (Suc n) x = transc.exp (transc.ln x / real (Suc n))"
   by (import transc ROOT_LN)
 
-lemma ROOT_0: "ALL n::nat. root (Suc n) 0 = 0"
+lemma ROOT_0: "transc.root (Suc n) 0 = 0"
   by (import transc ROOT_0)
 
-lemma ROOT_1: "ALL n::nat. root (Suc n) 1 = 1"
+lemma ROOT_1: "transc.root (Suc n) 1 = 1"
   by (import transc ROOT_1)
 
-lemma ROOT_POS_LT: "ALL (n::nat) x::real. 0 < x --> 0 < root (Suc n) x"
+lemma ROOT_POS_LT: "0 < x ==> 0 < transc.root (Suc n) x"
   by (import transc ROOT_POS_LT)
 
-lemma ROOT_POW_POS: "ALL (n::nat) x::real. 0 <= x --> root (Suc n) x ^ Suc n = x"
+lemma ROOT_POW_POS: "0 <= x ==> transc.root (Suc n) x ^ Suc n = x"
   by (import transc ROOT_POW_POS)
 
-lemma POW_ROOT_POS: "ALL (n::nat) x::real. 0 <= x --> root (Suc n) (x ^ Suc n) = x"
+lemma POW_ROOT_POS: "0 <= x ==> transc.root (Suc n) (x ^ Suc n) = x"
   by (import transc POW_ROOT_POS)
 
-lemma ROOT_POS: "ALL (n::nat) x::real. 0 <= x --> 0 <= root (Suc n) x"
+lemma ROOT_POS: "0 <= x ==> 0 <= transc.root (Suc n) x"
   by (import transc ROOT_POS)
 
-lemma ROOT_POS_UNIQ: "ALL (n::nat) (x::real) y::real.
-   0 <= x & 0 <= y & y ^ Suc n = x --> root (Suc n) x = y"
+lemma ROOT_POS_UNIQ: "0 <= x & 0 <= y & y ^ Suc n = x ==> transc.root (Suc n) x = y"
   by (import transc ROOT_POS_UNIQ)
 
-lemma ROOT_MUL: "ALL (n::nat) (x::real) y::real.
-   0 <= x & 0 <= y -->
-   root (Suc n) (x * y) = root (Suc n) x * root (Suc n) y"
+lemma ROOT_MUL: "0 <= x & 0 <= y
+==> transc.root (Suc n) (x * y) =
+    transc.root (Suc n) x * transc.root (Suc n) y"
   by (import transc ROOT_MUL)
 
-lemma ROOT_INV: "ALL (n::nat) x::real.
-   0 <= x --> root (Suc n) (inverse x) = inverse (root (Suc n) x)"
+lemma ROOT_INV: "0 <= x ==> transc.root (Suc n) (inverse x) = inverse (transc.root (Suc n) x)"
   by (import transc ROOT_INV)
 
-lemma ROOT_DIV: "ALL (x::nat) (xa::real) xb::real.
-   0 <= xa & 0 <= xb -->
-   root (Suc x) (xa / xb) = root (Suc x) xa / root (Suc x) xb"
+lemma ROOT_DIV: "0 <= xa & 0 <= xb
+==> transc.root (Suc x) (xa / xb) =
+    transc.root (Suc x) xa / transc.root (Suc x) xb"
   by (import transc ROOT_DIV)
 
-lemma ROOT_MONO_LE: "ALL (x::real) y::real.
-   0 <= x & x <= y --> root (Suc (n::nat)) x <= root (Suc n) y"
+lemma ROOT_MONO_LE: "0 <= x & x <= y ==> transc.root (Suc n) x <= transc.root (Suc n) y"
   by (import transc ROOT_MONO_LE)
 
-lemma SQRT_0: "sqrt 0 = 0"
+lemma SQRT_0: "transc.sqrt 0 = 0"
   by (import transc SQRT_0)
 
-lemma SQRT_1: "sqrt 1 = 1"
+lemma SQRT_1: "transc.sqrt 1 = 1"
   by (import transc SQRT_1)
 
-lemma SQRT_POS_LT: "ALL x>0. 0 < sqrt x"
+lemma SQRT_POS_LT: "0 < x ==> 0 < transc.sqrt x"
   by (import transc SQRT_POS_LT)
 
-lemma SQRT_POS_LE: "ALL x>=0. 0 <= sqrt x"
+lemma SQRT_POS_LE: "0 <= x ==> 0 <= transc.sqrt x"
   by (import transc SQRT_POS_LE)
 
-lemma SQRT_POW2: "ALL x::real. (sqrt x ^ 2 = x) = (0 <= x)"
+lemma SQRT_POW2: "(transc.sqrt x ^ 2 = x) = (0 <= x)"
   by (import transc SQRT_POW2)
 
-lemma SQRT_POW_2: "ALL x>=0. sqrt x ^ 2 = x"
+lemma SQRT_POW_2: "0 <= x ==> transc.sqrt x ^ 2 = x"
   by (import transc SQRT_POW_2)
 
-lemma POW_2_SQRT: "0 <= (x::real) --> sqrt (x ^ 2) = x"
+lemma POW_2_SQRT: "0 <= x ==> transc.sqrt (x ^ 2) = x"
   by (import transc POW_2_SQRT)
 
-lemma SQRT_POS_UNIQ: "ALL (x::real) xa::real. 0 <= x & 0 <= xa & xa ^ 2 = x --> sqrt x = xa"
+lemma SQRT_POS_UNIQ: "0 <= x & 0 <= xa & xa ^ 2 = x ==> transc.sqrt x = xa"
   by (import transc SQRT_POS_UNIQ)
 
-lemma SQRT_MUL: "ALL (x::real) xa::real.
-   0 <= x & 0 <= xa --> sqrt (x * xa) = sqrt x * sqrt xa"
+lemma SQRT_MUL: "0 <= x & 0 <= xa ==> transc.sqrt (x * xa) = transc.sqrt x * transc.sqrt xa"
   by (import transc SQRT_MUL)
 
-lemma SQRT_INV: "ALL x>=0. sqrt (inverse x) = inverse (sqrt x)"
+lemma SQRT_INV: "0 <= x ==> transc.sqrt (inverse x) = inverse (transc.sqrt x)"
   by (import transc SQRT_INV)
 
-lemma SQRT_DIV: "ALL (x::real) xa::real.
-   0 <= x & 0 <= xa --> sqrt (x / xa) = sqrt x / sqrt xa"
+lemma SQRT_DIV: "0 <= x & 0 <= xa ==> transc.sqrt (x / xa) = transc.sqrt x / transc.sqrt xa"
   by (import transc SQRT_DIV)
 
-lemma SQRT_MONO_LE: "ALL (x::real) xa::real. 0 <= x & x <= xa --> sqrt x <= sqrt xa"
+lemma SQRT_MONO_LE: "0 <= x & x <= xa ==> transc.sqrt x <= transc.sqrt xa"
   by (import transc SQRT_MONO_LE)
 
-lemma SQRT_EVEN_POW2: "ALL n::nat. EVEN n --> sqrt (2 ^ n) = 2 ^ (n div 2)"
+lemma SQRT_EVEN_POW2: "EVEN n ==> transc.sqrt (2 ^ n) = 2 ^ (n div 2)"
   by (import transc SQRT_EVEN_POW2)
 
-lemma REAL_DIV_SQRT: "ALL x>=0. x / sqrt x = sqrt x"
+lemma REAL_DIV_SQRT: "0 <= x ==> x / transc.sqrt x = transc.sqrt x"
   by (import transc REAL_DIV_SQRT)
 
-lemma SQRT_EQ: "ALL (x::real) y::real. x ^ 2 = y & 0 <= x --> x = sqrt y"
+lemma SQRT_EQ: "x ^ 2 = y & 0 <= x ==> x = transc.sqrt y"
   by (import transc SQRT_EQ)
 
-lemma SIN_0: "sin 0 = 0"
+lemma SIN_0: "transc.sin 0 = 0"
   by (import transc SIN_0)
 
-lemma COS_0: "cos 0 = 1"
+lemma COS_0: "transc.cos 0 = 1"
   by (import transc COS_0)
 
-lemma SIN_CIRCLE: "ALL x::real. sin x ^ 2 + cos x ^ 2 = 1"
+lemma SIN_CIRCLE: "transc.sin x ^ 2 + transc.cos x ^ 2 = 1"
   by (import transc SIN_CIRCLE)
 
-lemma SIN_BOUND: "ALL x::real. abs (sin x) <= 1"
+lemma SIN_BOUND: "abs (transc.sin x) <= 1"
   by (import transc SIN_BOUND)
 
-lemma SIN_BOUNDS: "ALL x::real. - 1 <= sin x & sin x <= 1"
+lemma SIN_BOUNDS: "- 1 <= transc.sin x & transc.sin x <= 1"
   by (import transc SIN_BOUNDS)
 
-lemma COS_BOUND: "ALL x::real. abs (cos x) <= 1"
+lemma COS_BOUND: "abs (transc.cos x) <= 1"
   by (import transc COS_BOUND)
 
-lemma COS_BOUNDS: "ALL x::real. - 1 <= cos x & cos x <= 1"
+lemma COS_BOUNDS: "- 1 <= transc.cos x & transc.cos x <= 1"
   by (import transc COS_BOUNDS)
 
-lemma SIN_COS_ADD: "ALL (x::real) y::real.
-   (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +
-   (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 =
-   0"
+lemma SIN_COS_ADD: "(transc.sin (x + y) -
+ (transc.sin x * transc.cos y + transc.cos x * transc.sin y)) ^
+2 +
+(transc.cos (x + y) -
+ (transc.cos x * transc.cos y - transc.sin x * transc.sin y)) ^
+2 =
+0"
   by (import transc SIN_COS_ADD)
 
-lemma SIN_COS_NEG: "ALL x::real. (sin (- x) + sin x) ^ 2 + (cos (- x) - cos x) ^ 2 = 0"
+lemma SIN_COS_NEG: "(transc.sin (- x) + transc.sin x) ^ 2 +
+(transc.cos (- x) - transc.cos x) ^ 2 =
+0"
   by (import transc SIN_COS_NEG)
 
-lemma SIN_ADD: "ALL (x::real) y::real. sin (x + y) = sin x * cos y + cos x * sin y"
+lemma SIN_ADD: "transc.sin (x + y) =
+transc.sin x * transc.cos y + transc.cos x * transc.sin y"
   by (import transc SIN_ADD)
 
-lemma COS_ADD: "ALL (x::real) y::real. cos (x + y) = cos x * cos y - sin x * sin y"
+lemma COS_ADD: "transc.cos (x + y) =
+transc.cos x * transc.cos y - transc.sin x * transc.sin y"
   by (import transc COS_ADD)
 
-lemma SIN_NEG: "ALL x::real. sin (- x) = - sin x"
+lemma SIN_NEG: "transc.sin (- x) = - transc.sin x"
   by (import transc SIN_NEG)
 
-lemma COS_NEG: "ALL x::real. cos (- x) = cos x"
+lemma COS_NEG: "transc.cos (- x) = transc.cos x"
   by (import transc COS_NEG)
 
-lemma SIN_DOUBLE: "ALL x::real. sin (2 * x) = 2 * (sin x * cos x)"
+lemma SIN_DOUBLE: "transc.sin (2 * x) = 2 * (transc.sin x * transc.cos x)"
   by (import transc SIN_DOUBLE)
 
-lemma COS_DOUBLE: "ALL x::real. cos (2 * x) = cos x ^ 2 - sin x ^ 2"
+lemma COS_DOUBLE: "transc.cos (2 * x) = transc.cos x ^ 2 - transc.sin x ^ 2"
   by (import transc COS_DOUBLE)
 
-lemma SIN_PAIRED: "ALL x::real.
-   sums (%n::nat. (- 1) ^ n / real (FACT (2 * n + 1)) * x ^ (2 * n + 1))
-    (sin x)"
+lemma SIN_PAIRED: "seq.sums (%n. (- 1) ^ n / real (FACT (2 * n + 1)) * x ^ (2 * n + 1))
+ (transc.sin x)"
   by (import transc SIN_PAIRED)
 
-lemma SIN_POS: "ALL x::real. 0 < x & x < 2 --> 0 < sin x"
+lemma SIN_POS: "0 < x & x < 2 ==> 0 < transc.sin x"
   by (import transc SIN_POS)
 
-lemma COS_PAIRED: "ALL x::real.
-   sums (%n::nat. (- 1) ^ n / real (FACT (2 * n)) * x ^ (2 * n)) (cos x)"
+lemma COS_PAIRED: "seq.sums (%n. (- 1) ^ n / real (FACT (2 * n)) * x ^ (2 * n)) (transc.cos x)"
   by (import transc COS_PAIRED)
 
-lemma COS_2: "cos 2 < 0"
+lemma COS_2: "transc.cos 2 < 0"
   by (import transc COS_2)
 
-lemma COS_ISZERO: "EX! x::real. 0 <= x & x <= 2 & cos x = 0"
+lemma COS_ISZERO: "EX! x. 0 <= x & x <= 2 & transc.cos x = 0"
   by (import transc COS_ISZERO)
 
-definition pi :: "real" where 
-  "pi == 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)"
-
-lemma pi: "pi = 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)"
+consts
+  pi :: "real" 
+
+defs
+  pi_def: "transc.pi == 2 * (SOME x. 0 <= x & x <= 2 & transc.cos x = 0)"
+
+lemma pi: "transc.pi = 2 * (SOME x. 0 <= x & x <= 2 & transc.cos x = 0)"
   by (import transc pi)
 
-lemma PI2: "pi / 2 = (SOME x::real. 0 <= x & x <= 2 & cos x = 0)"
+lemma PI2: "transc.pi / 2 = (SOME x. 0 <= x & x <= 2 & transc.cos x = 0)"
   by (import transc PI2)
 
-lemma COS_PI2: "cos (pi / 2) = 0"
+lemma COS_PI2: "transc.cos (transc.pi / 2) = 0"
   by (import transc COS_PI2)
 
-lemma PI2_BOUNDS: "0 < pi / 2 & pi / 2 < 2"
+lemma PI2_BOUNDS: "0 < transc.pi / 2 & transc.pi / 2 < 2"
   by (import transc PI2_BOUNDS)
 
-lemma PI_POS: "0 < pi"
+lemma PI_POS: "0 < transc.pi"
   by (import transc PI_POS)
 
-lemma SIN_PI2: "sin (pi / 2) = 1"
+lemma SIN_PI2: "transc.sin (transc.pi / 2) = 1"
   by (import transc SIN_PI2)
 
-lemma COS_PI: "cos pi = - 1"
+lemma COS_PI: "transc.cos transc.pi = - 1"
   by (import transc COS_PI)
 
-lemma SIN_PI: "sin pi = 0"
+lemma SIN_PI: "transc.sin transc.pi = 0"
   by (import transc SIN_PI)
 
-lemma SIN_COS: "ALL x::real. sin x = cos (pi / 2 - x)"
+lemma SIN_COS: "transc.sin x = transc.cos (transc.pi / 2 - x)"
   by (import transc SIN_COS)
 
-lemma COS_SIN: "ALL x::real. cos x = sin (pi / 2 - x)"
+lemma COS_SIN: "transc.cos x = transc.sin (transc.pi / 2 - x)"
   by (import transc COS_SIN)
 
-lemma SIN_PERIODIC_PI: "ALL x::real. sin (x + pi) = - sin x"
+lemma SIN_PERIODIC_PI: "transc.sin (x + transc.pi) = - transc.sin x"
   by (import transc SIN_PERIODIC_PI)
 
-lemma COS_PERIODIC_PI: "ALL x::real. cos (x + pi) = - cos x"
+lemma COS_PERIODIC_PI: "transc.cos (x + transc.pi) = - transc.cos x"
   by (import transc COS_PERIODIC_PI)
 
-lemma SIN_PERIODIC: "ALL x::real. sin (x + 2 * pi) = sin x"
+lemma SIN_PERIODIC: "transc.sin (x + 2 * transc.pi) = transc.sin x"
   by (import transc SIN_PERIODIC)
 
-lemma COS_PERIODIC: "ALL x::real. cos (x + 2 * pi) = cos x"
+lemma COS_PERIODIC: "transc.cos (x + 2 * transc.pi) = transc.cos x"
   by (import transc COS_PERIODIC)
 
-lemma COS_NPI: "ALL n::nat. cos (real n * pi) = (- 1) ^ n"
+lemma COS_NPI: "transc.cos (real n * transc.pi) = (- 1) ^ n"
   by (import transc COS_NPI)
 
-lemma SIN_NPI: "ALL n::nat. sin (real n * pi) = 0"
+lemma SIN_NPI: "transc.sin (real (n::nat) * transc.pi) = (0::real)"
   by (import transc SIN_NPI)
 
-lemma SIN_POS_PI2: "ALL x::real. 0 < x & x < pi / 2 --> 0 < sin x"
+lemma SIN_POS_PI2: "0 < x & x < transc.pi / 2 ==> 0 < transc.sin x"
   by (import transc SIN_POS_PI2)
 
-lemma COS_POS_PI2: "ALL x::real. 0 < x & x < pi / 2 --> 0 < cos x"
+lemma COS_POS_PI2: "0 < x & x < transc.pi / 2 ==> 0 < transc.cos x"
   by (import transc COS_POS_PI2)
 
-lemma COS_POS_PI: "ALL x::real. - (pi / 2) < x & x < pi / 2 --> 0 < cos x"
+lemma COS_POS_PI: "- (transc.pi / 2) < x & x < transc.pi / 2 ==> 0 < transc.cos x"
   by (import transc COS_POS_PI)
 
-lemma SIN_POS_PI: "ALL x::real. 0 < x & x < pi --> 0 < sin x"
+lemma SIN_POS_PI: "0 < x & x < transc.pi ==> 0 < transc.sin x"
   by (import transc SIN_POS_PI)
 
-lemma COS_POS_PI2_LE: "ALL x::real. 0 <= x & x <= pi / 2 --> 0 <= cos x"
+lemma COS_POS_PI2_LE: "0 <= x & x <= transc.pi / 2 ==> 0 <= transc.cos x"
   by (import transc COS_POS_PI2_LE)
 
-lemma COS_POS_PI_LE: "ALL x::real. - (pi / 2) <= x & x <= pi / 2 --> 0 <= cos x"
+lemma COS_POS_PI_LE: "- (transc.pi / 2) <= x & x <= transc.pi / 2 ==> 0 <= transc.cos x"
   by (import transc COS_POS_PI_LE)
 
-lemma SIN_POS_PI2_LE: "ALL x::real. 0 <= x & x <= pi / 2 --> 0 <= sin x"
+lemma SIN_POS_PI2_LE: "0 <= x & x <= transc.pi / 2 ==> 0 <= transc.sin x"
   by (import transc SIN_POS_PI2_LE)
 
-lemma SIN_POS_PI_LE: "ALL x::real. 0 <= x & x <= pi --> 0 <= sin x"
+lemma SIN_POS_PI_LE: "0 <= x & x <= transc.pi ==> 0 <= transc.sin x"
   by (import transc SIN_POS_PI_LE)
 
-lemma COS_TOTAL: "ALL y::real.
-   - 1 <= y & y <= 1 --> (EX! x::real. 0 <= x & x <= pi & cos x = y)"
+lemma COS_TOTAL: "- 1 <= y & y <= 1 ==> EX! x. 0 <= x & x <= transc.pi & transc.cos x = y"
   by (import transc COS_TOTAL)
 
-lemma SIN_TOTAL: "ALL y::real.
-   - 1 <= y & y <= 1 -->
-   (EX! x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y)"
+lemma SIN_TOTAL: "- 1 <= y & y <= 1
+==> EX! x. - (transc.pi / 2) <= x & x <= transc.pi / 2 & transc.sin x = y"
   by (import transc SIN_TOTAL)
 
-lemma COS_ZERO_LEMMA: "ALL x::real.
-   0 <= x & cos x = 0 --> (EX n::nat. ~ EVEN n & x = real n * (pi / 2))"
+lemma COS_ZERO_LEMMA: "0 <= x & transc.cos x = 0 ==> EX n. ~ EVEN n & x = real n * (transc.pi / 2)"
   by (import transc COS_ZERO_LEMMA)
 
-lemma SIN_ZERO_LEMMA: "ALL x::real.
-   0 <= x & sin x = 0 --> (EX n::nat. EVEN n & x = real n * (pi / 2))"
+lemma SIN_ZERO_LEMMA: "0 <= x & transc.sin x = 0 ==> EX n. EVEN n & x = real n * (transc.pi / 2)"
   by (import transc SIN_ZERO_LEMMA)
 
-lemma COS_ZERO: "ALL x::real.
-   (cos x = 0) =
-   ((EX n::nat. ~ EVEN n & x = real n * (pi / 2)) |
-    (EX n::nat. ~ EVEN n & x = - (real n * (pi / 2))))"
+lemma COS_ZERO: "(transc.cos x = 0) =
+((EX n. ~ EVEN n & x = real n * (transc.pi / 2)) |
+ (EX n. ~ EVEN n & x = - (real n * (transc.pi / 2))))"
   by (import transc COS_ZERO)
 
-lemma SIN_ZERO: "ALL x::real.
-   (sin x = 0) =
-   ((EX n::nat. EVEN n & x = real n * (pi / 2)) |
-    (EX n::nat. EVEN n & x = - (real n * (pi / 2))))"
+lemma SIN_ZERO: "(transc.sin x = 0) =
+((EX n. EVEN n & x = real n * (transc.pi / 2)) |
+ (EX n. EVEN n & x = - (real n * (transc.pi / 2))))"
   by (import transc SIN_ZERO)
 
-definition tan :: "real => real" where 
-  "tan == %x::real. sin x / cos x"
-
-lemma tan: "ALL x::real. tan x = sin x / cos x"
+consts
+  tan :: "real => real" 
+
+defs
+  tan_def: "transc.tan == %x. transc.sin x / transc.cos x"
+
+lemma tan: "transc.tan x = transc.sin x / transc.cos x"
   by (import transc tan)
 
-lemma TAN_0: "tan 0 = 0"
+lemma TAN_0: "transc.tan 0 = 0"
   by (import transc TAN_0)
 
-lemma TAN_PI: "tan pi = 0"
+lemma TAN_PI: "transc.tan transc.pi = 0"
   by (import transc TAN_PI)
 
-lemma TAN_NPI: "ALL n::nat. tan (real n * pi) = 0"
+lemma TAN_NPI: "transc.tan (real (n::nat) * transc.pi) = (0::real)"
   by (import transc TAN_NPI)
 
-lemma TAN_NEG: "ALL x::real. tan (- x) = - tan x"
+lemma TAN_NEG: "transc.tan (- x) = - transc.tan x"
   by (import transc TAN_NEG)
 
-lemma TAN_PERIODIC: "ALL x::real. tan (x + 2 * pi) = tan x"
+lemma TAN_PERIODIC: "transc.tan (x + 2 * transc.pi) = transc.tan x"
   by (import transc TAN_PERIODIC)
 
-lemma TAN_ADD: "ALL (x::real) y::real.
-   cos x ~= 0 & cos y ~= 0 & cos (x + y) ~= 0 -->
-   tan (x + y) = (tan x + tan y) / (1 - tan x * tan y)"
+lemma TAN_ADD: "transc.cos x ~= 0 & transc.cos y ~= 0 & transc.cos (x + y) ~= 0
+==> transc.tan (x + y) =
+    (transc.tan x + transc.tan y) / (1 - transc.tan x * transc.tan y)"
   by (import transc TAN_ADD)
 
-lemma TAN_DOUBLE: "ALL x::real.
-   cos x ~= 0 & cos (2 * x) ~= 0 -->
-   tan (2 * x) = 2 * tan x / (1 - tan x ^ 2)"
+lemma TAN_DOUBLE: "transc.cos x ~= 0 & transc.cos (2 * x) ~= 0
+==> transc.tan (2 * x) = 2 * transc.tan x / (1 - transc.tan x ^ 2)"
   by (import transc TAN_DOUBLE)
 
-lemma TAN_POS_PI2: "ALL x::real. 0 < x & x < pi / 2 --> 0 < tan x"
+lemma TAN_POS_PI2: "0 < x & x < transc.pi / 2 ==> 0 < transc.tan x"
   by (import transc TAN_POS_PI2)
 
-lemma DIFF_TAN: "ALL x::real. cos x ~= 0 --> diffl tan (inverse (cos x ^ 2)) x"
+lemma DIFF_TAN: "transc.cos x ~= 0 ==> diffl transc.tan (inverse (transc.cos x ^ 2)) x"
   by (import transc DIFF_TAN)
 
-lemma TAN_TOTAL_LEMMA: "ALL y>0. EX x>0. x < pi / 2 & y < tan x"
+lemma TAN_TOTAL_LEMMA: "0 < y ==> EX x>0. x < transc.pi / 2 & y < transc.tan x"
   by (import transc TAN_TOTAL_LEMMA)
 
-lemma TAN_TOTAL_POS: "ALL y>=0. EX x>=0. x < pi / 2 & tan x = y"
+lemma TAN_TOTAL_POS: "0 <= y ==> EX x>=0. x < transc.pi / 2 & transc.tan x = y"
   by (import transc TAN_TOTAL_POS)
 
-lemma TAN_TOTAL: "ALL y::real. EX! x::real. - (pi / 2) < x & x < pi / 2 & tan x = y"
+lemma TAN_TOTAL: "EX! x. - (transc.pi / 2) < x & x < transc.pi / 2 & transc.tan x = y"
   by (import transc TAN_TOTAL)
 
-definition asn :: "real => real" where 
-  "asn == %y::real. SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y"
-
-lemma asn: "ALL y::real.
-   asn y = (SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y)"
+definition
+  asn :: "real => real"  where
+  "asn ==
+%y. SOME x. - (transc.pi / 2) <= x & x <= transc.pi / 2 & transc.sin x = y"
+
+lemma asn: "asn y =
+(SOME x. - (transc.pi / 2) <= x & x <= transc.pi / 2 & transc.sin x = y)"
   by (import transc asn)
 
-definition acs :: "real => real" where 
-  "acs == %y::real. SOME x::real. 0 <= x & x <= pi & cos x = y"
-
-lemma acs: "ALL y::real. acs y = (SOME x::real. 0 <= x & x <= pi & cos x = y)"
+definition
+  acs :: "real => real"  where
+  "acs == %y. SOME x. 0 <= x & x <= transc.pi & transc.cos x = y"
+
+lemma acs: "acs y = (SOME x. 0 <= x & x <= transc.pi & transc.cos x = y)"
   by (import transc acs)
 
-definition atn :: "real => real" where 
-  "atn == %y::real. SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y"
-
-lemma atn: "ALL y::real. atn y = (SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y)"
+definition
+  atn :: "real => real"  where
+  "atn ==
+%y. SOME x. - (transc.pi / 2) < x & x < transc.pi / 2 & transc.tan x = y"
+
+lemma atn: "atn y =
+(SOME x. - (transc.pi / 2) < x & x < transc.pi / 2 & transc.tan x = y)"
   by (import transc atn)
 
-lemma ASN: "ALL y::real.
-   - 1 <= y & y <= 1 -->
-   - (pi / 2) <= asn y & asn y <= pi / 2 & sin (asn y) = y"
+lemma ASN: "- 1 <= y & y <= 1
+==> - (transc.pi / 2) <= asn y &
+    asn y <= transc.pi / 2 & transc.sin (asn y) = y"
   by (import transc ASN)
 
-lemma ASN_SIN: "ALL y::real. - 1 <= y & y <= 1 --> sin (asn y) = y"
+lemma ASN_SIN: "- 1 <= y & y <= 1 ==> transc.sin (asn y) = y"
   by (import transc ASN_SIN)
 
-lemma ASN_BOUNDS: "ALL y::real. - 1 <= y & y <= 1 --> - (pi / 2) <= asn y & asn y <= pi / 2"
+lemma ASN_BOUNDS: "- 1 <= y & y <= 1 ==> - (transc.pi / 2) <= asn y & asn y <= transc.pi / 2"
   by (import transc ASN_BOUNDS)
 
-lemma ASN_BOUNDS_LT: "ALL y::real. - 1 < y & y < 1 --> - (pi / 2) < asn y & asn y < pi / 2"
+lemma ASN_BOUNDS_LT: "- 1 < y & y < 1 ==> - (transc.pi / 2) < asn y & asn y < transc.pi / 2"
   by (import transc ASN_BOUNDS_LT)
 
-lemma SIN_ASN: "ALL x::real. - (pi / 2) <= x & x <= pi / 2 --> asn (sin x) = x"
+lemma SIN_ASN: "- (transc.pi / 2) <= x & x <= transc.pi / 2 ==> asn (transc.sin x) = x"
   by (import transc SIN_ASN)
 
-lemma ACS: "ALL y::real.
-   - 1 <= y & y <= 1 --> 0 <= acs y & acs y <= pi & cos (acs y) = y"
+lemma ACS: "- 1 <= y & y <= 1
+==> 0 <= acs y & acs y <= transc.pi & transc.cos (acs y) = y"
   by (import transc ACS)
 
-lemma ACS_COS: "ALL y::real. - 1 <= y & y <= 1 --> cos (acs y) = y"
+lemma ACS_COS: "- 1 <= y & y <= 1 ==> transc.cos (acs y) = y"
   by (import transc ACS_COS)
 
-lemma ACS_BOUNDS: "ALL y::real. - 1 <= y & y <= 1 --> 0 <= acs y & acs y <= pi"
+lemma ACS_BOUNDS: "- 1 <= y & y <= 1 ==> 0 <= acs y & acs y <= transc.pi"
   by (import transc ACS_BOUNDS)
 
-lemma ACS_BOUNDS_LT: "ALL y::real. - 1 < y & y < 1 --> 0 < acs y & acs y < pi"
+lemma ACS_BOUNDS_LT: "- 1 < y & y < 1 ==> 0 < acs y & acs y < transc.pi"
   by (import transc ACS_BOUNDS_LT)
 
-lemma COS_ACS: "ALL x::real. 0 <= x & x <= pi --> acs (cos x) = x"
+lemma COS_ACS: "0 <= x & x <= transc.pi ==> acs (transc.cos x) = x"
   by (import transc COS_ACS)
 
-lemma ATN: "ALL y::real. - (pi / 2) < atn y & atn y < pi / 2 & tan (atn y) = y"
+lemma ATN: "- (transc.pi / 2) < atn y & atn y < transc.pi / 2 & transc.tan (atn y) = y"
   by (import transc ATN)
 
-lemma ATN_TAN: "ALL x::real. tan (atn x) = x"
+lemma ATN_TAN: "transc.tan (atn x) = x"
   by (import transc ATN_TAN)
 
-lemma ATN_BOUNDS: "ALL x::real. - (pi / 2) < atn x & atn x < pi / 2"
+lemma ATN_BOUNDS: "- (transc.pi / 2) < atn x & atn x < transc.pi / 2"
   by (import transc ATN_BOUNDS)
 
-lemma TAN_ATN: "ALL x::real. - (pi / 2) < x & x < pi / 2 --> atn (tan x) = x"
+lemma TAN_ATN: "- (transc.pi / 2) < x & x < transc.pi / 2 ==> atn (transc.tan x) = x"
   by (import transc TAN_ATN)
 
-lemma TAN_SEC: "ALL x::real. cos x ~= 0 --> 1 + tan x ^ 2 = inverse (cos x) ^ 2"
+lemma TAN_SEC: "transc.cos x ~= 0 ==> 1 + transc.tan x ^ 2 = inverse (transc.cos x) ^ 2"
   by (import transc TAN_SEC)
 
-lemma SIN_COS_SQ: "ALL x::real. 0 <= x & x <= pi --> sin x = sqrt (1 - cos x ^ 2)"
+lemma SIN_COS_SQ: "0 <= x & x <= transc.pi
+==> transc.sin x = transc.sqrt (1 - transc.cos x ^ 2)"
   by (import transc SIN_COS_SQ)
 
-lemma COS_SIN_SQ: "ALL x::real. - (pi / 2) <= x & x <= pi / 2 --> cos x = sqrt (1 - sin x ^ 2)"
+lemma COS_SIN_SQ: "- (transc.pi / 2) <= x & x <= transc.pi / 2
+==> transc.cos x = transc.sqrt (1 - transc.sin x ^ 2)"
   by (import transc COS_SIN_SQ)
 
-lemma COS_ATN_NZ: "ALL x::real. cos (atn x) ~= 0"
+lemma COS_ATN_NZ: "transc.cos (atn x) ~= 0"
   by (import transc COS_ATN_NZ)
 
-lemma COS_ASN_NZ: "ALL x::real. - 1 < x & x < 1 --> cos (asn x) ~= 0"
+lemma COS_ASN_NZ: "- 1 < x & x < 1 ==> transc.cos (asn x) ~= 0"
   by (import transc COS_ASN_NZ)
 
-lemma SIN_ACS_NZ: "ALL x::real. - 1 < x & x < 1 --> sin (acs x) ~= 0"
+lemma SIN_ACS_NZ: "- 1 < x & x < 1 ==> transc.sin (acs x) ~= 0"
   by (import transc SIN_ACS_NZ)
 
-lemma COS_SIN_SQRT: "ALL x::real. 0 <= cos x --> cos x = sqrt (1 - sin x ^ 2)"
+lemma COS_SIN_SQRT: "0 <= transc.cos x ==> transc.cos x = transc.sqrt (1 - transc.sin x ^ 2)"
   by (import transc COS_SIN_SQRT)
 
-lemma SIN_COS_SQRT: "ALL x::real. 0 <= sin x --> sin x = sqrt (1 - cos x ^ 2)"
+lemma SIN_COS_SQRT: "0 <= transc.sin x ==> transc.sin x = transc.sqrt (1 - transc.cos x ^ 2)"
   by (import transc SIN_COS_SQRT)
 
-lemma DIFF_LN: "ALL x>0. diffl ln (inverse x) x"
+lemma DIFF_LN: "0 < x ==> diffl transc.ln (inverse x) x"
   by (import transc DIFF_LN)
 
-lemma DIFF_ASN_LEMMA: "ALL x::real. - 1 < x & x < 1 --> diffl asn (inverse (cos (asn x))) x"
+lemma DIFF_ASN_LEMMA: "- 1 < x & x < 1 ==> diffl asn (inverse (transc.cos (asn x))) x"
   by (import transc DIFF_ASN_LEMMA)
 
-lemma DIFF_ASN: "ALL x::real. - 1 < x & x < 1 --> diffl asn (inverse (sqrt (1 - x ^ 2))) x"
+lemma DIFF_ASN: "- 1 < x & x < 1 ==> diffl asn (inverse (transc.sqrt (1 - x ^ 2))) x"
   by (import transc DIFF_ASN)
 
-lemma DIFF_ACS_LEMMA: "ALL x::real. - 1 < x & x < 1 --> diffl acs (inverse (- sin (acs x))) x"
+lemma DIFF_ACS_LEMMA: "- 1 < x & x < 1 ==> diffl acs (inverse (- transc.sin (acs x))) x"
   by (import transc DIFF_ACS_LEMMA)
 
-lemma DIFF_ACS: "ALL x::real. - 1 < x & x < 1 --> diffl acs (- inverse (sqrt (1 - x ^ 2))) x"
+lemma DIFF_ACS: "- 1 < x & x < 1 ==> diffl acs (- inverse (transc.sqrt (1 - x ^ 2))) x"
   by (import transc DIFF_ACS)
 
-lemma DIFF_ATN: "ALL x::real. diffl atn (inverse (1 + x ^ 2)) x"
+lemma DIFF_ATN: "diffl atn (inverse (1 + x ^ 2)) x"
   by (import transc DIFF_ATN)
 
-definition division :: "real * real => (nat => real) => bool" where 
-  "(op ==::(real * real => (nat => real) => bool)
-        => (real * real => (nat => real) => bool) => prop)
- (division::real * real => (nat => real) => bool)
- ((split::(real => real => (nat => real) => bool)
-          => real * real => (nat => real) => bool)
-   (%(a::real) (b::real) D::nat => real.
-       (op &::bool => bool => bool)
-        ((op =::real => real => bool) (D (0::nat)) a)
-        ((Ex::(nat => bool) => bool)
-          (%N::nat.
-              (op &::bool => bool => bool)
-               ((All::(nat => bool) => bool)
-                 (%n::nat.
-                     (op -->::bool => bool => bool)
-                      ((op <::nat => nat => bool) n N)
-                      ((op <::real => real => bool) (D n)
-                        (D ((Suc::nat => nat) n)))))
-               ((All::(nat => bool) => bool)
-                 (%n::nat.
-                     (op -->::bool => bool => bool)
-                      ((op <=::nat => nat => bool) N n)
-                      ((op =::real => real => bool) (D n) b)))))))"
-
-lemma division: "(All::(real => bool) => bool)
- (%a::real.
-     (All::(real => bool) => bool)
-      (%b::real.
-          (All::((nat => real) => bool) => bool)
-           (%D::nat => real.
-               (op =::bool => bool => bool)
-                ((division::real * real => (nat => real) => bool)
-                  ((Pair::real => real => real * real) a b) D)
-                ((op &::bool => bool => bool)
-                  ((op =::real => real => bool) (D (0::nat)) a)
-                  ((Ex::(nat => bool) => bool)
-                    (%N::nat.
-                        (op &::bool => bool => bool)
-                         ((All::(nat => bool) => bool)
-                           (%n::nat.
-                               (op -->::bool => bool => bool)
-                                ((op <::nat => nat => bool) n N)
-                                ((op <::real => real => bool) (D n)
-                                  (D ((Suc::nat => nat) n)))))
-                         ((All::(nat => bool) => bool)
-                           (%n::nat.
-                               (op -->::bool => bool => bool)
-                                ((op <=::nat => nat => bool) N n)
-                                ((op =::real => real => bool) (D n)
-                                  b)))))))))"
+definition
+  division :: "real * real => (nat => real) => bool"  where
+  "division ==
+%(a, b) D.
+   D 0 = a & (EX N. (ALL n<N. D n < D (Suc n)) & (ALL n>=N. D n = b))"
+
+lemma division: "division (a, b) D =
+(D 0 = a & (EX N. (ALL n<N. D n < D (Suc n)) & (ALL n>=N. D n = b)))"
   by (import transc division)
 
-definition dsize :: "(nat => real) => nat" where 
-  "(op ==::((nat => real) => nat) => ((nat => real) => nat) => prop)
- (dsize::(nat => real) => nat)
- (%D::nat => real.
-     (Eps::(nat => bool) => nat)
-      (%N::nat.
-          (op &::bool => bool => bool)
-           ((All::(nat => bool) => bool)
-             (%n::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <::nat => nat => bool) n N)
-                  ((op <::real => real => bool) (D n)
-                    (D ((Suc::nat => nat) n)))))
-           ((All::(nat => bool) => bool)
-             (%n::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) N n)
-                  ((op =::real => real => bool) (D n) (D N))))))"
-
-lemma dsize: "(All::((nat => real) => bool) => bool)
- (%D::nat => real.
-     (op =::nat => nat => bool) ((dsize::(nat => real) => nat) D)
-      ((Eps::(nat => bool) => nat)
-        (%N::nat.
-            (op &::bool => bool => bool)
-             ((All::(nat => bool) => bool)
-               (%n::nat.
-                   (op -->::bool => bool => bool)
-                    ((op <::nat => nat => bool) n N)
-                    ((op <::real => real => bool) (D n)
-                      (D ((Suc::nat => nat) n)))))
-             ((All::(nat => bool) => bool)
-               (%n::nat.
-                   (op -->::bool => bool => bool)
-                    ((op <=::nat => nat => bool) N n)
-                    ((op =::real => real => bool) (D n) (D N)))))))"
+definition
+  dsize :: "(nat => real) => nat"  where
+  "dsize == %D. SOME N. (ALL n<N. D n < D (Suc n)) & (ALL n>=N. D n = D N)"
+
+lemma dsize: "dsize D = (SOME N. (ALL n<N. D n < D (Suc n)) & (ALL n>=N. D n = D N))"
   by (import transc dsize)
 
-definition tdiv :: "real * real => (nat => real) * (nat => real) => bool" where 
+definition
+  tdiv :: "real * real => (nat => real) * (nat => real) => bool"  where
   "tdiv ==
-%(a::real, b::real) (D::nat => real, p::nat => real).
-   division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n))"
-
-lemma tdiv: "ALL (a::real) (b::real) (D::nat => real) p::nat => real.
-   tdiv (a, b) (D, p) =
-   (division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n)))"
+%(a, b) (D, p). division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n))"
+
+lemma tdiv: "tdiv (a, b) (D, p) =
+(division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n)))"
   by (import transc tdiv)
 
-definition gauge :: "(real => bool) => (real => real) => bool" where 
-  "gauge == %(E::real => bool) g::real => real. ALL x::real. E x --> 0 < g x"
-
-lemma gauge: "ALL (E::real => bool) g::real => real.
-   gauge E g = (ALL x::real. E x --> 0 < g x)"
+definition
+  gauge :: "(real => bool) => (real => real) => bool"  where
+  "gauge == %E g. ALL x. E x --> 0 < g x"
+
+lemma gauge: "gauge E g = (ALL x. E x --> 0 < g x)"
   by (import transc gauge)
 
-definition fine :: "(real => real) => (nat => real) * (nat => real) => bool" where 
-  "(op ==::((real => real) => (nat => real) * (nat => real) => bool)
-        => ((real => real) => (nat => real) * (nat => real) => bool)
-           => prop)
- (fine::(real => real) => (nat => real) * (nat => real) => bool)
- (%g::real => real.
-     (split::((nat => real) => (nat => real) => bool)
-             => (nat => real) * (nat => real) => bool)
-      (%(D::nat => real) p::nat => real.
-          (All::(nat => bool) => bool)
-           (%n::nat.
-               (op -->::bool => bool => bool)
-                ((op <::nat => nat => bool) n
-                  ((dsize::(nat => real) => nat) D))
-                ((op <::real => real => bool)
-                  ((op -::real => real => real) (D ((Suc::nat => nat) n))
-                    (D n))
-                  (g (p n))))))"
-
-lemma fine: "(All::((real => real) => bool) => bool)
- (%g::real => real.
-     (All::((nat => real) => bool) => bool)
-      (%D::nat => real.
-          (All::((nat => real) => bool) => bool)
-           (%p::nat => real.
-               (op =::bool => bool => bool)
-                ((fine::(real => real)
-                        => (nat => real) * (nat => real) => bool)
-                  g ((Pair::(nat => real)
-                            => (nat => real)
-                               => (nat => real) * (nat => real))
-                      D p))
-                ((All::(nat => bool) => bool)
-                  (%n::nat.
-                      (op -->::bool => bool => bool)
-                       ((op <::nat => nat => bool) n
-                         ((dsize::(nat => real) => nat) D))
-                       ((op <::real => real => bool)
-                         ((op -::real => real => real)
-                           (D ((Suc::nat => nat) n)) (D n))
-                         (g (p n))))))))"
+definition
+  fine :: "(real => real) => (nat => real) * (nat => real) => bool"  where
+  "fine == %g (D, p). ALL n<dsize D. D (Suc n) - D n < g (p n)"
+
+lemma fine: "fine g (D, p) = (ALL n<dsize D. D (Suc n) - D n < g (p n))"
   by (import transc fine)
 
-definition rsum :: "(nat => real) * (nat => real) => (real => real) => real" where 
-  "rsum ==
-%(D::nat => real, p::nat => real) f::real => real.
-   real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))"
-
-lemma rsum: "ALL (D::nat => real) (p::nat => real) f::real => real.
-   rsum (D, p) f =
-   real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))"
+definition
+  rsum :: "(nat => real) * (nat => real) => (real => real) => real"  where
+  "rsum == %(D, p) f. real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))"
+
+lemma rsum: "rsum (D, p) f = real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))"
   by (import transc rsum)
 
-definition Dint :: "real * real => (real => real) => real => bool" where 
+definition
+  Dint :: "real * real => (real => real) => real => bool"  where
   "Dint ==
-%(a::real, b::real) (f::real => real) k::real.
+%(a, b) f k.
    ALL e>0.
-      EX g::real => real.
-         gauge (%x::real. a <= x & x <= b) g &
-         (ALL (D::nat => real) p::nat => real.
-             tdiv (a, b) (D, p) & fine g (D, p) -->
-             abs (rsum (D, p) f - k) < e)"
-
-lemma Dint: "ALL (a::real) (b::real) (f::real => real) k::real.
-   Dint (a, b) f k =
-   (ALL e>0.
-       EX g::real => real.
-          gauge (%x::real. a <= x & x <= b) g &
-          (ALL (D::nat => real) p::nat => real.
+      EX g. gauge (%x. a <= x & x <= b) g &
+            (ALL D p.
+                tdiv (a, b) (D, p) & fine g (D, p) -->
+                abs (rsum (D, p) f - k) < e)"
+
+lemma Dint: "Dint (a, b) f k =
+(ALL e>0.
+    EX g. gauge (%x. a <= x & x <= b) g &
+          (ALL D p.
               tdiv (a, b) (D, p) & fine g (D, p) -->
               abs (rsum (D, p) f - k) < e))"
   by (import transc Dint)
 
-lemma DIVISION_0: "ALL (a::real) b::real. a = b --> dsize (%n::nat. if n = 0 then a else b) = 0"
+lemma DIVISION_0: "a = b ==> dsize (%n. if n = 0 then a else b) = 0"
   by (import transc DIVISION_0)
 
-lemma DIVISION_1: "ALL (a::real) b::real. a < b --> dsize (%n::nat. if n = 0 then a else b) = 1"
+lemma DIVISION_1: "a < b ==> dsize (%n. if n = 0 then a else b) = 1"
   by (import transc DIVISION_1)
 
-lemma DIVISION_SINGLE: "ALL (a::real) b::real.
-   a <= b --> division (a, b) (%n::nat. if n = 0 then a else b)"
+lemma DIVISION_SINGLE: "a <= b ==> division (a, b) (%n. if n = 0 then a else b)"
   by (import transc DIVISION_SINGLE)
 
-lemma DIVISION_LHS: "ALL (D::nat => real) (a::real) b::real. division (a, b) D --> D 0 = a"
+lemma DIVISION_LHS: "division (a, b) D ==> D 0 = a"
   by (import transc DIVISION_LHS)
 
-lemma DIVISION_THM: "(All::((nat => real) => bool) => bool)
- (%D::nat => real.
-     (All::(real => bool) => bool)
-      (%a::real.
-          (All::(real => bool) => bool)
-           (%b::real.
-               (op =::bool => bool => bool)
-                ((division::real * real => (nat => real) => bool)
-                  ((Pair::real => real => real * real) a b) D)
-                ((op &::bool => bool => bool)
-                  ((op =::real => real => bool) (D (0::nat)) a)
-                  ((op &::bool => bool => bool)
-                    ((All::(nat => bool) => bool)
-                      (%n::nat.
-                          (op -->::bool => bool => bool)
-                           ((op <::nat => nat => bool) n
-                             ((dsize::(nat => real) => nat) D))
-                           ((op <::real => real => bool) (D n)
-                             (D ((Suc::nat => nat) n)))))
-                    ((All::(nat => bool) => bool)
-                      (%n::nat.
-                          (op -->::bool => bool => bool)
-                           ((op <=::nat => nat => bool)
-                             ((dsize::(nat => real) => nat) D) n)
-                           ((op =::real => real => bool) (D n) b))))))))"
+lemma DIVISION_THM: "division (a, b) D =
+(D 0 = a & (ALL n<dsize D. D n < D (Suc n)) & (ALL n>=dsize D. D n = b))"
   by (import transc DIVISION_THM)
 
-lemma DIVISION_RHS: "ALL (D::nat => real) (a::real) b::real.
-   division (a, b) D --> D (dsize D) = b"
+lemma DIVISION_RHS: "division (a, b) D ==> D (dsize D) = b"
   by (import transc DIVISION_RHS)
 
-lemma DIVISION_LT_GEN: "ALL (D::nat => real) (a::real) (b::real) (m::nat) n::nat.
-   division (a, b) D & m < n & n <= dsize D --> D m < D n"
+lemma DIVISION_LT_GEN: "division (a, b) D & m < n & n <= dsize D ==> D m < D n"
   by (import transc DIVISION_LT_GEN)
 
-lemma DIVISION_LT: "(All::((nat => real) => bool) => bool)
- (%D::nat => real.
-     (All::(real => bool) => bool)
-      (%a::real.
-          (All::(real => bool) => bool)
-           (%b::real.
-               (op -->::bool => bool => bool)
-                ((division::real * real => (nat => real) => bool)
-                  ((Pair::real => real => real * real) a b) D)
-                ((All::(nat => bool) => bool)
-                  (%n::nat.
-                      (op -->::bool => bool => bool)
-                       ((op <::nat => nat => bool) n
-                         ((dsize::(nat => real) => nat) D))
-                       ((op <::real => real => bool) (D (0::nat))
-                         (D ((Suc::nat => nat) n))))))))"
+lemma DIVISION_LT: "[| division (a, b) D; n < dsize D |] ==> D 0 < D (Suc n)"
   by (import transc DIVISION_LT)
 
-lemma DIVISION_LE: "ALL (D::nat => real) (a::real) b::real. division (a, b) D --> a <= b"
+lemma DIVISION_LE: "division (a, b) D ==> a <= b"
   by (import transc DIVISION_LE)
 
-lemma DIVISION_GT: "(All::((nat => real) => bool) => bool)
- (%D::nat => real.
-     (All::(real => bool) => bool)
-      (%a::real.
-          (All::(real => bool) => bool)
-           (%b::real.
-               (op -->::bool => bool => bool)
-                ((division::real * real => (nat => real) => bool)
-                  ((Pair::real => real => real * real) a b) D)
-                ((All::(nat => bool) => bool)
-                  (%n::nat.
-                      (op -->::bool => bool => bool)
-                       ((op <::nat => nat => bool) n
-                         ((dsize::(nat => real) => nat) D))
-                       ((op <::real => real => bool) (D n)
-                         (D ((dsize::(nat => real) => nat) D))))))))"
+lemma DIVISION_GT: "[| division (a, b) D; n < dsize D |] ==> D n < D (dsize D)"
   by (import transc DIVISION_GT)
 
-lemma DIVISION_EQ: "ALL (D::nat => real) (a::real) b::real.
-   division (a, b) D --> (a = b) = (dsize D = 0)"
+lemma DIVISION_EQ: "division (a, b) D ==> (a = b) = (dsize D = 0)"
   by (import transc DIVISION_EQ)
 
-lemma DIVISION_LBOUND: "ALL (D::nat => real) (a::real) b::real.
-   division (a, b) D --> (ALL r::nat. a <= D r)"
+lemma DIVISION_LBOUND: "division (a, b) D ==> a <= D r"
   by (import transc DIVISION_LBOUND)
 
-lemma DIVISION_LBOUND_LT: "ALL (D::nat => real) (a::real) b::real.
-   division (a, b) D & dsize D ~= 0 --> (ALL n::nat. a < D (Suc n))"
+lemma DIVISION_LBOUND_LT: "division (a, b) D & dsize D ~= 0 ==> a < D (Suc n)"
   by (import transc DIVISION_LBOUND_LT)
 
-lemma DIVISION_UBOUND: "ALL (D::nat => real) (a::real) b::real.
-   division (a, b) D --> (ALL r::nat. D r <= b)"
+lemma DIVISION_UBOUND: "division (a, b) D ==> D r <= b"
   by (import transc DIVISION_UBOUND)
 
-lemma DIVISION_UBOUND_LT: "ALL (D::nat => real) (a::real) (b::real) n::nat.
-   division (a, b) D & n < dsize D --> D n < b"
+lemma DIVISION_UBOUND_LT: "division (a, b) D & n < dsize D ==> D n < b"
   by (import transc DIVISION_UBOUND_LT)
 
-lemma DIVISION_APPEND: "ALL (a::real) (b::real) c::real.
-   (EX (D1::nat => real) p1::nat => real.
-       tdiv (a, b) (D1, p1) & fine (g::real => real) (D1, p1)) &
-   (EX (D2::nat => real) p2::nat => real.
-       tdiv (b, c) (D2, p2) & fine g (D2, p2)) -->
-   (EX (x::nat => real) p::nat => real. tdiv (a, c) (x, p) & fine g (x, p))"
+lemma DIVISION_APPEND: "(EX D1 p1. tdiv (a, b) (D1, p1) & fine g (D1, p1)) &
+(EX D2 p2. tdiv (b, c) (D2, p2) & fine g (D2, p2))
+==> EX x p. tdiv (a, c) (x, p) & fine g (x, p)"
   by (import transc DIVISION_APPEND)
 
-lemma DIVISION_EXISTS: "ALL (a::real) (b::real) g::real => real.
-   a <= b & gauge (%x::real. a <= x & x <= b) g -->
-   (EX (D::nat => real) p::nat => real. tdiv (a, b) (D, p) & fine g (D, p))"
+lemma DIVISION_EXISTS: "a <= b & gauge (%x. a <= x & x <= b) g
+==> EX D p. tdiv (a, b) (D, p) & fine g (D, p)"
   by (import transc DIVISION_EXISTS)
 
-lemma GAUGE_MIN: "ALL (E::real => bool) (g1::real => real) g2::real => real.
-   gauge E g1 & gauge E g2 -->
-   gauge E (%x::real. if g1 x < g2 x then g1 x else g2 x)"
+lemma GAUGE_MIN: "gauge E g1 & gauge E g2 ==> gauge E (%x. if g1 x < g2 x then g1 x else g2 x)"
   by (import transc GAUGE_MIN)
 
-lemma FINE_MIN: "ALL (g1::real => real) (g2::real => real) (D::nat => real) p::nat => real.
-   fine (%x::real. if g1 x < g2 x then g1 x else g2 x) (D, p) -->
-   fine g1 (D, p) & fine g2 (D, p)"
+lemma FINE_MIN: "fine (%x. if g1 x < g2 x then g1 x else g2 x) (D, p)
+==> fine g1 (D, p) & fine g2 (D, p)"
   by (import transc FINE_MIN)
 
-lemma DINT_UNIQ: "ALL (a::real) (b::real) (f::real => real) (k1::real) k2::real.
-   a <= b & Dint (a, b) f k1 & Dint (a, b) f k2 --> k1 = k2"
+lemma DINT_UNIQ: "a <= b & Dint (a, b) f k1 & Dint (a, b) f k2 ==> k1 = k2"
   by (import transc DINT_UNIQ)
 
-lemma INTEGRAL_NULL: "ALL (f::real => real) a::real. Dint (a, a) f 0"
+lemma INTEGRAL_NULL: "Dint (a, a) f 0"
   by (import transc INTEGRAL_NULL)
 
-lemma FTC1: "ALL (f::real => real) (f'::real => real) (a::real) b::real.
-   a <= b & (ALL x::real. a <= x & x <= b --> diffl f (f' x) x) -->
-   Dint (a, b) f' (f b - f a)"
+lemma FTC1: "a <= b & (ALL x. a <= x & x <= b --> diffl f (f' x) x)
+==> Dint (a, b) f' (f b - f a)"
   by (import transc FTC1)
 
-lemma MCLAURIN: "ALL (f::real => real) (diff::nat => real => real) (h::real) n::nat.
-   0 < h &
-   0 < n &
-   diff 0 = f &
-   (ALL (m::nat) t::real.
-       m < n & 0 <= t & t <= h --> diffl (diff m) (diff (Suc m) t) t) -->
-   (EX t>0.
+lemma MCLAURIN: "0 < h &
+0 < n &
+diff 0 = f &
+(ALL m t. m < n & 0 <= t & t <= h --> diffl (diff m) (diff (Suc m) t) t)
+==> EX t>0.
        t < h &
        f h =
-       real.sum (0, n) (%m::nat. diff m 0 / real (FACT m) * h ^ m) +
-       diff n t / real (FACT n) * h ^ n)"
+       real.sum (0, n) (%m. diff m 0 / real (FACT m) * h ^ m) +
+       diff n t / real (FACT n) * h ^ n"
   by (import transc MCLAURIN)
 
-lemma MCLAURIN_NEG: "ALL (f::real => real) (diff::nat => real => real) (h::real) n::nat.
-   h < 0 &
-   0 < n &
-   diff 0 = f &
-   (ALL (m::nat) t::real.
-       m < n & h <= t & t <= 0 --> diffl (diff m) (diff (Suc m) t) t) -->
-   (EX t::real.
-       h < t &
+lemma MCLAURIN_NEG: "h < 0 &
+0 < n &
+diff 0 = f &
+(ALL m t. m < n & h <= t & t <= 0 --> diffl (diff m) (diff (Suc m) t) t)
+==> EX t>h.
        t < 0 &
        f h =
-       real.sum (0, n) (%m::nat. diff m 0 / real (FACT m) * h ^ m) +
-       diff n t / real (FACT n) * h ^ n)"
+       real.sum (0, n) (%m. diff m 0 / real (FACT m) * h ^ m) +
+       diff n t / real (FACT n) * h ^ n"
   by (import transc MCLAURIN_NEG)
 
-lemma MCLAURIN_ALL_LT: "ALL (f::real => real) diff::nat => real => real.
-   diff 0 = f &
-   (ALL (m::nat) x::real. diffl (diff m) (diff (Suc m) x) x) -->
-   (ALL (x::real) n::nat.
-       x ~= 0 & 0 < n -->
-       (EX t::real.
-           0 < abs t &
-           abs t < abs x &
-           f x =
-           real.sum (0, n) (%m::nat. diff m 0 / real (FACT m) * x ^ m) +
-           diff n t / real (FACT n) * x ^ n))"
+lemma MCLAURIN_ALL_LT: "[| diff 0 = f & (ALL m x. diffl (diff m) (diff (Suc m) x) x);
+   x ~= 0 & 0 < n |]
+==> EX t. 0 < abs t &
+          abs t < abs x &
+          f x =
+          real.sum (0, n) (%m. diff m 0 / real (FACT m) * x ^ m) +
+          diff n t / real (FACT n) * x ^ n"
   by (import transc MCLAURIN_ALL_LT)
 
-lemma MCLAURIN_ZERO: "ALL (diff::nat => real => real) (n::nat) x::real.
-   x = 0 & 0 < n -->
-   real.sum (0, n) (%m::nat. diff m 0 / real (FACT m) * x ^ m) = diff 0 0"
+lemma MCLAURIN_ZERO: "(x::real) = (0::real) & (0::nat) < (n::nat)
+==> real.sum (0::nat, n)
+     (%m::nat.
+         (diff::nat => real => real) m (0::real) / real (FACT m) * x ^ m) =
+    diff (0::nat) (0::real)"
   by (import transc MCLAURIN_ZERO)
 
-lemma MCLAURIN_ALL_LE: "ALL (f::real => real) diff::nat => real => real.
-   diff 0 = f &
-   (ALL (m::nat) x::real. diffl (diff m) (diff (Suc m) x) x) -->
-   (ALL (x::real) n::nat.
-       EX t::real.
-          abs t <= abs x &
+lemma MCLAURIN_ALL_LE: "diff 0 = f & (ALL m x. diffl (diff m) (diff (Suc m) x) x)
+==> EX t. abs t <= abs x &
           f x =
-          real.sum (0, n) (%m::nat. diff m 0 / real (FACT m) * x ^ m) +
-          diff n t / real (FACT n) * x ^ n)"
+          real.sum (0, n) (%m. diff m 0 / real (FACT m) * x ^ m) +
+          diff n t / real (FACT n) * x ^ n"
   by (import transc MCLAURIN_ALL_LE)
 
-lemma MCLAURIN_EXP_LT: "ALL (x::real) n::nat.
-   x ~= 0 & 0 < n -->
-   (EX xa::real.
+lemma MCLAURIN_EXP_LT: "x ~= 0 & 0 < n
+==> EX xa.
        0 < abs xa &
        abs xa < abs x &
-       exp x =
-       real.sum (0, n) (%m::nat. x ^ m / real (FACT m)) +
-       exp xa / real (FACT n) * x ^ n)"
+       transc.exp x =
+       real.sum (0, n) (%m. x ^ m / real (FACT m)) +
+       transc.exp xa / real (FACT n) * x ^ n"
   by (import transc MCLAURIN_EXP_LT)
 
-lemma MCLAURIN_EXP_LE: "ALL (x::real) n::nat.
-   EX xa::real.
-      abs xa <= abs x &
-      exp x =
-      real.sum (0, n) (%m::nat. x ^ m / real (FACT m)) +
-      exp xa / real (FACT n) * x ^ n"
+lemma MCLAURIN_EXP_LE: "EX xa.
+   abs xa <= abs x &
+   transc.exp x =
+   real.sum (0, n) (%m. x ^ m / real (FACT m)) +
+   transc.exp xa / real (FACT n) * x ^ n"
   by (import transc MCLAURIN_EXP_LE)
 
-lemma DIFF_LN_COMPOSITE: "ALL (g::real => real) (m::real) x::real.
-   diffl g m x & 0 < g x -->
-   diffl (%x::real. ln (g x)) (inverse (g x) * m) x"
+lemma DIFF_LN_COMPOSITE: "diffl g m x & 0 < g x ==> diffl (%x. transc.ln (g x)) (inverse (g x) * m) x"
   by (import transc DIFF_LN_COMPOSITE)
 
 ;end_setup
@@ -3200,15 +2558,14 @@
 consts
   poly :: "real list => real => real" 
 
-specification (poly_primdef: poly) poly_def: "(ALL x::real. poly [] x = 0) &
-(ALL (h::real) (t::real list) x::real. poly (h # t) x = h + x * poly t x)"
+specification (poly_primdef: poly) poly_def: "(ALL x. poly [] x = 0) & (ALL h t x. poly (h # t) x = h + x * poly t x)"
   by (import poly poly_def)
 
 consts
   poly_add :: "real list => real list => real list" 
 
-specification (poly_add_primdef: poly_add) poly_add_def: "(ALL l2::real list. poly_add [] l2 = l2) &
-(ALL (h::real) (t::real list) l2::real list.
+specification (poly_add_primdef: poly_add) poly_add_def: "(ALL l2. poly_add [] l2 = l2) &
+(ALL h t l2.
     poly_add (h # t) l2 =
     (if l2 = [] then h # t else (h + hd l2) # poly_add t (tl l2)))"
   by (import poly poly_add_def)
@@ -3216,8 +2573,7 @@
 consts
   "##" :: "real => real list => real list" ("##")
 
-specification ("##") poly_cmul_def: "(ALL c::real. ## c [] = []) &
-(ALL (c::real) (h::real) t::real list. ## c (h # t) = c * h # ## c t)"
+specification ("##") poly_cmul_def: "(ALL c. ## c [] = []) & (ALL c h t. ## c (h # t) = c * h # ## c t)"
   by (import poly poly_cmul_def)
 
 consts
@@ -3232,8 +2588,8 @@
 consts
   poly_mul :: "real list => real list => real list" 
 
-specification (poly_mul_primdef: poly_mul) poly_mul_def: "(ALL l2::real list. poly_mul [] l2 = []) &
-(ALL (h::real) (t::real list) l2::real list.
+specification (poly_mul_primdef: poly_mul) poly_mul_def: "(ALL l2. poly_mul [] l2 = []) &
+(ALL h t l2.
     poly_mul (h # t) l2 =
     (if t = [] then ## h l2 else poly_add (## h l2) (0 # poly_mul t l2)))"
   by (import poly poly_mul_def)
@@ -3241,514 +2597,351 @@
 consts
   poly_exp :: "real list => nat => real list" 
 
-specification (poly_exp_primdef: poly_exp) poly_exp_def: "(ALL p::real list. poly_exp p 0 = [1]) &
-(ALL (p::real list) n::nat. poly_exp p (Suc n) = poly_mul p (poly_exp p n))"
+specification (poly_exp_primdef: poly_exp) poly_exp_def: "(ALL p. poly_exp p 0 = [1]) &
+(ALL p n. poly_exp p (Suc n) = poly_mul p (poly_exp p n))"
   by (import poly poly_exp_def)
 
 consts
   poly_diff_aux :: "nat => real list => real list" 
 
-specification (poly_diff_aux_primdef: poly_diff_aux) poly_diff_aux_def: "(ALL n::nat. poly_diff_aux n [] = []) &
-(ALL (n::nat) (h::real) t::real list.
-    poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)"
+specification (poly_diff_aux_primdef: poly_diff_aux) poly_diff_aux_def: "(ALL n. poly_diff_aux n [] = []) &
+(ALL n h t. poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)"
   by (import poly poly_diff_aux_def)
 
-definition diff :: "real list => real list" where 
-  "diff == %l::real list. if l = [] then [] else poly_diff_aux 1 (tl l)"
-
-lemma poly_diff_def: "ALL l::real list. diff l = (if l = [] then [] else poly_diff_aux 1 (tl l))"
+definition
+  diff :: "real list => real list"  where
+  "diff == %l. if l = [] then [] else poly_diff_aux 1 (tl l)"
+
+lemma poly_diff_def: "diff l = (if l = [] then [] else poly_diff_aux 1 (tl l))"
   by (import poly poly_diff_def)
 
-lemma POLY_ADD_CLAUSES: "poly_add [] (p2::real list) = p2 &
-poly_add (p1::real list) [] = p1 &
-poly_add ((h1::real) # (t1::real list)) ((h2::real) # (t2::real list)) =
-(h1 + h2) # poly_add t1 t2"
+lemma POLY_ADD_CLAUSES: "poly_add [] p2 = p2 &
+poly_add p1 [] = p1 &
+poly_add (h1 # t1) (h2 # t2) = (h1 + h2) # poly_add t1 t2"
   by (import poly POLY_ADD_CLAUSES)
 
-lemma POLY_CMUL_CLAUSES: "## (c::real) [] = [] & ## c ((h::real) # (t::real list)) = c * h # ## c t"
+lemma POLY_CMUL_CLAUSES: "## c [] = [] & ## c (h # t) = c * h # ## c t"
   by (import poly POLY_CMUL_CLAUSES)
 
-lemma POLY_NEG_CLAUSES: "poly_neg [] = [] & poly_neg ((h::real) # (t::real list)) = - h # poly_neg t"
+lemma POLY_NEG_CLAUSES: "poly_neg [] = [] & poly_neg (h # t) = - h # poly_neg t"
   by (import poly POLY_NEG_CLAUSES)
 
-lemma POLY_MUL_CLAUSES: "poly_mul [] (p2::real list) = [] &
-poly_mul [h1::real] p2 = ## h1 p2 &
-poly_mul (h1 # (k1::real) # (t1::real list)) p2 =
-poly_add (## h1 p2) (0 # poly_mul (k1 # t1) p2)"
+lemma POLY_MUL_CLAUSES: "poly_mul [] p2 = [] &
+poly_mul [h1] p2 = ## h1 p2 &
+poly_mul (h1 # k1 # t1) p2 = poly_add (## h1 p2) (0 # poly_mul (k1 # t1) p2)"
   by (import poly POLY_MUL_CLAUSES)
 
-lemma POLY_DIFF_CLAUSES: "diff [] = [] &
-diff [c::real] = [] & diff ((h::real) # (t::real list)) = poly_diff_aux 1 t"
+lemma POLY_DIFF_CLAUSES: "diff [] = [] & diff [c] = [] & diff (h # t) = poly_diff_aux 1 t"
   by (import poly POLY_DIFF_CLAUSES)
 
-lemma POLY_ADD: "ALL (t::real list) (p2::real list) x::real.
-   poly (poly_add t p2) x = poly t x + poly p2 x"
+lemma POLY_ADD: "poly (poly_add t p2) x = poly t x + poly p2 x"
   by (import poly POLY_ADD)
 
-lemma POLY_CMUL: "ALL (t::real list) (c::real) x::real. poly (## c t) x = c * poly t x"
+lemma POLY_CMUL: "poly (## c t) x = c * poly t x"
   by (import poly POLY_CMUL)
 
-lemma POLY_NEG: "ALL (x::real list) xa::real. poly (poly_neg x) xa = - poly x xa"
+lemma POLY_NEG: "poly (poly_neg x) xa = - poly x xa"
   by (import poly POLY_NEG)
 
-lemma POLY_MUL: "ALL (x::real) (t::real list) p2::real list.
-   poly (poly_mul t p2) x = poly t x * poly p2 x"
+lemma POLY_MUL: "poly (poly_mul t p2) x = poly t x * poly p2 x"
   by (import poly POLY_MUL)
 
-lemma POLY_EXP: "ALL (p::real list) (n::nat) x::real. poly (poly_exp p n) x = poly p x ^ n"
+lemma POLY_EXP: "poly (poly_exp p n) x = poly p x ^ n"
   by (import poly POLY_EXP)
 
-lemma POLY_DIFF_LEMMA: "ALL (t::real list) (n::nat) x::real.
-   diffl (%x::real. x ^ Suc n * poly t x)
-    (x ^ n * poly (poly_diff_aux (Suc n) t) x) x"
+lemma POLY_DIFF_LEMMA: "diffl (%x. x ^ Suc n * poly t x) (x ^ n * poly (poly_diff_aux (Suc n) t) x)
+ x"
   by (import poly POLY_DIFF_LEMMA)
 
-lemma POLY_DIFF: "ALL (t::real list) x::real. diffl (poly t) (poly (diff t) x) x"
+lemma POLY_DIFF: "diffl (poly t) (poly (diff t) x) x"
   by (import poly POLY_DIFF)
 
-lemma POLY_DIFFERENTIABLE: "ALL l::real list. All (differentiable (poly l))"
+lemma POLY_DIFFERENTIABLE: "lim.differentiable (poly l) x"
   by (import poly POLY_DIFFERENTIABLE)
 
-lemma POLY_CONT: "ALL l::real list. All (contl (poly l))"
+lemma POLY_CONT: "contl (poly l) x"
   by (import poly POLY_CONT)
 
-lemma POLY_IVT_POS: "ALL (x::real list) (xa::real) xb::real.
-   xa < xb & poly x xa < 0 & 0 < poly x xb -->
-   (EX xc::real. xa < xc & xc < xb & poly x xc = 0)"
+lemma POLY_IVT_POS: "xa < xb & poly x xa < 0 & 0 < poly x xb
+==> EX xc>xa. xc < xb & poly x xc = 0"
   by (import poly POLY_IVT_POS)
 
-lemma POLY_IVT_NEG: "ALL (p::real list) (a::real) b::real.
-   a < b & 0 < poly p a & poly p b < 0 -->
-   (EX x::real. a < x & x < b & poly p x = 0)"
+lemma POLY_IVT_NEG: "a < b & 0 < poly p a & poly p b < 0 ==> EX x>a. x < b & poly p x = 0"
   by (import poly POLY_IVT_NEG)
 
-lemma POLY_MVT: "ALL (p::real list) (a::real) b::real.
-   a < b -->
-   (EX x::real.
-       a < x & x < b & poly p b - poly p a = (b - a) * poly (diff p) x)"
+lemma POLY_MVT: "a < b ==> EX x>a. x < b & poly p b - poly p a = (b - a) * poly (diff p) x"
   by (import poly POLY_MVT)
 
-lemma POLY_ADD_RZERO: "ALL x::real list. poly (poly_add x []) = poly x"
+lemma POLY_ADD_RZERO: "poly (poly_add x []) = poly x"
   by (import poly POLY_ADD_RZERO)
 
-lemma POLY_MUL_ASSOC: "ALL (x::real list) (xa::real list) xb::real list.
-   poly (poly_mul x (poly_mul xa xb)) = poly (poly_mul (poly_mul x xa) xb)"
+lemma POLY_MUL_ASSOC: "poly (poly_mul x (poly_mul xa xb)) = poly (poly_mul (poly_mul x xa) xb)"
   by (import poly POLY_MUL_ASSOC)
 
-lemma POLY_EXP_ADD: "ALL (x::nat) (xa::nat) xb::real list.
-   poly (poly_exp xb (xa + x)) =
-   poly (poly_mul (poly_exp xb xa) (poly_exp xb x))"
+lemma POLY_EXP_ADD: "poly (poly_exp xb (xa + x)) =
+poly (poly_mul (poly_exp xb xa) (poly_exp xb x))"
   by (import poly POLY_EXP_ADD)
 
-lemma POLY_DIFF_AUX_ADD: "ALL (t::real list) (p2::real list) n::nat.
-   poly (poly_diff_aux n (poly_add t p2)) =
-   poly (poly_add (poly_diff_aux n t) (poly_diff_aux n p2))"
+lemma POLY_DIFF_AUX_ADD: "poly (poly_diff_aux n (poly_add t p2)) =
+poly (poly_add (poly_diff_aux n t) (poly_diff_aux n p2))"
   by (import poly POLY_DIFF_AUX_ADD)
 
-lemma POLY_DIFF_AUX_CMUL: "ALL (t::real list) (c::real) n::nat.
-   poly (poly_diff_aux n (## c t)) = poly (## c (poly_diff_aux n t))"
+lemma POLY_DIFF_AUX_CMUL: "poly (poly_diff_aux n (## c t)) = poly (## c (poly_diff_aux n t))"
   by (import poly POLY_DIFF_AUX_CMUL)
 
-lemma POLY_DIFF_AUX_NEG: "ALL (x::real list) xa::nat.
-   poly (poly_diff_aux xa (poly_neg x)) =
-   poly (poly_neg (poly_diff_aux xa x))"
+lemma POLY_DIFF_AUX_NEG: "poly (poly_diff_aux xa (poly_neg x)) = poly (poly_neg (poly_diff_aux xa x))"
   by (import poly POLY_DIFF_AUX_NEG)
 
-lemma POLY_DIFF_AUX_MUL_LEMMA: "ALL (t::real list) n::nat.
-   poly (poly_diff_aux (Suc n) t) = poly (poly_add (poly_diff_aux n t) t)"
+lemma POLY_DIFF_AUX_MUL_LEMMA: "poly (poly_diff_aux (Suc n) t) = poly (poly_add (poly_diff_aux n t) t)"
   by (import poly POLY_DIFF_AUX_MUL_LEMMA)
 
-lemma POLY_DIFF_ADD: "ALL (t::real list) p2::real list.
-   poly (diff (poly_add t p2)) = poly (poly_add (diff t) (diff p2))"
+lemma POLY_DIFF_ADD: "poly (diff (poly_add t p2)) = poly (poly_add (diff t) (diff p2))"
   by (import poly POLY_DIFF_ADD)
 
-lemma POLY_DIFF_CMUL: "ALL (t::real list) c::real. poly (diff (## c t)) = poly (## c (diff t))"
+lemma POLY_DIFF_CMUL: "poly (diff (## c t)) = poly (## c (diff t))"
   by (import poly POLY_DIFF_CMUL)
 
-lemma POLY_DIFF_NEG: "ALL x::real list. poly (diff (poly_neg x)) = poly (poly_neg (diff x))"
+lemma POLY_DIFF_NEG: "poly (diff (poly_neg x)) = poly (poly_neg (diff x))"
   by (import poly POLY_DIFF_NEG)
 
-lemma POLY_DIFF_MUL_LEMMA: "ALL (x::real list) xa::real.
-   poly (diff (xa # x)) = poly (poly_add (0 # diff x) x)"
+lemma POLY_DIFF_MUL_LEMMA: "poly (diff (xa # x)) = poly (poly_add (0 # diff x) x)"
   by (import poly POLY_DIFF_MUL_LEMMA)
 
-lemma POLY_DIFF_MUL: "ALL (t::real list) p2::real list.
-   poly (diff (poly_mul t p2)) =
-   poly (poly_add (poly_mul t (diff p2)) (poly_mul (diff t) p2))"
+lemma POLY_DIFF_MUL: "poly (diff (poly_mul t p2)) =
+poly (poly_add (poly_mul t (diff p2)) (poly_mul (diff t) p2))"
   by (import poly POLY_DIFF_MUL)
 
-lemma POLY_DIFF_EXP: "ALL (p::real list) n::nat.
-   poly (diff (poly_exp p (Suc n))) =
-   poly (poly_mul (## (real (Suc n)) (poly_exp p n)) (diff p))"
+lemma POLY_DIFF_EXP: "poly (diff (poly_exp p (Suc n))) =
+poly (poly_mul (## (real (Suc n)) (poly_exp p n)) (diff p))"
   by (import poly POLY_DIFF_EXP)
 
-lemma POLY_DIFF_EXP_PRIME: "ALL (n::nat) a::real.
-   poly (diff (poly_exp [- a, 1] (Suc n))) =
-   poly (## (real (Suc n)) (poly_exp [- a, 1] n))"
+lemma POLY_DIFF_EXP_PRIME: "poly (diff (poly_exp [- a, 1] (Suc n))) =
+poly (## (real (Suc n)) (poly_exp [- a, 1] n))"
   by (import poly POLY_DIFF_EXP_PRIME)
 
-lemma POLY_LINEAR_REM: "ALL (t::real list) h::real.
-   EX (q::real list) r::real.
-      h # t = poly_add [r] (poly_mul [- (a::real), 1] q)"
+lemma POLY_LINEAR_REM: "EX q r. h # t = poly_add [r] (poly_mul [- a, 1] q)"
   by (import poly POLY_LINEAR_REM)
 
-lemma POLY_LINEAR_DIVIDES: "ALL (a::real) t::real list.
-   (poly t a = 0) = (t = [] | (EX q::real list. t = poly_mul [- a, 1] q))"
+lemma POLY_LINEAR_DIVIDES: "(poly t a = 0) = (t = [] | (EX q. t = poly_mul [- a, 1] q))"
   by (import poly POLY_LINEAR_DIVIDES)
 
-lemma POLY_LENGTH_MUL: "ALL x::real list. length (poly_mul [- (a::real), 1] x) = Suc (length x)"
+lemma POLY_LENGTH_MUL: "length (poly_mul [- a, 1] x) = Suc (length x)"
   by (import poly POLY_LENGTH_MUL)
 
-lemma POLY_ROOTS_INDEX_LEMMA: "(All::(nat => bool) => bool)
- (%n::nat.
-     (All::(real list => bool) => bool)
-      (%p::real list.
-          (op -->::bool => bool => bool)
-           ((op &::bool => bool => bool)
-             ((Not::bool => bool)
-               ((op =::(real => real) => (real => real) => bool)
-                 ((poly::real list => real => real) p)
-                 ((poly::real list => real => real) ([]::real list))))
-             ((op =::nat => nat => bool) ((size::real list => nat) p) n))
-           ((Ex::((nat => real) => bool) => bool)
-             (%i::nat => real.
-                 (All::(real => bool) => bool)
-                  (%x::real.
-                      (op -->::bool => bool => bool)
-                       ((op =::real => real => bool)
-                         ((poly::real list => real => real) p x) (0::real))
-                       ((Ex::(nat => bool) => bool)
-                         (%m::nat.
-                             (op &::bool => bool => bool)
-                              ((op <=::nat => nat => bool) m n)
-                              ((op =::real => real => bool) x (i m)))))))))"
+lemma POLY_ROOTS_INDEX_LEMMA: "poly p ~= poly [] & length p = n
+==> EX i. ALL x. poly p x = 0 --> (EX m<=n. x = i m)"
   by (import poly POLY_ROOTS_INDEX_LEMMA)
 
-lemma POLY_ROOTS_INDEX_LENGTH: "(All::(real list => bool) => bool)
- (%p::real list.
-     (op -->::bool => bool => bool)
-      ((Not::bool => bool)
-        ((op =::(real => real) => (real => real) => bool)
-          ((poly::real list => real => real) p)
-          ((poly::real list => real => real) ([]::real list))))
-      ((Ex::((nat => real) => bool) => bool)
-        (%i::nat => real.
-            (All::(real => bool) => bool)
-             (%x::real.
-                 (op -->::bool => bool => bool)
-                  ((op =::real => real => bool)
-                    ((poly::real list => real => real) p x) (0::real))
-                  ((Ex::(nat => bool) => bool)
-                    (%n::nat.
-                        (op &::bool => bool => bool)
-                         ((op <=::nat => nat => bool) n
-                           ((size::real list => nat) p))
-                         ((op =::real => real => bool) x (i n))))))))"
+lemma POLY_ROOTS_INDEX_LENGTH: "poly p ~= poly []
+==> EX i. ALL x. poly p x = 0 --> (EX n<=length p. x = i n)"
   by (import poly POLY_ROOTS_INDEX_LENGTH)
 
-lemma POLY_ROOTS_FINITE_LEMMA: "(All::(real list => bool) => bool)
- (%p::real list.
-     (op -->::bool => bool => bool)
-      ((Not::bool => bool)
-        ((op =::(real => real) => (real => real) => bool)
-          ((poly::real list => real => real) p)
-          ((poly::real list => real => real) ([]::real list))))
-      ((Ex::(nat => bool) => bool)
-        (%N::nat.
-            (Ex::((nat => real) => bool) => bool)
-             (%i::nat => real.
-                 (All::(real => bool) => bool)
-                  (%x::real.
-                      (op -->::bool => bool => bool)
-                       ((op =::real => real => bool)
-                         ((poly::real list => real => real) p x) (0::real))
-                       ((Ex::(nat => bool) => bool)
-                         (%n::nat.
-                             (op &::bool => bool => bool)
-                              ((op <::nat => nat => bool) n N)
-                              ((op =::real => real => bool) x (i n)))))))))"
+lemma POLY_ROOTS_FINITE_LEMMA: "poly (p::real list) ~= poly []
+==> EX (N::nat) i::nat => real.
+       ALL x::real. poly p x = (0::real) --> (EX n<N. x = i n)"
   by (import poly POLY_ROOTS_FINITE_LEMMA)
 
-lemma FINITE_LEMMA: "(All::((nat => real) => bool) => bool)
- (%i::nat => real.
-     (All::(nat => bool) => bool)
-      (%x::nat.
-          (All::((real => bool) => bool) => bool)
-           (%xa::real => bool.
-               (op -->::bool => bool => bool)
-                ((All::(real => bool) => bool)
-                  (%xb::real.
-                      (op -->::bool => bool => bool) (xa xb)
-                       ((Ex::(nat => bool) => bool)
-                         (%n::nat.
-                             (op &::bool => bool => bool)
-                              ((op <::nat => nat => bool) n x)
-                              ((op =::real => real => bool) xb (i n))))))
-                ((Ex::(real => bool) => bool)
-                  (%a::real.
-                      (All::(real => bool) => bool)
-                       (%x::real.
-                           (op -->::bool => bool => bool) (xa x)
-                            ((op <::real => real => bool) x a)))))))"
+lemma FINITE_LEMMA: "(!!xb::real. (xa::real => bool) xb ==> EX n<x::nat. xb = (i::nat => real) n)
+==> EX a::real. ALL x::real. xa x --> x < a"
   by (import poly FINITE_LEMMA)
 
-lemma POLY_ROOTS_FINITE: "(All::(real list => bool) => bool)
- (%p::real list.
-     (op =::bool => bool => bool)
-      ((Not::bool => bool)
-        ((op =::(real => real) => (real => real) => bool)
-          ((poly::real list => real => real) p)
-          ((poly::real list => real => real) ([]::real list))))
-      ((Ex::(nat => bool) => bool)
-        (%N::nat.
-            (Ex::((nat => real) => bool) => bool)
-             (%i::nat => real.
-                 (All::(real => bool) => bool)
-                  (%x::real.
-                      (op -->::bool => bool => bool)
-                       ((op =::real => real => bool)
-                         ((poly::real list => real => real) p x) (0::real))
-                       ((Ex::(nat => bool) => bool)
-                         (%n::nat.
-                             (op &::bool => bool => bool)
-                              ((op <::nat => nat => bool) n N)
-                              ((op =::real => real => bool) x (i n)))))))))"
+lemma POLY_ROOTS_FINITE: "(poly (p::real list) ~= poly []) =
+(EX (N::nat) i::nat => real.
+    ALL x::real. poly p x = (0::real) --> (EX n<N. x = i n))"
   by (import poly POLY_ROOTS_FINITE)
 
-lemma POLY_ENTIRE_LEMMA: "ALL (p::real list) q::real list.
-   poly p ~= poly [] & poly q ~= poly [] --> poly (poly_mul p q) ~= poly []"
+lemma POLY_ENTIRE_LEMMA: "poly p ~= poly [] & poly q ~= poly [] ==> poly (poly_mul p q) ~= poly []"
   by (import poly POLY_ENTIRE_LEMMA)
 
-lemma POLY_ENTIRE: "ALL (p::real list) q::real list.
-   (poly (poly_mul p q) = poly []) = (poly p = poly [] | poly q = poly [])"
+lemma POLY_ENTIRE: "(poly (poly_mul p q) = poly []) = (poly p = poly [] | poly q = poly [])"
   by (import poly POLY_ENTIRE)
 
-lemma POLY_MUL_LCANCEL: "ALL (x::real list) (xa::real list) xb::real list.
-   (poly (poly_mul x xa) = poly (poly_mul x xb)) =
-   (poly x = poly [] | poly xa = poly xb)"
+lemma POLY_MUL_LCANCEL: "(poly (poly_mul x xa) = poly (poly_mul x xb)) =
+(poly x = poly [] | poly xa = poly xb)"
   by (import poly POLY_MUL_LCANCEL)
 
-lemma POLY_EXP_EQ_0: "ALL (p::real list) n::nat.
-   (poly (poly_exp p n) = poly []) = (poly p = poly [] & n ~= 0)"
+lemma POLY_EXP_EQ_0: "(poly (poly_exp p n) = poly []) = (poly p = poly [] & n ~= 0)"
   by (import poly POLY_EXP_EQ_0)
 
-lemma POLY_PRIME_EQ_0: "ALL a::real. poly [a, 1] ~= poly []"
+lemma POLY_PRIME_EQ_0: "poly [a, 1] ~= poly []"
   by (import poly POLY_PRIME_EQ_0)
 
-lemma POLY_EXP_PRIME_EQ_0: "ALL (a::real) n::nat. poly (poly_exp [a, 1] n) ~= poly []"
+lemma POLY_EXP_PRIME_EQ_0: "poly (poly_exp [a, 1] n) ~= poly []"
   by (import poly POLY_EXP_PRIME_EQ_0)
 
-lemma POLY_ZERO_LEMMA: "ALL (h::real) t::real list.
-   poly (h # t) = poly [] --> h = 0 & poly t = poly []"
+lemma POLY_ZERO_LEMMA: "poly (h # t) = poly [] ==> h = 0 & poly t = poly []"
   by (import poly POLY_ZERO_LEMMA)
 
-lemma POLY_ZERO: "ALL t::real list. (poly t = poly []) = list_all (%c::real. c = 0) t"
+lemma POLY_ZERO: "(poly t = poly []) = list_all (%c. c = 0) t"
   by (import poly POLY_ZERO)
 
-lemma POLY_DIFF_AUX_ISZERO: "ALL (t::real list) n::nat.
-   list_all (%c::real. c = 0) (poly_diff_aux (Suc n) t) =
-   list_all (%c::real. c = 0) t"
+lemma POLY_DIFF_AUX_ISZERO: "list_all (%c. c = 0) (poly_diff_aux (Suc n) t) = list_all (%c. c = 0) t"
   by (import poly POLY_DIFF_AUX_ISZERO)
 
-lemma POLY_DIFF_ISZERO: "ALL x::real list.
-   poly (diff x) = poly [] --> (EX h::real. poly x = poly [h])"
+lemma POLY_DIFF_ISZERO: "poly (diff x) = poly [] ==> EX h. poly x = poly [h]"
   by (import poly POLY_DIFF_ISZERO)
 
-lemma POLY_DIFF_ZERO: "ALL x::real list. poly x = poly [] --> poly (diff x) = poly []"
+lemma POLY_DIFF_ZERO: "poly x = poly [] ==> poly (diff x) = poly []"
   by (import poly POLY_DIFF_ZERO)
 
-lemma POLY_DIFF_WELLDEF: "ALL (p::real list) q::real list.
-   poly p = poly q --> poly (diff p) = poly (diff q)"
+lemma POLY_DIFF_WELLDEF: "poly p = poly q ==> poly (diff p) = poly (diff q)"
   by (import poly POLY_DIFF_WELLDEF)
 
-definition poly_divides :: "real list => real list => bool" where 
-  "poly_divides ==
-%(p1::real list) p2::real list.
-   EX q::real list. poly p2 = poly (poly_mul p1 q)"
-
-lemma poly_divides: "ALL (p1::real list) p2::real list.
-   poly_divides p1 p2 = (EX q::real list. poly p2 = poly (poly_mul p1 q))"
+definition
+  poly_divides :: "real list => real list => bool"  where
+  "poly_divides == %p1 p2. EX q. poly p2 = poly (poly_mul p1 q)"
+
+lemma poly_divides: "poly_divides p1 p2 = (EX q. poly p2 = poly (poly_mul p1 q))"
   by (import poly poly_divides)
 
-lemma POLY_PRIMES: "ALL (a::real) (p::real list) q::real list.
-   poly_divides [a, 1] (poly_mul p q) =
-   (poly_divides [a, 1] p | poly_divides [a, 1] q)"
+lemma POLY_PRIMES: "poly_divides [a, 1] (poly_mul p q) =
+(poly_divides [a, 1] p | poly_divides [a, 1] q)"
   by (import poly POLY_PRIMES)
 
-lemma POLY_DIVIDES_REFL: "ALL p::real list. poly_divides p p"
+lemma POLY_DIVIDES_REFL: "poly_divides p p"
   by (import poly POLY_DIVIDES_REFL)
 
-lemma POLY_DIVIDES_TRANS: "ALL (p::real list) (q::real list) r::real list.
-   poly_divides p q & poly_divides q r --> poly_divides p r"
+lemma POLY_DIVIDES_TRANS: "poly_divides p q & poly_divides q r ==> poly_divides p r"
   by (import poly POLY_DIVIDES_TRANS)
 
-lemma POLY_DIVIDES_EXP: "ALL (p::real list) (m::nat) n::nat.
-   m <= n --> poly_divides (poly_exp p m) (poly_exp p n)"
+lemma POLY_DIVIDES_EXP: "m <= n ==> poly_divides (poly_exp p m) (poly_exp p n)"
   by (import poly POLY_DIVIDES_EXP)
 
-lemma POLY_EXP_DIVIDES: "ALL (p::real list) (q::real list) (m::nat) n::nat.
-   poly_divides (poly_exp p n) q & m <= n --> poly_divides (poly_exp p m) q"
+lemma POLY_EXP_DIVIDES: "poly_divides (poly_exp p n) q & m <= n ==> poly_divides (poly_exp p m) q"
   by (import poly POLY_EXP_DIVIDES)
 
-lemma POLY_DIVIDES_ADD: "ALL (p::real list) (q::real list) r::real list.
-   poly_divides p q & poly_divides p r --> poly_divides p (poly_add q r)"
+lemma POLY_DIVIDES_ADD: "poly_divides p q & poly_divides p r ==> poly_divides p (poly_add q r)"
   by (import poly POLY_DIVIDES_ADD)
 
-lemma POLY_DIVIDES_SUB: "ALL (p::real list) (q::real list) r::real list.
-   poly_divides p q & poly_divides p (poly_add q r) --> poly_divides p r"
+lemma POLY_DIVIDES_SUB: "poly_divides p q & poly_divides p (poly_add q r) ==> poly_divides p r"
   by (import poly POLY_DIVIDES_SUB)
 
-lemma POLY_DIVIDES_SUB2: "ALL (p::real list) (q::real list) r::real list.
-   poly_divides p r & poly_divides p (poly_add q r) --> poly_divides p q"
+lemma POLY_DIVIDES_SUB2: "poly_divides p r & poly_divides p (poly_add q r) ==> poly_divides p q"
   by (import poly POLY_DIVIDES_SUB2)
 
-lemma POLY_DIVIDES_ZERO: "ALL (p::real list) q::real list. poly p = poly [] --> poly_divides q p"
+lemma POLY_DIVIDES_ZERO: "poly p = poly [] ==> poly_divides q p"
   by (import poly POLY_DIVIDES_ZERO)
 
-lemma POLY_ORDER_EXISTS: "ALL (a::real) (d::nat) p::real list.
-   length p = d & poly p ~= poly [] -->
-   (EX x::nat.
-       poly_divides (poly_exp [- a, 1] x) p &
-       ~ poly_divides (poly_exp [- a, 1] (Suc x)) p)"
+lemma POLY_ORDER_EXISTS: "length p = d & poly p ~= poly []
+==> EX x. poly_divides (poly_exp [- a, 1] x) p &
+          ~ poly_divides (poly_exp [- a, 1] (Suc x)) p"
   by (import poly POLY_ORDER_EXISTS)
 
-lemma POLY_ORDER: "ALL (p::real list) a::real.
-   poly p ~= poly [] -->
-   (EX! n::nat.
+lemma POLY_ORDER: "poly p ~= poly []
+==> EX! n.
        poly_divides (poly_exp [- a, 1] n) p &
-       ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)"
+       ~ poly_divides (poly_exp [- a, 1] (Suc n)) p"
   by (import poly POLY_ORDER)
 
-definition poly_order :: "real => real list => nat" where 
+definition
+  poly_order :: "real => real list => nat"  where
   "poly_order ==
-%(a::real) p::real list.
-   SOME n::nat.
-      poly_divides (poly_exp [- a, 1] n) p &
-      ~ poly_divides (poly_exp [- a, 1] (Suc n)) p"
-
-lemma poly_order: "ALL (a::real) p::real list.
-   poly_order a p =
-   (SOME n::nat.
-       poly_divides (poly_exp [- a, 1] n) p &
-       ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)"
+%a p. SOME n.
+         poly_divides (poly_exp [- a, 1] n) p &
+         ~ poly_divides (poly_exp [- a, 1] (Suc n)) p"
+
+lemma poly_order: "poly_order a p =
+(SOME n.
+    poly_divides (poly_exp [- a, 1] n) p &
+    ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)"
   by (import poly poly_order)
 
-lemma ORDER: "ALL (p::real list) (a::real) n::nat.
-   (poly_divides (poly_exp [- a, 1] n) p &
-    ~ poly_divides (poly_exp [- a, 1] (Suc n)) p) =
-   (n = poly_order a p & poly p ~= poly [])"
+lemma ORDER: "(poly_divides (poly_exp [- a, 1] n) p &
+ ~ poly_divides (poly_exp [- a, 1] (Suc n)) p) =
+(n = poly_order a p & poly p ~= poly [])"
   by (import poly ORDER)
 
-lemma ORDER_THM: "ALL (p::real list) a::real.
-   poly p ~= poly [] -->
-   poly_divides (poly_exp [- a, 1] (poly_order a p)) p &
-   ~ poly_divides (poly_exp [- a, 1] (Suc (poly_order a p))) p"
+lemma ORDER_THM: "poly p ~= poly []
+==> poly_divides (poly_exp [- a, 1] (poly_order a p)) p &
+    ~ poly_divides (poly_exp [- a, 1] (Suc (poly_order a p))) p"
   by (import poly ORDER_THM)
 
-lemma ORDER_UNIQUE: "ALL (p::real list) (a::real) n::nat.
-   poly p ~= poly [] &
-   poly_divides (poly_exp [- a, 1] n) p &
-   ~ poly_divides (poly_exp [- a, 1] (Suc n)) p -->
-   n = poly_order a p"
+lemma ORDER_UNIQUE: "poly p ~= poly [] &
+poly_divides (poly_exp [- a, 1] n) p &
+~ poly_divides (poly_exp [- a, 1] (Suc n)) p
+==> n = poly_order a p"
   by (import poly ORDER_UNIQUE)
 
-lemma ORDER_POLY: "ALL (p::real list) (q::real list) a::real.
-   poly p = poly q --> poly_order a p = poly_order a q"
+lemma ORDER_POLY: "poly p = poly q ==> poly_order a p = poly_order a q"
   by (import poly ORDER_POLY)
 
-lemma ORDER_ROOT: "ALL (p::real list) a::real.
-   (poly p a = 0) = (poly p = poly [] | poly_order a p ~= 0)"
+lemma ORDER_ROOT: "(poly p a = 0) = (poly p = poly [] | poly_order a p ~= 0)"
   by (import poly ORDER_ROOT)
 
-lemma ORDER_DIVIDES: "ALL (p::real list) (a::real) n::nat.
-   poly_divides (poly_exp [- a, 1] n) p =
-   (poly p = poly [] | n <= poly_order a p)"
+lemma ORDER_DIVIDES: "poly_divides (poly_exp [- a, 1] n) p =
+(poly p = poly [] | n <= poly_order a p)"
   by (import poly ORDER_DIVIDES)
 
-lemma ORDER_DECOMP: "ALL (p::real list) a::real.
-   poly p ~= poly [] -->
-   (EX x::real list.
-       poly p = poly (poly_mul (poly_exp [- a, 1] (poly_order a p)) x) &
-       ~ poly_divides [- a, 1] x)"
+lemma ORDER_DECOMP: "poly p ~= poly []
+==> EX x. poly p = poly (poly_mul (poly_exp [- a, 1] (poly_order a p)) x) &
+          ~ poly_divides [- a, 1] x"
   by (import poly ORDER_DECOMP)
 
-lemma ORDER_MUL: "ALL (a::real) (p::real list) q::real list.
-   poly (poly_mul p q) ~= poly [] -->
-   poly_order a (poly_mul p q) = poly_order a p + poly_order a q"
+lemma ORDER_MUL: "poly (poly_mul p q) ~= poly []
+==> poly_order a (poly_mul p q) = poly_order a p + poly_order a q"
   by (import poly ORDER_MUL)
 
-lemma ORDER_DIFF: "ALL (p::real list) a::real.
-   poly (diff p) ~= poly [] & poly_order a p ~= 0 -->
-   poly_order a p = Suc (poly_order a (diff p))"
+lemma ORDER_DIFF: "poly (diff p) ~= poly [] & poly_order a p ~= 0
+==> poly_order a p = Suc (poly_order a (diff p))"
   by (import poly ORDER_DIFF)
 
-lemma POLY_SQUAREFREE_DECOMP_ORDER: "ALL (p::real list) (q::real list) (d::real list) (e::real list)
-   (r::real list) s::real list.
-   poly (diff p) ~= poly [] &
-   poly p = poly (poly_mul q d) &
-   poly (diff p) = poly (poly_mul e d) &
-   poly d = poly (poly_add (poly_mul r p) (poly_mul s (diff p))) -->
-   (ALL a::real. poly_order a q = (if poly_order a p = 0 then 0 else 1))"
+lemma POLY_SQUAREFREE_DECOMP_ORDER: "poly (diff p) ~= poly [] &
+poly p = poly (poly_mul q d) &
+poly (diff p) = poly (poly_mul e d) &
+poly d = poly (poly_add (poly_mul r p) (poly_mul s (diff p)))
+==> poly_order a q = (if poly_order a p = 0 then 0 else 1)"
   by (import poly POLY_SQUAREFREE_DECOMP_ORDER)
 
-definition rsquarefree :: "real list => bool" where 
+definition
+  rsquarefree :: "real list => bool"  where
   "rsquarefree ==
-%p::real list.
-   poly p ~= poly [] &
-   (ALL a::real. poly_order a p = 0 | poly_order a p = 1)"
-
-lemma rsquarefree: "ALL p::real list.
-   rsquarefree p =
-   (poly p ~= poly [] &
-    (ALL a::real. poly_order a p = 0 | poly_order a p = 1))"
+%p. poly p ~= poly [] & (ALL a. poly_order a p = 0 | poly_order a p = 1)"
+
+lemma rsquarefree: "rsquarefree p =
+(poly p ~= poly [] & (ALL a. poly_order a p = 0 | poly_order a p = 1))"
   by (import poly rsquarefree)
 
-lemma RSQUAREFREE_ROOTS: "ALL p::real list.
-   rsquarefree p = (ALL a::real. ~ (poly p a = 0 & poly (diff p) a = 0))"
+lemma RSQUAREFREE_ROOTS: "rsquarefree p = (ALL a. ~ (poly p a = 0 & poly (diff p) a = 0))"
   by (import poly RSQUAREFREE_ROOTS)
 
-lemma RSQUAREFREE_DECOMP: "ALL (p::real list) a::real.
-   rsquarefree p & poly p a = 0 -->
-   (EX q::real list. poly p = poly (poly_mul [- a, 1] q) & poly q a ~= 0)"
+lemma RSQUAREFREE_DECOMP: "rsquarefree p & poly p a = 0
+==> EX q. poly p = poly (poly_mul [- a, 1] q) & poly q a ~= 0"
   by (import poly RSQUAREFREE_DECOMP)
 
-lemma POLY_SQUAREFREE_DECOMP: "ALL (p::real list) (q::real list) (d::real list) (e::real list)
-   (r::real list) s::real list.
-   poly (diff p) ~= poly [] &
-   poly p = poly (poly_mul q d) &
-   poly (diff p) = poly (poly_mul e d) &
-   poly d = poly (poly_add (poly_mul r p) (poly_mul s (diff p))) -->
-   rsquarefree q & (ALL x::real. (poly q x = 0) = (poly p x = 0))"
+lemma POLY_SQUAREFREE_DECOMP: "poly (diff p) ~= poly [] &
+poly p = poly (poly_mul q d) &
+poly (diff p) = poly (poly_mul e d) &
+poly d = poly (poly_add (poly_mul r p) (poly_mul s (diff p)))
+==> rsquarefree q & (ALL x. (poly q x = 0) = (poly p x = 0))"
   by (import poly POLY_SQUAREFREE_DECOMP)
 
 consts
   normalize :: "real list => real list" 
 
 specification (normalize) normalize: "normalize [] = [] &
-(ALL (h::real) t::real list.
+(ALL h t.
     normalize (h # t) =
     (if normalize t = [] then if h = 0 then [] else [h]
      else h # normalize t))"
   by (import poly normalize)
 
-lemma POLY_NORMALIZE: "ALL t::real list. poly (normalize t) = poly t"
+lemma POLY_NORMALIZE: "poly (normalize t) = poly t"
   by (import poly POLY_NORMALIZE)
 
-definition degree :: "real list => nat" where 
-  "degree == %p::real list. PRE (length (normalize p))"
-
-lemma degree: "ALL p::real list. degree p = PRE (length (normalize p))"
+definition
+  degree :: "real list => nat"  where
+  "degree == %p. PRE (length (normalize p))"
+
+lemma degree: "degree p = PRE (length (normalize p))"
   by (import poly degree)
 
-lemma DEGREE_ZERO: "ALL p::real list. poly p = poly [] --> degree p = 0"
+lemma DEGREE_ZERO: "poly p = poly [] ==> degree p = 0"
   by (import poly DEGREE_ZERO)
 
-lemma POLY_ROOTS_FINITE_SET: "ALL p::real list.
-   poly p ~= poly [] --> FINITE (GSPEC (%x::real. (x, poly p x = 0)))"
+lemma POLY_ROOTS_FINITE_SET: "poly p ~= poly [] ==> FINITE (GSPEC (%x. (x, poly p x = 0)))"
   by (import poly POLY_ROOTS_FINITE_SET)
 
-lemma POLY_MONO: "ALL (x::real) (k::real) xa::real list.
-   abs x <= k --> abs (poly xa x) <= poly (map abs xa) k"
+lemma POLY_MONO: "abs x <= k ==> abs (poly xa x) <= poly (map abs xa) k"
   by (import poly POLY_MONO)
 
 ;end_setup
--- a/src/HOL/Import/HOL/HOL4Vec.thy	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/HOL4Vec.thy	Tue Sep 06 22:41:35 2011 -0700
@@ -4,139 +4,96 @@
 
 ;setup_theory res_quan
 
-lemma RES_FORALL_CONJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool.
-   RES_FORALL P (%i::'a::type. Q i & R i) =
-   (RES_FORALL P Q & RES_FORALL P R)"
+lemma RES_FORALL_CONJ_DIST: "RES_FORALL P (%i. Q i & R i) = (RES_FORALL P Q & RES_FORALL P R)"
   by (import res_quan RES_FORALL_CONJ_DIST)
 
-lemma RES_FORALL_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool.
-   RES_FORALL (%j::'a::type. P j | Q j) R =
-   (RES_FORALL P R & RES_FORALL Q R)"
+lemma RES_FORALL_DISJ_DIST: "RES_FORALL (%j. P j | Q j) R = (RES_FORALL P R & RES_FORALL Q R)"
   by (import res_quan RES_FORALL_DISJ_DIST)
 
-lemma RES_FORALL_UNIQUE: "ALL (x::'a::type => bool) xa::'a::type. RES_FORALL (op = xa) x = x xa"
+lemma RES_FORALL_UNIQUE: "RES_FORALL (op = xa) x = x xa"
   by (import res_quan RES_FORALL_UNIQUE)
 
-lemma RES_FORALL_FORALL: "ALL (P::'a::type => bool) (R::'a::type => 'b::type => bool) x::'b::type.
-   (ALL x::'b::type. RES_FORALL P (%i::'a::type. R i x)) =
-   RES_FORALL P (%i::'a::type. All (R i))"
+lemma RES_FORALL_FORALL: "(ALL x::'b.
+    RES_FORALL (P::'a => bool) (%i::'a. (R::'a => 'b => bool) i x)) =
+RES_FORALL P (%i::'a. All (R i))"
   by (import res_quan RES_FORALL_FORALL)
 
-lemma RES_FORALL_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool)
-   R::'a::type => 'b::type => bool.
-   RES_FORALL P (%i::'a::type. RES_FORALL Q (R i)) =
-   RES_FORALL Q (%j::'b::type. RES_FORALL P (%i::'a::type. R i j))"
+lemma RES_FORALL_REORDER: "RES_FORALL P (%i. RES_FORALL Q (R i)) =
+RES_FORALL Q (%j. RES_FORALL P (%i. R i j))"
   by (import res_quan RES_FORALL_REORDER)
 
-lemma RES_FORALL_EMPTY: "All (RES_FORALL EMPTY)"
+lemma RES_FORALL_EMPTY: "RES_FORALL EMPTY x"
   by (import res_quan RES_FORALL_EMPTY)
 
-lemma RES_FORALL_UNIV: "ALL p::'a::type => bool. RES_FORALL pred_set.UNIV p = All p"
+lemma RES_FORALL_UNIV: "RES_FORALL pred_set.UNIV p = All p"
   by (import res_quan RES_FORALL_UNIV)
 
-lemma RES_FORALL_NULL: "ALL (p::'a::type => bool) m::bool.
-   RES_FORALL p (%x::'a::type. m) = (p = EMPTY | m)"
+lemma RES_FORALL_NULL: "RES_FORALL p (%x. m) = (p = EMPTY | m)"
   by (import res_quan RES_FORALL_NULL)
 
-lemma RES_EXISTS_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool.
-   RES_EXISTS P (%i::'a::type. Q i | R i) =
-   (RES_EXISTS P Q | RES_EXISTS P R)"
+lemma RES_EXISTS_DISJ_DIST: "RES_EXISTS P (%i. Q i | R i) = (RES_EXISTS P Q | RES_EXISTS P R)"
   by (import res_quan RES_EXISTS_DISJ_DIST)
 
-lemma RES_DISJ_EXISTS_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool.
-   RES_EXISTS (%i::'a::type. P i | Q i) R =
-   (RES_EXISTS P R | RES_EXISTS Q R)"
+lemma RES_DISJ_EXISTS_DIST: "RES_EXISTS (%i. P i | Q i) R = (RES_EXISTS P R | RES_EXISTS Q R)"
   by (import res_quan RES_DISJ_EXISTS_DIST)
 
-lemma RES_EXISTS_EQUAL: "ALL (x::'a::type => bool) xa::'a::type. RES_EXISTS (op = xa) x = x xa"
+lemma RES_EXISTS_EQUAL: "RES_EXISTS (op = xa) x = x xa"
   by (import res_quan RES_EXISTS_EQUAL)
 
-lemma RES_EXISTS_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool)
-   R::'a::type => 'b::type => bool.
-   RES_EXISTS P (%i::'a::type. RES_EXISTS Q (R i)) =
-   RES_EXISTS Q (%j::'b::type. RES_EXISTS P (%i::'a::type. R i j))"
+lemma RES_EXISTS_REORDER: "RES_EXISTS P (%i. RES_EXISTS Q (R i)) =
+RES_EXISTS Q (%j. RES_EXISTS P (%i. R i j))"
   by (import res_quan RES_EXISTS_REORDER)
 
-lemma RES_EXISTS_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS EMPTY p"
+lemma RES_EXISTS_EMPTY: "~ RES_EXISTS EMPTY p"
   by (import res_quan RES_EXISTS_EMPTY)
 
-lemma RES_EXISTS_UNIV: "ALL p::'a::type => bool. RES_EXISTS pred_set.UNIV p = Ex p"
+lemma RES_EXISTS_UNIV: "RES_EXISTS pred_set.UNIV p = Ex p"
   by (import res_quan RES_EXISTS_UNIV)
 
-lemma RES_EXISTS_NULL: "ALL (p::'a::type => bool) m::bool.
-   RES_EXISTS p (%x::'a::type. m) = (p ~= EMPTY & m)"
+lemma RES_EXISTS_NULL: "RES_EXISTS p (%x. m) = (p ~= EMPTY & m)"
   by (import res_quan RES_EXISTS_NULL)
 
-lemma RES_EXISTS_ALT: "ALL (p::'a::type => bool) m::'a::type => bool.
-   RES_EXISTS p m = (IN (RES_SELECT p m) p & m (RES_SELECT p m))"
+lemma RES_EXISTS_ALT: "RES_EXISTS p m = (IN (RES_SELECT p m) p & m (RES_SELECT p m))"
   by (import res_quan RES_EXISTS_ALT)
 
-lemma RES_EXISTS_UNIQUE_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS_UNIQUE EMPTY p"
+lemma RES_EXISTS_UNIQUE_EMPTY: "~ RES_EXISTS_UNIQUE EMPTY p"
   by (import res_quan RES_EXISTS_UNIQUE_EMPTY)
 
-lemma RES_EXISTS_UNIQUE_UNIV: "ALL p::'a::type => bool. RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p"
+lemma RES_EXISTS_UNIQUE_UNIV: "RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p"
   by (import res_quan RES_EXISTS_UNIQUE_UNIV)
 
-lemma RES_EXISTS_UNIQUE_NULL: "ALL (p::'a::type => bool) m::bool.
-   RES_EXISTS_UNIQUE p (%x::'a::type. m) =
-   ((EX x::'a::type. p = INSERT x EMPTY) & m)"
+lemma RES_EXISTS_UNIQUE_NULL: "RES_EXISTS_UNIQUE p (%x. m) = ((EX x. p = INSERT x EMPTY) & m)"
   by (import res_quan RES_EXISTS_UNIQUE_NULL)
 
-lemma RES_EXISTS_UNIQUE_ALT: "ALL (p::'a::type => bool) m::'a::type => bool.
-   RES_EXISTS_UNIQUE p m =
-   RES_EXISTS p
-    (%x::'a::type. m x & RES_FORALL p (%y::'a::type. m y --> y = x))"
+lemma RES_EXISTS_UNIQUE_ALT: "RES_EXISTS_UNIQUE p m =
+RES_EXISTS p (%x. m x & RES_FORALL p (%y. m y --> y = x))"
   by (import res_quan RES_EXISTS_UNIQUE_ALT)
 
-lemma RES_SELECT_EMPTY: "ALL p::'a::type => bool. RES_SELECT EMPTY p = (SOME x::'a::type. False)"
+lemma RES_SELECT_EMPTY: "RES_SELECT EMPTY p = (SOME x. False)"
   by (import res_quan RES_SELECT_EMPTY)
 
-lemma RES_SELECT_UNIV: "ALL p::'a::type => bool. RES_SELECT pred_set.UNIV p = Eps p"
+lemma RES_SELECT_UNIV: "RES_SELECT pred_set.UNIV p = Eps p"
   by (import res_quan RES_SELECT_UNIV)
 
-lemma RES_ABSTRACT: "ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type.
-   IN x p --> RES_ABSTRACT p m x = m x"
+lemma RES_ABSTRACT: "IN x p ==> RES_ABSTRACT p m x = m x"
   by (import res_quan RES_ABSTRACT)
 
-lemma RES_ABSTRACT_EQUAL: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type)
-   m2::'a::type => 'b::type.
-   (ALL x::'a::type. IN x p --> m1 x = m2 x) -->
-   RES_ABSTRACT p m1 = RES_ABSTRACT p m2"
+lemma RES_ABSTRACT_EQUAL: "(!!x. IN x p ==> m1 x = m2 x) ==> RES_ABSTRACT p m1 = RES_ABSTRACT p m2"
   by (import res_quan RES_ABSTRACT_EQUAL)
 
-lemma RES_ABSTRACT_IDEMPOT: "ALL (p::'a::type => bool) m::'a::type => 'b::type.
-   RES_ABSTRACT p (RES_ABSTRACT p m) = RES_ABSTRACT p m"
+lemma RES_ABSTRACT_IDEMPOT: "RES_ABSTRACT p (RES_ABSTRACT p m) = RES_ABSTRACT p m"
   by (import res_quan RES_ABSTRACT_IDEMPOT)
 
-lemma RES_ABSTRACT_EQUAL_EQ: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type)
-   m2::'a::type => 'b::type.
-   (RES_ABSTRACT p m1 = RES_ABSTRACT p m2) =
-   (ALL x::'a::type. IN x p --> m1 x = m2 x)"
+lemma RES_ABSTRACT_EQUAL_EQ: "(RES_ABSTRACT p m1 = RES_ABSTRACT p m2) = (ALL x. IN x p --> m1 x = m2 x)"
   by (import res_quan RES_ABSTRACT_EQUAL_EQ)
 
 ;end_setup
 
 ;setup_theory word_base
 
-typedef (open) ('a) word = "(Collect::('a::type list recspace => bool) => 'a::type list recspace set)
- (%x::'a::type list recspace.
-     (All::(('a::type list recspace => bool) => bool) => bool)
-      (%word::'a::type list recspace => bool.
-          (op -->::bool => bool => bool)
-           ((All::('a::type list recspace => bool) => bool)
-             (%a0::'a::type list recspace.
-                 (op -->::bool => bool => bool)
-                  ((Ex::('a::type list => bool) => bool)
-                    (%a::'a::type list.
-                        (op =::'a::type list recspace
-                               => 'a::type list recspace => bool)
-                         a0 ((CONSTR::nat
-=> 'a::type list
-   => (nat => 'a::type list recspace) => 'a::type list recspace)
-                              (0::nat) a
-                              (%n::nat. BOTTOM::'a::type list recspace))))
-                  (word a0)))
-           (word x)))" 
+typedef (open) ('a) word = "{x. ALL word.
+       (ALL a0. (EX a. a0 = CONSTR 0 a (%n. BOTTOM)) --> word a0) -->
+       word x} :: ('a::type list recspace set)"
   by (rule typedef_helper,import word_base word_TY_DEF)
 
 lemmas word_TY_DEF = typedef_hol2hol4 [OF type_definition_word]
@@ -145,11 +102,11 @@
   mk_word :: "'a list recspace => 'a word" 
   dest_word :: "'a word => 'a list recspace" 
 
-specification (dest_word mk_word) word_repfns: "(ALL a::'a::type word. mk_word (dest_word a) = a) &
-(ALL r::'a::type list recspace.
-    (ALL word::'a::type list recspace => bool.
-        (ALL a0::'a::type list recspace.
-            (EX a::'a::type list. a0 = CONSTR 0 a (%n::nat. BOTTOM)) -->
+specification (dest_word mk_word) word_repfns: "(ALL a::'a word. mk_word (dest_word a) = a) &
+(ALL r::'a list recspace.
+    (ALL word::'a list recspace => bool.
+        (ALL a0::'a list recspace.
+            (EX a::'a list. a0 = CONSTR (0::nat) a (%n::nat. BOTTOM)) -->
             word a0) -->
         word r) =
     (dest_word (mk_word r) = r))"
@@ -159,12 +116,13 @@
   word_base0 :: "'a list => 'a word" 
 
 defs
-  word_base0_primdef: "word_base0 == %a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM))"
+  word_base0_primdef: "word_base0 == %a. mk_word (CONSTR 0 a (%n. BOTTOM))"
 
-lemma word_base0_def: "word_base0 = (%a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM)))"
+lemma word_base0_def: "word_base0 = (%a. mk_word (CONSTR 0 a (%n. BOTTOM)))"
   by (import word_base word_base0_def)
 
-definition WORD :: "'a list => 'a word" where 
+definition
+  WORD :: "'a list => 'a word"  where
   "WORD == word_base0"
 
 lemma WORD: "WORD = word_base0"
@@ -173,601 +131,350 @@
 consts
   word_case :: "('a list => 'b) => 'a word => 'b" 
 
-specification (word_case_primdef: word_case) word_case_def: "ALL (f::'a::type list => 'b::type) a::'a::type list.
-   word_case f (WORD a) = f a"
+specification (word_case_primdef: word_case) word_case_def: "ALL f a. word_base.word_case f (WORD a) = f a"
   by (import word_base word_case_def)
 
 consts
   word_size :: "('a => nat) => 'a word => nat" 
 
-specification (word_size_primdef: word_size) word_size_def: "ALL (f::'a::type => nat) a::'a::type list.
-   word_size f (WORD a) = 1 + list_size f a"
+specification (word_size_primdef: word_size) word_size_def: "ALL f a. word_base.word_size f (WORD a) = 1 + HOL4Compat.list_size f a"
   by (import word_base word_size_def)
 
-lemma word_11: "ALL (a::'a::type list) a'::'a::type list. (WORD a = WORD a') = (a = a')"
+lemma word_11: "(WORD a = WORD a') = (a = a')"
   by (import word_base word_11)
 
-lemma word_case_cong: "ALL (M::'a::type word) (M'::'a::type word) f::'a::type list => 'b::type.
-   M = M' &
-   (ALL a::'a::type list.
-       M' = WORD a --> f a = (f'::'a::type list => 'b::type) a) -->
-   word_case f M = word_case f' M'"
+lemma word_case_cong: "M = M' & (ALL a. M' = WORD a --> f a = f' a)
+==> word_base.word_case f M = word_base.word_case f' M'"
   by (import word_base word_case_cong)
 
-lemma word_nchotomy: "ALL x::'a::type word. EX l::'a::type list. x = WORD l"
+lemma word_nchotomy: "EX l. x = WORD l"
   by (import word_base word_nchotomy)
 
-lemma word_Axiom: "ALL f::'a::type list => 'b::type.
-   EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a"
+lemma word_Axiom: "EX fn. ALL a. fn (WORD a) = f a"
   by (import word_base word_Axiom)
 
-lemma word_induction: "ALL P::'a::type word => bool. (ALL a::'a::type list. P (WORD a)) --> All P"
+lemma word_induction: "(!!a. P (WORD a)) ==> P x"
   by (import word_base word_induction)
 
-lemma word_Ax: "ALL f::'a::type list => 'b::type.
-   EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a"
+lemma word_Ax: "EX fn. ALL a. fn (WORD a) = f a"
   by (import word_base word_Ax)
 
-lemma WORD_11: "ALL (x::'a::type list) xa::'a::type list. (WORD x = WORD xa) = (x = xa)"
+lemma WORD_11: "(WORD x = WORD xa) = (x = xa)"
   by (import word_base WORD_11)
 
-lemma word_induct: "ALL x::'a::type word => bool. (ALL l::'a::type list. x (WORD l)) --> All x"
+lemma word_induct: "(!!l. x (WORD l)) ==> x xa"
   by (import word_base word_induct)
 
-lemma word_cases: "ALL x::'a::type word. EX l::'a::type list. x = WORD l"
+lemma word_cases: "EX l. x = WORD l"
   by (import word_base word_cases)
 
 consts
   WORDLEN :: "'a word => nat" 
 
-specification (WORDLEN) WORDLEN_DEF: "ALL l::'a::type list. WORDLEN (WORD l) = length l"
+specification (WORDLEN) WORDLEN_DEF: "ALL l. WORDLEN (WORD l) = length l"
   by (import word_base WORDLEN_DEF)
 
 consts
   PWORDLEN :: "nat => 'a word => bool" 
 
 defs
-  PWORDLEN_primdef: "PWORDLEN == %n::nat. GSPEC (%w::'a::type word. (w, WORDLEN w = n))"
+  PWORDLEN_primdef: "PWORDLEN == %n. GSPEC (%w. (w, WORDLEN w = n))"
 
-lemma PWORDLEN_def: "ALL n::nat. PWORDLEN n = GSPEC (%w::'a::type word. (w, WORDLEN w = n))"
+lemma PWORDLEN_def: "PWORDLEN n = GSPEC (%w. (w, WORDLEN w = n))"
   by (import word_base PWORDLEN_def)
 
-lemma IN_PWORDLEN: "ALL (n::nat) l::'a::type list. IN (WORD l) (PWORDLEN n) = (length l = n)"
+lemma IN_PWORDLEN: "IN (WORD l) (PWORDLEN n) = (length l = n)"
   by (import word_base IN_PWORDLEN)
 
-lemma PWORDLEN: "ALL (n::nat) w::'a::type word. IN w (PWORDLEN n) = (WORDLEN w = n)"
+lemma PWORDLEN: "IN w (PWORDLEN n) = (WORDLEN w = n)"
   by (import word_base PWORDLEN)
 
-lemma PWORDLEN0: "ALL w::'a::type word. IN w (PWORDLEN 0) --> w = WORD []"
+lemma PWORDLEN0: "IN w (PWORDLEN 0) ==> w = WORD []"
   by (import word_base PWORDLEN0)
 
-lemma PWORDLEN1: "ALL x::'a::type. IN (WORD [x]) (PWORDLEN 1)"
+lemma PWORDLEN1: "IN (WORD [x]) (PWORDLEN 1)"
   by (import word_base PWORDLEN1)
 
 consts
   WSEG :: "nat => nat => 'a word => 'a word" 
 
-specification (WSEG) WSEG_DEF: "ALL (m::nat) (k::nat) l::'a::type list.
-   WSEG m k (WORD l) = WORD (LASTN m (BUTLASTN k l))"
+specification (WSEG) WSEG_DEF: "ALL m k l. WSEG m k (WORD l) = WORD (LASTN m (BUTLASTN k l))"
   by (import word_base WSEG_DEF)
 
-lemma WSEG0: "ALL (k::nat) w::'a::type word. WSEG 0 k w = WORD []"
+lemma WSEG0: "WSEG 0 k w = WORD []"
   by (import word_base WSEG0)
 
-lemma WSEG_PWORDLEN: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat. m + k <= n --> IN (WSEG m k w) (PWORDLEN m))"
+lemma WSEG_PWORDLEN: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k. m + k <= n --> IN (WSEG m k w) (PWORDLEN m))"
   by (import word_base WSEG_PWORDLEN)
 
-lemma WSEG_WORDLEN: "ALL x::nat.
-   RES_FORALL (PWORDLEN x)
-    (%xa::'a::type word.
-        ALL (xb::nat) xc::nat.
-           xb + xc <= x --> WORDLEN (WSEG xb xc xa) = xb)"
+lemma WSEG_WORDLEN: "RES_FORALL (PWORDLEN x)
+ (%xa. ALL xb xc. xb + xc <= x --> WORDLEN (WSEG xb xc xa) = xb)"
   by (import word_base WSEG_WORDLEN)
 
-lemma WSEG_WORD_LENGTH: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::'a::type word. WSEG n 0 w = w)"
+lemma WSEG_WORD_LENGTH: "RES_FORALL (PWORDLEN n) (%w. WSEG n 0 w = w)"
   by (import word_base WSEG_WORD_LENGTH)
 
 consts
   bit :: "nat => 'a word => 'a" 
 
-specification (bit) BIT_DEF: "ALL (k::nat) l::'a::type list. bit k (WORD l) = ELL k l"
+specification (bit) BIT_DEF: "ALL k l. bit k (WORD l) = ELL k l"
   by (import word_base BIT_DEF)
 
-lemma BIT0: "ALL x::'a::type. bit 0 (WORD [x]) = x"
+lemma BIT0: "bit 0 (WORD [x]) = x"
   by (import word_base BIT0)
 
-lemma WSEG_BIT: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%k::nat.
-               (op -->::bool => bool => bool)
-                ((op <::nat => nat => bool) k n)
-                ((op =::'a::type word => 'a::type word => bool)
-                  ((WSEG::nat => nat => 'a::type word => 'a::type word)
-                    (1::nat) k w)
-                  ((WORD::'a::type list => 'a::type word)
-                    ((op #::'a::type => 'a::type list => 'a::type list)
-                      ((bit::nat => 'a::type word => 'a::type) k w)
-                      ([]::'a::type list)))))))"
+lemma WSEG_BIT: "RES_FORALL (PWORDLEN n) (%w. ALL k<n. WSEG 1 k w = WORD [bit k w])"
   by (import word_base WSEG_BIT)
 
-lemma BIT_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) (k::nat) j::nat.
-           m + k <= n --> j < m --> bit j (WSEG m k w) = bit (j + k) w)"
+lemma BIT_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k j.
+         m + k <= n --> j < m --> bit j (WSEG m k w) = bit (j + k) w)"
   by (import word_base BIT_WSEG)
 
 consts
   MSB :: "'a word => 'a" 
 
-specification (MSB) MSB_DEF: "ALL l::'a::type list. MSB (WORD l) = hd l"
+specification (MSB) MSB_DEF: "ALL l. MSB (WORD l) = hd l"
   by (import word_base MSB_DEF)
 
-lemma MSB: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word. 0 < n --> MSB w = bit (PRE n) w)"
+lemma MSB: "RES_FORALL (PWORDLEN n) (%w. 0 < n --> MSB w = bit (PRE n) w)"
   by (import word_base MSB)
 
 consts
   LSB :: "'a word => 'a" 
 
-specification (LSB) LSB_DEF: "ALL l::'a::type list. LSB (WORD l) = last l"
+specification (LSB) LSB_DEF: "ALL l. LSB (WORD l) = last l"
   by (import word_base LSB_DEF)
 
-lemma LSB: "ALL n::nat.
-   RES_FORALL (PWORDLEN n) (%w::'a::type word. 0 < n --> LSB w = bit 0 w)"
+lemma LSB: "RES_FORALL (PWORDLEN n) (%w. 0 < n --> LSB w = bit 0 w)"
   by (import word_base LSB)
 
 consts
   WSPLIT :: "nat => 'a word => 'a word * 'a word" 
 
-specification (WSPLIT) WSPLIT_DEF: "ALL (m::nat) l::'a::type list.
-   WSPLIT m (WORD l) = (WORD (BUTLASTN m l), WORD (LASTN m l))"
+specification (WSPLIT) WSPLIT_DEF: "ALL m l. WSPLIT m (WORD l) = (WORD (BUTLASTN m l), WORD (LASTN m l))"
   by (import word_base WSPLIT_DEF)
 
 consts
   WCAT :: "'a word * 'a word => 'a word" 
 
-specification (WCAT) WCAT_DEF: "ALL (l1::'a::type list) l2::'a::type list.
-   WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)"
+specification (WCAT) WCAT_DEF: "ALL l1 l2. WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)"
   by (import word_base WCAT_DEF)
 
-lemma WORD_PARTITION: "(op &::bool => bool => bool)
- ((All::(nat => bool) => bool)
-   (%n::nat.
-       (RES_FORALL::('a::type word => bool)
-                    => ('a::type word => bool) => bool)
-        ((PWORDLEN::nat => 'a::type word => bool) n)
-        (%w::'a::type word.
-            (All::(nat => bool) => bool)
-             (%m::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <=::nat => nat => bool) m n)
-                  ((op =::'a::type word => 'a::type word => bool)
-                    ((WCAT::'a::type word * 'a::type word => 'a::type word)
-                      ((WSPLIT::nat
-                                => 'a::type word
-                                   => 'a::type word * 'a::type word)
-                        m w))
-                    w)))))
- ((All::(nat => bool) => bool)
-   (%n::nat.
-       (All::(nat => bool) => bool)
-        (%m::nat.
-            (RES_FORALL::('a::type word => bool)
-                         => ('a::type word => bool) => bool)
-             ((PWORDLEN::nat => 'a::type word => bool) n)
-             (%w1::'a::type word.
-                 (RES_FORALL::('a::type word => bool)
-                              => ('a::type word => bool) => bool)
-                  ((PWORDLEN::nat => 'a::type word => bool) m)
-                  (%w2::'a::type word.
-                      (op =::'a::type word * 'a::type word
-                             => 'a::type word * 'a::type word => bool)
-                       ((WSPLIT::nat
-                                 => 'a::type word
-                                    => 'a::type word * 'a::type word)
-                         m ((WCAT::'a::type word * 'a::type word
-                                   => 'a::type word)
-                             ((Pair::'a::type word
-                                     => 'a::type word
-  => 'a::type word * 'a::type word)
-                               w1 w2)))
-                       ((Pair::'a::type word
-                               => 'a::type word
-                                  => 'a::type word * 'a::type word)
-                         w1 w2))))))"
+lemma WORD_PARTITION: "(ALL n::nat.
+    RES_FORALL (PWORDLEN n)
+     (%w::'a word. ALL m<=n. WCAT (WSPLIT m w) = w)) &
+(ALL (n::nat) m::nat.
+    RES_FORALL (PWORDLEN n)
+     (%w1::'a word.
+         RES_FORALL (PWORDLEN m)
+          (%w2::'a word. WSPLIT m (WCAT (w1, w2)) = (w1, w2))))"
   by (import word_base WORD_PARTITION)
 
-lemma WCAT_ASSOC: "ALL (w1::'a::type word) (w2::'a::type word) w3::'a::type word.
-   WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)"
+lemma WCAT_ASSOC: "WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)"
   by (import word_base WCAT_ASSOC)
 
-lemma WCAT0: "ALL w::'a::type word. WCAT (WORD [], w) = w & WCAT (w, WORD []) = w"
+lemma WCAT0: "WCAT (WORD [], w) = w & WCAT (w, WORD []) = w"
   by (import word_base WCAT0)
 
-lemma WCAT_11: "ALL (m::nat) n::nat.
-   RES_FORALL (PWORDLEN m)
-    (%wm1::'a::type word.
-        RES_FORALL (PWORDLEN m)
-         (%wm2::'a::type word.
-             RES_FORALL (PWORDLEN n)
-              (%wn1::'a::type word.
-                  RES_FORALL (PWORDLEN n)
-                   (%wn2::'a::type word.
-                       (WCAT (wm1, wn1) = WCAT (wm2, wn2)) =
-                       (wm1 = wm2 & wn1 = wn2)))))"
+lemma WCAT_11: "RES_FORALL (PWORDLEN m)
+ (%wm1. RES_FORALL (PWORDLEN m)
+         (%wm2. RES_FORALL (PWORDLEN n)
+                 (%wn1. RES_FORALL (PWORDLEN n)
+                         (%wn2. (WCAT (wm1, wn1) = WCAT (wm2, wn2)) =
+                                (wm1 = wm2 & wn1 = wn2)))))"
   by (import word_base WCAT_11)
 
-lemma WSPLIT_PWORDLEN: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%m::nat.
-               (op -->::bool => bool => bool)
-                ((op <=::nat => nat => bool) m n)
-                ((op &::bool => bool => bool)
-                  ((IN::'a::type word => ('a::type word => bool) => bool)
-                    ((fst::'a::type word * 'a::type word => 'a::type word)
-                      ((WSPLIT::nat
-                                => 'a::type word
-                                   => 'a::type word * 'a::type word)
-                        m w))
-                    ((PWORDLEN::nat => 'a::type word => bool)
-                      ((op -::nat => nat => nat) n m)))
-                  ((IN::'a::type word => ('a::type word => bool) => bool)
-                    ((snd::'a::type word * 'a::type word => 'a::type word)
-                      ((WSPLIT::nat
-                                => 'a::type word
-                                   => 'a::type word * 'a::type word)
-                        m w))
-                    ((PWORDLEN::nat => 'a::type word => bool) m))))))"
+lemma WSPLIT_PWORDLEN: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m<=n.
+         IN (fst (WSPLIT m w)) (PWORDLEN (n - m)) &
+         IN (snd (WSPLIT m w)) (PWORDLEN m))"
   by (import word_base WSPLIT_PWORDLEN)
 
-lemma WCAT_PWORDLEN: "ALL n1::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        ALL n2::nat.
-           RES_FORALL (PWORDLEN n2)
-            (%w2::'a::type word. IN (WCAT (w1, w2)) (PWORDLEN (n1 + n2))))"
+lemma WCAT_PWORDLEN: "RES_FORALL (PWORDLEN n1)
+ (%w1. ALL n2.
+          RES_FORALL (PWORDLEN n2)
+           (%w2. IN (WCAT (w1, w2)) (PWORDLEN (n1 + n2))))"
   by (import word_base WCAT_PWORDLEN)
 
-lemma WORDLEN_SUC_WCAT: "ALL (n::nat) w::'a::type word.
-   IN w (PWORDLEN (Suc n)) -->
-   RES_EXISTS (PWORDLEN 1)
-    (%b::'a::type word.
-        RES_EXISTS (PWORDLEN n) (%w'::'a::type word. w = WCAT (b, w')))"
+lemma WORDLEN_SUC_WCAT: "IN w (PWORDLEN (Suc n))
+==> RES_EXISTS (PWORDLEN 1)
+     (%b. RES_EXISTS (PWORDLEN n) (%w'. w = WCAT (b, w')))"
   by (import word_base WORDLEN_SUC_WCAT)
 
-lemma WSEG_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m1::nat) (k1::nat) (m2::nat) k2::nat.
-           m1 + k1 <= n & m2 + k2 <= m1 -->
-           WSEG m2 k2 (WSEG m1 k1 w) = WSEG m2 (k1 + k2) w)"
+lemma WSEG_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m1 k1 m2 k2.
+         m1 + k1 <= n & m2 + k2 <= m1 -->
+         WSEG m2 k2 (WSEG m1 k1 w) = WSEG m2 (k1 + k2) w)"
   by (import word_base WSEG_WSEG)
 
-lemma WSPLIT_WSEG: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%k::nat.
-               (op -->::bool => bool => bool)
-                ((op <=::nat => nat => bool) k n)
-                ((op =::'a::type word * 'a::type word
-                        => 'a::type word * 'a::type word => bool)
-                  ((WSPLIT::nat
-                            => 'a::type word
-                               => 'a::type word * 'a::type word)
-                    k w)
-                  ((Pair::'a::type word
-                          => 'a::type word => 'a::type word * 'a::type word)
-                    ((WSEG::nat => nat => 'a::type word => 'a::type word)
-                      ((op -::nat => nat => nat) n k) k w)
-                    ((WSEG::nat => nat => 'a::type word => 'a::type word) k
-                      (0::nat) w))))))"
+lemma WSPLIT_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL k<=n. WSPLIT k w = (WSEG (n - k) k w, WSEG k 0 w))"
   by (import word_base WSPLIT_WSEG)
 
-lemma WSPLIT_WSEG1: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%k::nat.
-               (op -->::bool => bool => bool)
-                ((op <=::nat => nat => bool) k n)
-                ((op =::'a::type word => 'a::type word => bool)
-                  ((fst::'a::type word * 'a::type word => 'a::type word)
-                    ((WSPLIT::nat
-                              => 'a::type word
-                                 => 'a::type word * 'a::type word)
-                      k w))
-                  ((WSEG::nat => nat => 'a::type word => 'a::type word)
-                    ((op -::nat => nat => nat) n k) k w)))))"
+lemma WSPLIT_WSEG1: "RES_FORALL (PWORDLEN n) (%w. ALL k<=n. fst (WSPLIT k w) = WSEG (n - k) k w)"
   by (import word_base WSPLIT_WSEG1)
 
-lemma WSPLIT_WSEG2: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%k::nat.
-               (op -->::bool => bool => bool)
-                ((op <=::nat => nat => bool) k n)
-                ((op =::'a::type word => 'a::type word => bool)
-                  ((snd::'a::type word * 'a::type word => 'a::type word)
-                    ((WSPLIT::nat
-                              => 'a::type word
-                                 => 'a::type word * 'a::type word)
-                      k w))
-                  ((WSEG::nat => nat => 'a::type word => 'a::type word) k
-                    (0::nat) w)))))"
+lemma WSPLIT_WSEG2: "RES_FORALL (PWORDLEN n) (%w. ALL k<=n. snd (WSPLIT k w) = WSEG k 0 w)"
   by (import word_base WSPLIT_WSEG2)
 
-lemma WCAT_WSEG_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m1::nat) (m2::nat) k::nat.
-           m1 + (m2 + k) <= n -->
-           WCAT (WSEG m2 (m1 + k) w, WSEG m1 k w) = WSEG (m1 + m2) k w)"
+lemma WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m1 m2 k.
+         m1 + (m2 + k) <= n -->
+         WCAT (WSEG m2 (m1 + k) w, WSEG m1 k w) = WSEG (m1 + m2) k w)"
   by (import word_base WCAT_WSEG_WSEG)
 
-lemma WORD_SPLIT: "ALL (x::nat) xa::nat.
-   RES_FORALL (PWORDLEN (x + xa))
-    (%w::'a::type word. w = WCAT (WSEG x xa w, WSEG xa 0 w))"
+lemma WORD_SPLIT: "RES_FORALL (PWORDLEN (x + xa)) (%w. w = WCAT (WSEG x xa w, WSEG xa 0 w))"
   by (import word_base WORD_SPLIT)
 
-lemma WORDLEN_SUC_WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN (Suc (n::nat)))
- (%w::'a::type word. w = WCAT (WSEG 1 n w, WSEG n 0 w))"
+lemma WORDLEN_SUC_WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG 1 n w, WSEG n 0 w))"
   by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG)
 
-lemma WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc (n::nat)))
- (%w::'a::type word. w = WCAT (WSEG n 1 w, WSEG 1 0 w))"
+lemma WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG n 1 w, WSEG 1 0 w))"
   by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT)
 
-lemma WORDLEN_SUC_WCAT_BIT_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w::'a::type word. w = WCAT (WORD [bit n w], WSEG n 0 w))"
+lemma WORDLEN_SUC_WCAT_BIT_WSEG: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WORD [bit n w], WSEG n 0 w))"
   by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG)
 
-lemma WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w::'a::type word. w = WCAT (WSEG n 1 w, WORD [bit 0 w]))"
+lemma WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG n 1 w, WORD [bit 0 w]))"
   by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT)
 
-lemma WSEG_WCAT1: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word. WSEG n1 n2 (WCAT (w1, w2)) = w1))"
+lemma WSEG_WCAT1: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2) (%w2. WSEG n1 n2 (WCAT (w1, w2)) = w1))"
   by (import word_base WSEG_WCAT1)
 
-lemma WSEG_WCAT2: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word. WSEG n2 0 (WCAT (w1, w2)) = w2))"
+lemma WSEG_WCAT2: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2) (%w2. WSEG n2 0 (WCAT (w1, w2)) = w2))"
   by (import word_base WSEG_WCAT2)
 
-lemma WSEG_SUC: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (k::nat) m1::nat.
-           k + Suc m1 < n -->
-           WSEG (Suc m1) k w = WCAT (WSEG 1 (k + m1) w, WSEG m1 k w))"
+lemma WSEG_SUC: "RES_FORALL (PWORDLEN n)
+ (%w. ALL k m1.
+         k + Suc m1 < n -->
+         WSEG (Suc m1) k w = WCAT (WSEG 1 (k + m1) w, WSEG m1 k w))"
   by (import word_base WSEG_SUC)
 
-lemma WORD_CONS_WCAT: "ALL (x::'a::type) l::'a::type list. WORD (x # l) = WCAT (WORD [x], WORD l)"
+lemma WORD_CONS_WCAT: "WORD (x # l) = WCAT (WORD [x], WORD l)"
   by (import word_base WORD_CONS_WCAT)
 
-lemma WORD_SNOC_WCAT: "ALL (l::'a::type list) x::'a::type.
-   WORD (SNOC x l) = WCAT (WORD l, WORD [x])"
+lemma WORD_SNOC_WCAT: "WORD (SNOC x l) = WCAT (WORD l, WORD [x])"
   by (import word_base WORD_SNOC_WCAT)
 
-lemma BIT_WCAT_FST: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word.
-             ALL k::nat.
-                n2 <= k & k < n1 + n2 -->
-                bit k (WCAT (w1, w2)) = bit (k - n2) w1))"
+lemma BIT_WCAT_FST: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. ALL k.
+                 n2 <= k & k < n1 + n2 -->
+                 bit k (WCAT (w1, w2)) = bit (k - n2) w1))"
   by (import word_base BIT_WCAT_FST)
 
-lemma BIT_WCAT_SND: "(All::(nat => bool) => bool)
- (%n1::nat.
-     (All::(nat => bool) => bool)
-      (%n2::nat.
-          (RES_FORALL::('a::type word => bool)
-                       => ('a::type word => bool) => bool)
-           ((PWORDLEN::nat => 'a::type word => bool) n1)
-           (%w1::'a::type word.
-               (RES_FORALL::('a::type word => bool)
-                            => ('a::type word => bool) => bool)
-                ((PWORDLEN::nat => 'a::type word => bool) n2)
-                (%w2::'a::type word.
-                    (All::(nat => bool) => bool)
-                     (%k::nat.
-                         (op -->::bool => bool => bool)
-                          ((op <::nat => nat => bool) k n2)
-                          ((op =::'a::type => 'a::type => bool)
-                            ((bit::nat => 'a::type word => 'a::type) k
-                              ((WCAT::'a::type word * 'a::type word
-=> 'a::type word)
-                                ((Pair::'a::type word
-  => 'a::type word => 'a::type word * 'a::type word)
-                                  w1 w2)))
-                            ((bit::nat => 'a::type word => 'a::type) k
-                              w2)))))))"
+lemma BIT_WCAT_SND: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. ALL k<n2. bit k (WCAT (w1, w2)) = bit k w2))"
   by (import word_base BIT_WCAT_SND)
 
-lemma BIT_WCAT1: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word. ALL b::'a::type. bit n (WCAT (WORD [b], w)) = b)"
+lemma BIT_WCAT1: "RES_FORALL (PWORDLEN n) (%w. ALL b. bit n (WCAT (WORD [b], w)) = b)"
   by (import word_base BIT_WCAT1)
 
-lemma WSEG_WCAT_WSEG1: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word.
-             ALL (m::nat) k::nat.
-                m <= n1 & n2 <= k -->
-                WSEG m k (WCAT (w1, w2)) = WSEG m (k - n2) w1))"
+lemma WSEG_WCAT_WSEG1: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. ALL m k.
+                 m <= n1 & n2 <= k -->
+                 WSEG m k (WCAT (w1, w2)) = WSEG m (k - n2) w1))"
   by (import word_base WSEG_WCAT_WSEG1)
 
-lemma WSEG_WCAT_WSEG2: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word.
-             ALL (m::nat) k::nat.
-                m + k <= n2 --> WSEG m k (WCAT (w1, w2)) = WSEG m k w2))"
+lemma WSEG_WCAT_WSEG2: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. ALL m k.
+                 m + k <= n2 --> WSEG m k (WCAT (w1, w2)) = WSEG m k w2))"
   by (import word_base WSEG_WCAT_WSEG2)
 
-lemma WSEG_WCAT_WSEG: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::'a::type word.
-             ALL (m::nat) k::nat.
-                m + k <= n1 + n2 & k < n2 & n2 <= m + k -->
-                WSEG m k (WCAT (w1, w2)) =
-                WCAT (WSEG (m + k - n2) 0 w1, WSEG (n2 - k) k w2)))"
+lemma WSEG_WCAT_WSEG: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. ALL m k.
+                 m + k <= n1 + n2 & k < n2 & n2 <= m + k -->
+                 WSEG m k (WCAT (w1, w2)) =
+                 WCAT (WSEG (m + k - n2) 0 w1, WSEG (n2 - k) k w2)))"
   by (import word_base WSEG_WCAT_WSEG)
 
-lemma BIT_EQ_IMP_WORD_EQ: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w1::'a::type word.
-          (RES_FORALL::('a::type word => bool)
-                       => ('a::type word => bool) => bool)
-           ((PWORDLEN::nat => 'a::type word => bool) n)
-           (%w2::'a::type word.
-               (op -->::bool => bool => bool)
-                ((All::(nat => bool) => bool)
-                  (%k::nat.
-                      (op -->::bool => bool => bool)
-                       ((op <::nat => nat => bool) k n)
-                       ((op =::'a::type => 'a::type => bool)
-                         ((bit::nat => 'a::type word => 'a::type) k w1)
-                         ((bit::nat => 'a::type word => 'a::type) k w2))))
-                ((op =::'a::type word => 'a::type word => bool) w1 w2))))"
+lemma BIT_EQ_IMP_WORD_EQ: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. (ALL k<n. bit k w1 = bit k w2) --> w1 = w2))"
   by (import word_base BIT_EQ_IMP_WORD_EQ)
 
 ;end_setup
 
 ;setup_theory word_num
 
-definition LVAL :: "('a => nat) => nat => 'a list => nat" where 
-  "LVAL ==
-%(f::'a::type => nat) b::nat. foldl (%(e::nat) x::'a::type. b * e + f x) 0"
+definition
+  LVAL :: "('a => nat) => nat => 'a list => nat"  where
+  "LVAL == %f b. foldl (%e x. b * e + f x) 0"
 
-lemma LVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list.
-   LVAL f b l = foldl (%(e::nat) x::'a::type. b * e + f x) 0 l"
+lemma LVAL_DEF: "LVAL f b l = foldl (%e x. b * e + f x) 0 l"
   by (import word_num LVAL_DEF)
 
 consts
   NVAL :: "('a => nat) => nat => 'a word => nat" 
 
-specification (NVAL) NVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list.
-   NVAL f b (WORD l) = LVAL f b l"
+specification (NVAL) NVAL_DEF: "ALL f b l. NVAL f b (WORD l) = LVAL f b l"
   by (import word_num NVAL_DEF)
 
-lemma LVAL: "(ALL (x::'a::type => nat) xa::nat. LVAL x xa [] = 0) &
-(ALL (x::'a::type list) (xa::'a::type => nat) (xb::nat) xc::'a::type.
+lemma LVAL: "(ALL (x::'a => nat) xa::nat. LVAL x xa [] = (0::nat)) &
+(ALL (x::'a list) (xa::'a => nat) (xb::nat) xc::'a.
     LVAL xa xb (xc # x) = xa xc * xb ^ length x + LVAL xa xb x)"
   by (import word_num LVAL)
 
-lemma LVAL_SNOC: "ALL (l::'a::type list) (h::'a::type) (f::'a::type => nat) b::nat.
-   LVAL f b (SNOC h l) = LVAL f b l * b + f h"
+lemma LVAL_SNOC: "LVAL f b (SNOC h l) = LVAL f b l * b + f h"
   by (import word_num LVAL_SNOC)
 
-lemma LVAL_MAX: "ALL (l::'a::type list) (f::'a::type => nat) b::nat.
-   (ALL x::'a::type. f x < b) --> LVAL f b l < b ^ length l"
+lemma LVAL_MAX: "(!!x. f x < b) ==> LVAL f b l < b ^ length l"
   by (import word_num LVAL_MAX)
 
-lemma NVAL_MAX: "ALL (f::'a::type => nat) b::nat.
-   (ALL x::'a::type. f x < b) -->
-   (ALL n::nat.
-       RES_FORALL (PWORDLEN n) (%w::'a::type word. NVAL f b w < b ^ n))"
+lemma NVAL_MAX: "(!!x. f x < b) ==> RES_FORALL (PWORDLEN n) (%w. NVAL f b w < b ^ n)"
   by (import word_num NVAL_MAX)
 
-lemma NVAL0: "ALL (x::'a::type => nat) xa::nat. NVAL x xa (WORD []) = 0"
+lemma NVAL0: "NVAL x xa (WORD []) = 0"
   by (import word_num NVAL0)
 
-lemma NVAL1: "ALL (x::'a::type => nat) (xa::nat) xb::'a::type.
-   NVAL x xa (WORD [xb]) = x xb"
+lemma NVAL1: "NVAL x xa (WORD [xb]) = x xb"
   by (import word_num NVAL1)
 
-lemma NVAL_WORDLEN_0: "RES_FORALL (PWORDLEN 0)
- (%w::'a::type word. ALL (fv::'a::type => nat) r::nat. NVAL fv r w = 0)"
+lemma NVAL_WORDLEN_0: "RES_FORALL (PWORDLEN 0) (%w. ALL fv r. NVAL fv r w = 0)"
   by (import word_num NVAL_WORDLEN_0)
 
-lemma NVAL_WCAT1: "ALL (w::'a::type word) (f::'a::type => nat) (b::nat) x::'a::type.
-   NVAL f b (WCAT (w, WORD [x])) = NVAL f b w * b + f x"
+lemma NVAL_WCAT1: "NVAL f b (WCAT (w, WORD [x])) = NVAL f b w * b + f x"
   by (import word_num NVAL_WCAT1)
 
-lemma NVAL_WCAT2: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (f::'a::type => nat) (b::nat) x::'a::type.
-           NVAL f b (WCAT (WORD [x], w)) = f x * b ^ n + NVAL f b w)"
+lemma NVAL_WCAT2: "RES_FORALL (PWORDLEN n)
+ (%w. ALL f b x. NVAL f b (WCAT (WORD [x], w)) = f x * b ^ n + NVAL f b w)"
   by (import word_num NVAL_WCAT2)
 
-lemma NVAL_WCAT: "ALL (n::nat) m::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::'a::type word.
-        RES_FORALL (PWORDLEN m)
-         (%w2::'a::type word.
-             ALL (f::'a::type => nat) b::nat.
-                NVAL f b (WCAT (w1, w2)) =
-                NVAL f b w1 * b ^ m + NVAL f b w2))"
+lemma NVAL_WCAT: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN m)
+        (%w2. ALL f b.
+                 NVAL f b (WCAT (w1, w2)) =
+                 NVAL f b w1 * b ^ m + NVAL f b w2))"
   by (import word_num NVAL_WCAT)
 
 consts
   NLIST :: "nat => (nat => 'a) => nat => nat => 'a list" 
 
-specification (NLIST) NLIST_DEF: "(ALL (frep::nat => 'a::type) (b::nat) m::nat. NLIST 0 frep b m = []) &
-(ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat.
+specification (NLIST) NLIST_DEF: "(ALL (frep::nat => 'a) (b::nat) m::nat. NLIST (0::nat) frep b m = []) &
+(ALL (n::nat) (frep::nat => 'a) (b::nat) m::nat.
     NLIST (Suc n) frep b m =
     SNOC (frep (m mod b)) (NLIST n frep b (m div b)))"
   by (import word_num NLIST_DEF)
 
-definition NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" where 
-  "NWORD ==
-%(n::nat) (frep::nat => 'a::type) (b::nat) m::nat. WORD (NLIST n frep b m)"
+definition
+  NWORD :: "nat => (nat => 'a) => nat => nat => 'a word"  where
+  "NWORD == %n frep b m. WORD (NLIST n frep b m)"
 
-lemma NWORD_DEF: "ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat.
-   NWORD n frep b m = WORD (NLIST n frep b m)"
+lemma NWORD_DEF: "NWORD n frep b m = WORD (NLIST n frep b m)"
   by (import word_num NWORD_DEF)
 
-lemma NWORD_LENGTH: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat.
-   WORDLEN (NWORD x xa xb xc) = x"
+lemma NWORD_LENGTH: "WORDLEN (NWORD x xa xb xc) = x"
   by (import word_num NWORD_LENGTH)
 
-lemma NWORD_PWORDLEN: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat.
-   IN (NWORD x xa xb xc) (PWORDLEN x)"
+lemma NWORD_PWORDLEN: "IN (NWORD x xa xb xc) (PWORDLEN x)"
   by (import word_num NWORD_PWORDLEN)
 
 ;end_setup
@@ -780,79 +487,49 @@
 defs
   PBITOP_primdef: "PBITOP ==
 GSPEC
- (%oper::'a::type word => 'b::type word.
+ (%oper.
      (oper,
-      ALL n::nat.
+      ALL n.
          RES_FORALL (PWORDLEN n)
-          (%w::'a::type word.
-              IN (oper w) (PWORDLEN n) &
-              (ALL (m::nat) k::nat.
-                  m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))"
+          (%w. IN (oper w) (PWORDLEN n) &
+               (ALL m k.
+                   m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))"
 
 lemma PBITOP_def: "PBITOP =
 GSPEC
- (%oper::'a::type word => 'b::type word.
+ (%oper.
      (oper,
-      ALL n::nat.
+      ALL n.
          RES_FORALL (PWORDLEN n)
-          (%w::'a::type word.
-              IN (oper w) (PWORDLEN n) &
-              (ALL (m::nat) k::nat.
-                  m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))"
+          (%w. IN (oper w) (PWORDLEN n) &
+               (ALL m k.
+                   m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))"
   by (import word_bitop PBITOP_def)
 
-lemma IN_PBITOP: "ALL oper::'a::type word => 'b::type word.
-   IN oper PBITOP =
-   (ALL n::nat.
-       RES_FORALL (PWORDLEN n)
-        (%w::'a::type word.
-            IN (oper w) (PWORDLEN n) &
-            (ALL (m::nat) k::nat.
-                m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w))))"
+lemma IN_PBITOP: "IN oper PBITOP =
+(ALL n.
+    RES_FORALL (PWORDLEN n)
+     (%w. IN (oper w) (PWORDLEN n) &
+          (ALL m k. m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w))))"
   by (import word_bitop IN_PBITOP)
 
 lemma PBITOP_PWORDLEN: "RES_FORALL PBITOP
- (%oper::'a::type word => 'b::type word.
-     ALL n::nat.
-        RES_FORALL (PWORDLEN n)
-         (%w::'a::type word. IN (oper w) (PWORDLEN n)))"
+ (%oper. ALL n. RES_FORALL (PWORDLEN n) (%w. IN (oper w) (PWORDLEN n)))"
   by (import word_bitop PBITOP_PWORDLEN)
 
 lemma PBITOP_WSEG: "RES_FORALL PBITOP
- (%oper::'a::type word => 'b::type word.
-     ALL n::nat.
+ (%oper.
+     ALL n.
         RES_FORALL (PWORDLEN n)
-         (%w::'a::type word.
-             ALL (m::nat) k::nat.
-                m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))"
+         (%w. ALL m k.
+                 m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))"
   by (import word_bitop PBITOP_WSEG)
 
-lemma PBITOP_BIT: "(RES_FORALL::(('a::type word => 'b::type word) => bool)
-             => (('a::type word => 'b::type word) => bool) => bool)
- (PBITOP::('a::type word => 'b::type word) => bool)
- (%oper::'a::type word => 'b::type word.
-     (All::(nat => bool) => bool)
-      (%n::nat.
-          (RES_FORALL::('a::type word => bool)
-                       => ('a::type word => bool) => bool)
-           ((PWORDLEN::nat => 'a::type word => bool) n)
-           (%w::'a::type word.
-               (All::(nat => bool) => bool)
-                (%k::nat.
-                    (op -->::bool => bool => bool)
-                     ((op <::nat => nat => bool) k n)
-                     ((op =::'b::type word => 'b::type word => bool)
-                       (oper
-                         ((WORD::'a::type list => 'a::type word)
-                           ((op #::'a::type
-                                   => 'a::type list => 'a::type list)
-                             ((bit::nat => 'a::type word => 'a::type) k w)
-                             ([]::'a::type list))))
-                       ((WORD::'b::type list => 'b::type word)
-                         ((op #::'b::type => 'b::type list => 'b::type list)
-                           ((bit::nat => 'b::type word => 'b::type) k
-                             (oper w))
-                           ([]::'b::type list))))))))"
+lemma PBITOP_BIT: "RES_FORALL PBITOP
+ (%oper.
+     ALL n.
+        RES_FORALL (PWORDLEN n)
+         (%w. ALL k<n. oper (WORD [bit k w]) = WORD [bit k (oper w)]))"
   by (import word_bitop PBITOP_BIT)
 
 consts
@@ -861,532 +538,383 @@
 defs
   PBITBOP_primdef: "PBITBOP ==
 GSPEC
- (%oper::'a::type word => 'b::type word => 'c::type word.
+ (%oper.
      (oper,
-      ALL n::nat.
+      ALL n.
          RES_FORALL (PWORDLEN n)
-          (%w1::'a::type word.
-              RES_FORALL (PWORDLEN n)
-               (%w2::'b::type word.
-                   IN (oper w1 w2) (PWORDLEN n) &
-                   (ALL (m::nat) k::nat.
-                       m + k <= n -->
-                       oper (WSEG m k w1) (WSEG m k w2) =
-                       WSEG m k (oper w1 w2))))))"
+          (%w1. RES_FORALL (PWORDLEN n)
+                 (%w2. IN (oper w1 w2) (PWORDLEN n) &
+                       (ALL m k.
+                           m + k <= n -->
+                           oper (WSEG m k w1) (WSEG m k w2) =
+                           WSEG m k (oper w1 w2))))))"
 
 lemma PBITBOP_def: "PBITBOP =
 GSPEC
- (%oper::'a::type word => 'b::type word => 'c::type word.
+ (%oper.
      (oper,
-      ALL n::nat.
+      ALL n.
          RES_FORALL (PWORDLEN n)
-          (%w1::'a::type word.
-              RES_FORALL (PWORDLEN n)
-               (%w2::'b::type word.
-                   IN (oper w1 w2) (PWORDLEN n) &
-                   (ALL (m::nat) k::nat.
-                       m + k <= n -->
-                       oper (WSEG m k w1) (WSEG m k w2) =
-                       WSEG m k (oper w1 w2))))))"
+          (%w1. RES_FORALL (PWORDLEN n)
+                 (%w2. IN (oper w1 w2) (PWORDLEN n) &
+                       (ALL m k.
+                           m + k <= n -->
+                           oper (WSEG m k w1) (WSEG m k w2) =
+                           WSEG m k (oper w1 w2))))))"
   by (import word_bitop PBITBOP_def)
 
-lemma IN_PBITBOP: "ALL oper::'a::type word => 'b::type word => 'c::type word.
-   IN oper PBITBOP =
-   (ALL n::nat.
-       RES_FORALL (PWORDLEN n)
-        (%w1::'a::type word.
-            RES_FORALL (PWORDLEN n)
-             (%w2::'b::type word.
-                 IN (oper w1 w2) (PWORDLEN n) &
-                 (ALL (m::nat) k::nat.
-                     m + k <= n -->
-                     oper (WSEG m k w1) (WSEG m k w2) =
-                     WSEG m k (oper w1 w2)))))"
+lemma IN_PBITBOP: "IN oper PBITBOP =
+(ALL n.
+    RES_FORALL (PWORDLEN n)
+     (%w1. RES_FORALL (PWORDLEN n)
+            (%w2. IN (oper w1 w2) (PWORDLEN n) &
+                  (ALL m k.
+                      m + k <= n -->
+                      oper (WSEG m k w1) (WSEG m k w2) =
+                      WSEG m k (oper w1 w2)))))"
   by (import word_bitop IN_PBITBOP)
 
 lemma PBITBOP_PWORDLEN: "RES_FORALL PBITBOP
- (%oper::'a::type word => 'b::type word => 'c::type word.
-     ALL n::nat.
+ (%oper.
+     ALL n.
         RES_FORALL (PWORDLEN n)
-         (%w1::'a::type word.
-             RES_FORALL (PWORDLEN n)
-              (%w2::'b::type word. IN (oper w1 w2) (PWORDLEN n))))"
+         (%w1. RES_FORALL (PWORDLEN n) (%w2. IN (oper w1 w2) (PWORDLEN n))))"
   by (import word_bitop PBITBOP_PWORDLEN)
 
 lemma PBITBOP_WSEG: "RES_FORALL PBITBOP
- (%oper::'a::type word => 'b::type word => 'c::type word.
-     ALL n::nat.
+ (%oper.
+     ALL n.
         RES_FORALL (PWORDLEN n)
-         (%w1::'a::type word.
-             RES_FORALL (PWORDLEN n)
-              (%w2::'b::type word.
-                  ALL (m::nat) k::nat.
-                     m + k <= n -->
-                     oper (WSEG m k w1) (WSEG m k w2) =
-                     WSEG m k (oper w1 w2))))"
+         (%w1. RES_FORALL (PWORDLEN n)
+                (%w2. ALL m k.
+                         m + k <= n -->
+                         oper (WSEG m k w1) (WSEG m k w2) =
+                         WSEG m k (oper w1 w2))))"
   by (import word_bitop PBITBOP_WSEG)
 
-lemma PBITBOP_EXISTS: "ALL f::'a::type => 'b::type => 'c::type.
-   EX x::'a::type word => 'b::type word => 'c::type word.
-      ALL (l1::'a::type list) l2::'b::type list.
-         x (WORD l1) (WORD l2) = WORD (map2 f l1 l2)"
+lemma PBITBOP_EXISTS: "EX x. ALL l1 l2. x (WORD l1) (WORD l2) = WORD (map2 f l1 l2)"
   by (import word_bitop PBITBOP_EXISTS)
 
 consts
   WMAP :: "('a => 'b) => 'a word => 'b word" 
 
-specification (WMAP) WMAP_DEF: "ALL (f::'a::type => 'b::type) l::'a::type list.
-   WMAP f (WORD l) = WORD (map f l)"
+specification (WMAP) WMAP_DEF: "ALL f l. WMAP f (WORD l) = WORD (map f l)"
   by (import word_bitop WMAP_DEF)
 
-lemma WMAP_PWORDLEN: "RES_FORALL (PWORDLEN (n::nat))
- (%w::'a::type word.
-     ALL f::'a::type => 'b::type. IN (WMAP f w) (PWORDLEN n))"
+lemma WMAP_PWORDLEN: "RES_FORALL (PWORDLEN n) (%w. ALL f. IN (WMAP f w) (PWORDLEN n))"
   by (import word_bitop WMAP_PWORDLEN)
 
-lemma WMAP_0: "ALL x::'a::type => 'b::type. WMAP x (WORD []) = WORD []"
+lemma WMAP_0: "WMAP (x::'a => 'b) (WORD []) = WORD []"
   by (import word_bitop WMAP_0)
 
-lemma WMAP_BIT: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(nat => bool) => bool)
-           (%k::nat.
-               (op -->::bool => bool => bool)
-                ((op <::nat => nat => bool) k n)
-                ((All::(('a::type => 'b::type) => bool) => bool)
-                  (%f::'a::type => 'b::type.
-                      (op =::'b::type => 'b::type => bool)
-                       ((bit::nat => 'b::type word => 'b::type) k
-                         ((WMAP::('a::type => 'b::type)
-                                 => 'a::type word => 'b::type word)
-                           f w))
-                       (f ((bit::nat => 'a::type word => 'a::type) k
-                            w)))))))"
+lemma WMAP_BIT: "RES_FORALL (PWORDLEN n) (%w. ALL k<n. ALL f. bit k (WMAP f w) = f (bit k w))"
   by (import word_bitop WMAP_BIT)
 
-lemma WMAP_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k <= n -->
-           (ALL f::'a::type => 'b::type.
-               WMAP f (WSEG m k w) = WSEG m k (WMAP f w)))"
+lemma WMAP_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k <= n --> (ALL f. WMAP f (WSEG m k w) = WSEG m k (WMAP f w)))"
   by (import word_bitop WMAP_WSEG)
 
-lemma WMAP_PBITOP: "ALL f::'a::type => 'b::type. IN (WMAP f) PBITOP"
+lemma WMAP_PBITOP: "IN (WMAP f) PBITOP"
   by (import word_bitop WMAP_PBITOP)
 
-lemma WMAP_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) f::'a::type => 'b::type.
-   WMAP f (WCAT (w1, w2)) = WCAT (WMAP f w1, WMAP f w2)"
+lemma WMAP_WCAT: "WMAP (f::'a => 'b) (WCAT (w1::'a word, w2::'a word)) =
+WCAT (WMAP f w1, WMAP f w2)"
   by (import word_bitop WMAP_WCAT)
 
-lemma WMAP_o: "ALL (w::'a::type word) (f::'a::type => 'b::type) g::'b::type => 'c::type.
-   WMAP g (WMAP f w) = WMAP (g o f) w"
+lemma WMAP_o: "WMAP (g::'b => 'c) (WMAP (f::'a => 'b) (w::'a word)) = WMAP (g o f) w"
   by (import word_bitop WMAP_o)
 
 consts
   FORALLBITS :: "('a => bool) => 'a word => bool" 
 
-specification (FORALLBITS) FORALLBITS_DEF: "ALL (P::'a::type => bool) l::'a::type list.
-   FORALLBITS P (WORD l) = list_all P l"
+specification (FORALLBITS) FORALLBITS_DEF: "ALL P l. FORALLBITS P (WORD l) = list_all P l"
   by (import word_bitop FORALLBITS_DEF)
 
-lemma FORALLBITS: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(('a::type => bool) => bool) => bool)
-           (%P::'a::type => bool.
-               (op =::bool => bool => bool)
-                ((FORALLBITS::('a::type => bool) => 'a::type word => bool) P
-                  w)
-                ((All::(nat => bool) => bool)
-                  (%k::nat.
-                      (op -->::bool => bool => bool)
-                       ((op <::nat => nat => bool) k n)
-                       (P ((bit::nat => 'a::type word => 'a::type) k
-                            w)))))))"
+lemma FORALLBITS: "RES_FORALL (PWORDLEN n) (%w. ALL P. FORALLBITS P w = (ALL k<n. P (bit k w)))"
   by (import word_bitop FORALLBITS)
 
-lemma FORALLBITS_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL P::'a::type => bool.
-           FORALLBITS P w -->
-           (ALL (m::nat) k::nat. m + k <= n --> FORALLBITS P (WSEG m k w)))"
+lemma FORALLBITS_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL P.
+         FORALLBITS P w -->
+         (ALL m k. m + k <= n --> FORALLBITS P (WSEG m k w)))"
   by (import word_bitop FORALLBITS_WSEG)
 
-lemma FORALLBITS_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool.
-   FORALLBITS P (WCAT (w1, w2)) = (FORALLBITS P w1 & FORALLBITS P w2)"
+lemma FORALLBITS_WCAT: "FORALLBITS P (WCAT (w1, w2)) = (FORALLBITS P w1 & FORALLBITS P w2)"
   by (import word_bitop FORALLBITS_WCAT)
 
 consts
   EXISTSABIT :: "('a => bool) => 'a word => bool" 
 
-specification (EXISTSABIT) EXISTSABIT_DEF: "ALL (P::'a::type => bool) l::'a::type list.
-   EXISTSABIT P (WORD l) = list_exists P l"
+specification (EXISTSABIT) EXISTSABIT_DEF: "ALL P l. EXISTSABIT P (WORD l) = list_ex P l"
   by (import word_bitop EXISTSABIT_DEF)
 
-lemma NOT_EXISTSABIT: "ALL (P::'a::type => bool) w::'a::type word.
-   (~ EXISTSABIT P w) = FORALLBITS (Not o P) w"
+lemma NOT_EXISTSABIT: "(~ EXISTSABIT P w) = FORALLBITS (Not o P) w"
   by (import word_bitop NOT_EXISTSABIT)
 
-lemma NOT_FORALLBITS: "ALL (P::'a::type => bool) w::'a::type word.
-   (~ FORALLBITS P w) = EXISTSABIT (Not o P) w"
+lemma NOT_FORALLBITS: "(~ FORALLBITS P w) = EXISTSABIT (Not o P) w"
   by (import word_bitop NOT_FORALLBITS)
 
-lemma EXISTSABIT: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::('a::type word => bool)
-                  => ('a::type word => bool) => bool)
-      ((PWORDLEN::nat => 'a::type word => bool) n)
-      (%w::'a::type word.
-          (All::(('a::type => bool) => bool) => bool)
-           (%P::'a::type => bool.
-               (op =::bool => bool => bool)
-                ((EXISTSABIT::('a::type => bool) => 'a::type word => bool) P
-                  w)
-                ((Ex::(nat => bool) => bool)
-                  (%k::nat.
-                      (op &::bool => bool => bool)
-                       ((op <::nat => nat => bool) k n)
-                       (P ((bit::nat => 'a::type word => 'a::type) k
-                            w)))))))"
+lemma EXISTSABIT: "RES_FORALL (PWORDLEN n) (%w. ALL P. EXISTSABIT P w = (EX k<n. P (bit k w)))"
   by (import word_bitop EXISTSABIT)
 
-lemma EXISTSABIT_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k <= n -->
-           (ALL P::'a::type => bool.
-               EXISTSABIT P (WSEG m k w) --> EXISTSABIT P w))"
+lemma EXISTSABIT_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k <= n -->
+         (ALL P. EXISTSABIT P (WSEG m k w) --> EXISTSABIT P w))"
   by (import word_bitop EXISTSABIT_WSEG)
 
-lemma EXISTSABIT_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool.
-   EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)"
+lemma EXISTSABIT_WCAT: "EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)"
   by (import word_bitop EXISTSABIT_WCAT)
 
-definition SHR :: "bool => 'a => 'a word => 'a word * 'a" where 
+definition
+  SHR :: "bool => 'a => 'a word => 'a word * 'a"  where
   "SHR ==
-%(f::bool) (b::'a::type) w::'a::type word.
+%f b w.
    (WCAT
      (if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b],
       WSEG (PRE (WORDLEN w)) 1 w),
     bit 0 w)"
 
-lemma SHR_DEF: "ALL (f::bool) (b::'a::type) w::'a::type word.
-   SHR f b w =
-   (WCAT
-     (if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b],
-      WSEG (PRE (WORDLEN w)) 1 w),
-    bit 0 w)"
+lemma SHR_DEF: "SHR f b w =
+(WCAT
+  (if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b],
+   WSEG (PRE (WORDLEN w)) 1 w),
+ bit 0 w)"
   by (import word_bitop SHR_DEF)
 
-definition SHL :: "bool => 'a word => 'a => 'a * 'a word" where 
+definition
+  SHL :: "bool => 'a word => 'a => 'a * 'a word"  where
   "SHL ==
-%(f::bool) (w::'a::type word) b::'a::type.
+%f w b.
    (bit (PRE (WORDLEN w)) w,
     WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))"
 
-lemma SHL_DEF: "ALL (f::bool) (w::'a::type word) b::'a::type.
-   SHL f w b =
-   (bit (PRE (WORDLEN w)) w,
-    WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))"
+lemma SHL_DEF: "SHL f w b =
+(bit (PRE (WORDLEN w)) w,
+ WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))"
   by (import word_bitop SHL_DEF)
 
-lemma SHR_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k <= n -->
-           0 < m -->
-           (ALL (f::bool) b::'a::type.
-               SHR f b (WSEG m k w) =
-               (if f
-                then WCAT (WSEG 1 (k + (m - 1)) w, WSEG (m - 1) (k + 1) w)
-                else WCAT (WORD [b], WSEG (m - 1) (k + 1) w),
-                bit k w)))"
+lemma SHR_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k <= n -->
+         0 < m -->
+         (ALL f b.
+             SHR f b (WSEG m k w) =
+             (if f
+              then WCAT (WSEG 1 (k + (m - 1)) w, WSEG (m - 1) (k + 1) w)
+              else WCAT (WORD [b], WSEG (m - 1) (k + 1) w),
+              bit k w)))"
   by (import word_bitop SHR_WSEG)
 
-lemma SHR_WSEG_1F: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (b::'a::type) (m::nat) k::nat.
-           m + k <= n -->
-           0 < m -->
-           SHR False b (WSEG m k w) =
-           (WCAT (WORD [b], WSEG (m - 1) (k + 1) w), bit k w))"
+lemma SHR_WSEG_1F: "RES_FORALL (PWORDLEN n)
+ (%w. ALL b m k.
+         m + k <= n -->
+         0 < m -->
+         SHR False b (WSEG m k w) =
+         (WCAT (WORD [b], WSEG (m - 1) (k + 1) w), bit k w))"
   by (import word_bitop SHR_WSEG_1F)
 
-lemma SHR_WSEG_NF: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k < n -->
-           0 < m -->
-           SHR False (bit (m + k) w) (WSEG m k w) =
-           (WSEG m (k + 1) w, bit k w))"
+lemma SHR_WSEG_NF: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k < n -->
+         0 < m -->
+         SHR False (bit (m + k) w) (WSEG m k w) =
+         (WSEG m (k + 1) w, bit k w))"
   by (import word_bitop SHR_WSEG_NF)
 
-lemma SHL_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k <= n -->
-           0 < m -->
-           (ALL (f::bool) b::'a::type.
-               SHL f (WSEG m k w) b =
-               (bit (k + (m - 1)) w,
-                if f then WCAT (WSEG (m - 1) k w, WSEG 1 k w)
-                else WCAT (WSEG (m - 1) k w, WORD [b]))))"
+lemma SHL_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k <= n -->
+         0 < m -->
+         (ALL f b.
+             SHL f (WSEG m k w) b =
+             (bit (k + (m - 1)) w,
+              if f then WCAT (WSEG (m - 1) k w, WSEG 1 k w)
+              else WCAT (WSEG (m - 1) k w, WORD [b]))))"
   by (import word_bitop SHL_WSEG)
 
-lemma SHL_WSEG_1F: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (b::'a::type) (m::nat) k::nat.
-           m + k <= n -->
-           0 < m -->
-           SHL False (WSEG m k w) b =
-           (bit (k + (m - 1)) w, WCAT (WSEG (m - 1) k w, WORD [b])))"
+lemma SHL_WSEG_1F: "RES_FORALL (PWORDLEN n)
+ (%w. ALL b m k.
+         m + k <= n -->
+         0 < m -->
+         SHL False (WSEG m k w) b =
+         (bit (k + (m - 1)) w, WCAT (WSEG (m - 1) k w, WORD [b])))"
   by (import word_bitop SHL_WSEG_1F)
 
-lemma SHL_WSEG_NF: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           m + k <= n -->
-           0 < m -->
-           0 < k -->
-           SHL False (WSEG m k w) (bit (k - 1) w) =
-           (bit (k + (m - 1)) w, WSEG m (k - 1) w))"
+lemma SHL_WSEG_NF: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m k.
+         m + k <= n -->
+         0 < m -->
+         0 < k -->
+         SHL False (WSEG m k w) (bit (k - 1) w) =
+         (bit (k + (m - 1)) w, WSEG m (k - 1) w))"
   by (import word_bitop SHL_WSEG_NF)
 
-lemma WSEG_SHL: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w::'a::type word.
-        ALL (m::nat) k::nat.
-           0 < k & m + k <= Suc n -->
-           (ALL b::'a::type.
-               WSEG m k (snd (SHL (f::bool) w b)) = WSEG m (k - 1) w))"
+lemma WSEG_SHL: "RES_FORALL (PWORDLEN (Suc n))
+ (%w. ALL m k.
+         0 < k & m + k <= Suc n -->
+         (ALL b. WSEG m k (snd (SHL f w b)) = WSEG m (k - 1) w))"
   by (import word_bitop WSEG_SHL)
 
-lemma WSEG_SHL_0: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w::'a::type word.
-        ALL (m::nat) b::'a::type.
-           0 < m & m <= Suc n -->
-           WSEG m 0 (snd (SHL (f::bool) w b)) =
-           WCAT (WSEG (m - 1) 0 w, if f then WSEG 1 0 w else WORD [b]))"
+lemma WSEG_SHL_0: "RES_FORALL (PWORDLEN (Suc n))
+ (%w. ALL m b.
+         0 < m & m <= Suc n -->
+         WSEG m 0 (snd (SHL f w b)) =
+         WCAT (WSEG (m - 1) 0 w, if f then WSEG 1 0 w else WORD [b]))"
   by (import word_bitop WSEG_SHL_0)
 
 ;end_setup
 
 ;setup_theory bword_num
 
-definition BV :: "bool => nat" where 
-  "BV == %b::bool. if b then Suc 0 else 0"
+definition
+  BV :: "bool => nat"  where
+  "BV == %b. if b then Suc 0 else 0"
 
-lemma BV_DEF: "ALL b::bool. BV b = (if b then Suc 0 else 0)"
+lemma BV_DEF: "BV b = (if b then Suc 0 else 0)"
   by (import bword_num BV_DEF)
 
 consts
   BNVAL :: "bool word => nat" 
 
-specification (BNVAL) BNVAL_DEF: "ALL l::bool list. BNVAL (WORD l) = LVAL BV 2 l"
+specification (BNVAL) BNVAL_DEF: "ALL l. BNVAL (WORD l) = LVAL BV 2 l"
   by (import bword_num BNVAL_DEF)
 
-lemma BV_LESS_2: "ALL x::bool. BV x < 2"
+lemma BV_LESS_2: "BV x < 2"
   by (import bword_num BV_LESS_2)
 
-lemma BNVAL_NVAL: "ALL w::bool word. BNVAL w = NVAL BV 2 w"
+lemma BNVAL_NVAL: "BNVAL w = NVAL BV 2 w"
   by (import bword_num BNVAL_NVAL)
 
 lemma BNVAL0: "BNVAL (WORD []) = 0"
   by (import bword_num BNVAL0)
 
-lemma BNVAL_11: "ALL (w1::bool word) w2::bool word.
-   WORDLEN w1 = WORDLEN w2 --> BNVAL w1 = BNVAL w2 --> w1 = w2"
+lemma BNVAL_11: "[| WORDLEN w1 = WORDLEN w2; BNVAL w1 = BNVAL w2 |] ==> w1 = w2"
   by (import bword_num BNVAL_11)
 
-lemma BNVAL_ONTO: "ALL w::bool word. Ex (op = (BNVAL w))"
+lemma BNVAL_ONTO: "Ex (op = (BNVAL w))"
   by (import bword_num BNVAL_ONTO)
 
-lemma BNVAL_MAX: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. BNVAL w < 2 ^ n)"
+lemma BNVAL_MAX: "RES_FORALL (PWORDLEN n) (%w. BNVAL w < 2 ^ n)"
   by (import bword_num BNVAL_MAX)
 
-lemma BNVAL_WCAT1: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::bool word.
-        ALL x::bool. BNVAL (WCAT (w, WORD [x])) = BNVAL w * 2 + BV x)"
+lemma BNVAL_WCAT1: "RES_FORALL (PWORDLEN n)
+ (%w. ALL x. BNVAL (WCAT (w, WORD [x])) = BNVAL w * 2 + BV x)"
   by (import bword_num BNVAL_WCAT1)
 
-lemma BNVAL_WCAT2: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w::bool word.
-        ALL x::bool. BNVAL (WCAT (WORD [x], w)) = BV x * 2 ^ n + BNVAL w)"
+lemma BNVAL_WCAT2: "RES_FORALL (PWORDLEN n)
+ (%w. ALL x. BNVAL (WCAT (WORD [x], w)) = BV x * 2 ^ n + BNVAL w)"
   by (import bword_num BNVAL_WCAT2)
 
-lemma BNVAL_WCAT: "ALL (n::nat) m::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN m)
-         (%w2::bool word.
-             BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))"
+lemma BNVAL_WCAT: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN m)
+        (%w2. BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))"
   by (import bword_num BNVAL_WCAT)
 
-definition VB :: "nat => bool" where 
-  "VB == %n::nat. n mod 2 ~= 0"
+definition
+  VB :: "nat => bool"  where
+  "VB == %n. n mod 2 ~= 0"
 
-lemma VB_DEF: "ALL n::nat. VB n = (n mod 2 ~= 0)"
+lemma VB_DEF: "VB n = (n mod 2 ~= 0)"
   by (import bword_num VB_DEF)
 
-definition NBWORD :: "nat => nat => bool word" where 
-  "NBWORD == %(n::nat) m::nat. WORD (NLIST n VB 2 m)"
+definition
+  NBWORD :: "nat => nat => bool word"  where
+  "NBWORD == %n m. WORD (NLIST n VB 2 m)"
 
-lemma NBWORD_DEF: "ALL (n::nat) m::nat. NBWORD n m = WORD (NLIST n VB 2 m)"
+lemma NBWORD_DEF: "NBWORD n m = WORD (NLIST n VB 2 m)"
   by (import bword_num NBWORD_DEF)
 
-lemma NBWORD0: "ALL x::nat. NBWORD 0 x = WORD []"
+lemma NBWORD0: "NBWORD 0 x = WORD []"
   by (import bword_num NBWORD0)
 
-lemma WORDLEN_NBWORD: "ALL (x::nat) xa::nat. WORDLEN (NBWORD x xa) = x"
+lemma WORDLEN_NBWORD: "WORDLEN (NBWORD x xa) = x"
   by (import bword_num WORDLEN_NBWORD)
 
-lemma PWORDLEN_NBWORD: "ALL (x::nat) xa::nat. IN (NBWORD x xa) (PWORDLEN x)"
+lemma PWORDLEN_NBWORD: "IN (NBWORD x xa) (PWORDLEN x)"
   by (import bword_num PWORDLEN_NBWORD)
 
-lemma NBWORD_SUC: "ALL (n::nat) m::nat.
-   NBWORD (Suc n) m = WCAT (NBWORD n (m div 2), WORD [VB (m mod 2)])"
+lemma NBWORD_SUC: "NBWORD (Suc n) m = WCAT (NBWORD n (m div 2), WORD [VB (m mod 2)])"
   by (import bword_num NBWORD_SUC)
 
-lemma VB_BV: "ALL x::bool. VB (BV x) = x"
+lemma VB_BV: "VB (BV x) = x"
   by (import bword_num VB_BV)
 
-lemma BV_VB: "(All::(nat => bool) => bool)
- (%x::nat.
-     (op -->::bool => bool => bool)
-      ((op <::nat => nat => bool) x
-        ((number_of \<Colon> int => nat)
-          ((Int.Bit0 \<Colon> int => int)
-            ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int)))))
-      ((op =::nat => nat => bool) ((BV::bool => nat) ((VB::nat => bool) x))
-        x))"
+lemma BV_VB: "x < 2 ==> BV (VB x) = x"
   by (import bword_num BV_VB)
 
-lemma NBWORD_BNVAL: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. NBWORD n (BNVAL w) = w)"
+lemma NBWORD_BNVAL: "RES_FORALL (PWORDLEN n) (%w. NBWORD n (BNVAL w) = w)"
   by (import bword_num NBWORD_BNVAL)
 
-lemma BNVAL_NBWORD: "ALL (n::nat) m::nat. m < 2 ^ n --> BNVAL (NBWORD n m) = m"
+lemma BNVAL_NBWORD: "m < 2 ^ n ==> BNVAL (NBWORD n m) = m"
   by (import bword_num BNVAL_NBWORD)
 
-lemma ZERO_WORD_VAL: "RES_FORALL (PWORDLEN (n::nat))
- (%w::bool word. (w = NBWORD n 0) = (BNVAL w = 0))"
+lemma ZERO_WORD_VAL: "RES_FORALL (PWORDLEN n) (%w. (w = NBWORD n 0) = (BNVAL w = 0))"
   by (import bword_num ZERO_WORD_VAL)
 
-lemma WCAT_NBWORD_0: "ALL (n1::nat) n2::nat. WCAT (NBWORD n1 0, NBWORD n2 0) = NBWORD (n1 + n2) 0"
+lemma WCAT_NBWORD_0: "WCAT (NBWORD n1 0, NBWORD n2 0) = NBWORD (n1 + n2) 0"
   by (import bword_num WCAT_NBWORD_0)
 
-lemma WSPLIT_NBWORD_0: "ALL (n::nat) m::nat.
-   m <= n --> WSPLIT m (NBWORD n 0) = (NBWORD (n - m) 0, NBWORD m 0)"
+lemma WSPLIT_NBWORD_0: "m <= n ==> WSPLIT m (NBWORD n 0) = (NBWORD (n - m) 0, NBWORD m 0)"
   by (import bword_num WSPLIT_NBWORD_0)
 
-lemma EQ_NBWORD0_SPLIT: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::(bool word => bool) => (bool word => bool) => bool)
-      ((PWORDLEN::nat => bool word => bool) n)
-      (%w::bool word.
-          (All::(nat => bool) => bool)
-           (%m::nat.
-               (op -->::bool => bool => bool)
-                ((op <=::nat => nat => bool) m n)
-                ((op =::bool => bool => bool)
-                  ((op =::bool word => bool word => bool) w
-                    ((NBWORD::nat => nat => bool word) n (0::nat)))
-                  ((op &::bool => bool => bool)
-                    ((op =::bool word => bool word => bool)
-                      ((WSEG::nat => nat => bool word => bool word)
-                        ((op -::nat => nat => nat) n m) m w)
-                      ((NBWORD::nat => nat => bool word)
-                        ((op -::nat => nat => nat) n m) (0::nat)))
-                    ((op =::bool word => bool word => bool)
-                      ((WSEG::nat => nat => bool word => bool word) m
-                        (0::nat) w)
-                      ((NBWORD::nat => nat => bool word) m (0::nat))))))))"
+lemma EQ_NBWORD0_SPLIT: "RES_FORALL (PWORDLEN n)
+ (%w. ALL m<=n.
+         (w = NBWORD n 0) =
+         (WSEG (n - m) m w = NBWORD (n - m) 0 & WSEG m 0 w = NBWORD m 0))"
   by (import bword_num EQ_NBWORD0_SPLIT)
 
-lemma NBWORD_MOD: "ALL (n::nat) m::nat. NBWORD n (m mod 2 ^ n) = NBWORD n m"
+lemma NBWORD_MOD: "NBWORD n (m mod 2 ^ n) = NBWORD n m"
   by (import bword_num NBWORD_MOD)
 
-lemma WSEG_NBWORD_SUC: "ALL (n::nat) m::nat. WSEG n 0 (NBWORD (Suc n) m) = NBWORD n m"
+lemma WSEG_NBWORD_SUC: "WSEG n 0 (NBWORD (Suc n) m) = NBWORD n m"
   by (import bword_num WSEG_NBWORD_SUC)
 
-lemma NBWORD_SUC_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w::bool word. NBWORD n (BNVAL w) = WSEG n 0 w)"
+lemma NBWORD_SUC_WSEG: "RES_FORALL (PWORDLEN (Suc n)) (%w. NBWORD n (BNVAL w) = WSEG n 0 w)"
   by (import bword_num NBWORD_SUC_WSEG)
 
-lemma DOUBL_EQ_SHL: "ALL x>0.
-   RES_FORALL (PWORDLEN x)
-    (%xa::bool word.
-        ALL xb::bool.
-           NBWORD x (BNVAL xa + BNVAL xa + BV xb) = snd (SHL False xa xb))"
+lemma DOUBL_EQ_SHL: "0 < x
+==> RES_FORALL (PWORDLEN x)
+     (%xa. ALL xb.
+              NBWORD x (BNVAL xa + BNVAL xa + BV xb) =
+              snd (SHL False xa xb))"
   by (import bword_num DOUBL_EQ_SHL)
 
-lemma MSB_NBWORD: "ALL (n::nat) m::nat. bit n (NBWORD (Suc n) m) = VB (m div 2 ^ n mod 2)"
+lemma MSB_NBWORD: "bit n (NBWORD (Suc n) m) = VB (m div 2 ^ n mod 2)"
   by (import bword_num MSB_NBWORD)
 
-lemma NBWORD_SPLIT: "ALL (n1::nat) (n2::nat) m::nat.
-   NBWORD (n1 + n2) m = WCAT (NBWORD n1 (m div 2 ^ n2), NBWORD n2 m)"
+lemma NBWORD_SPLIT: "NBWORD (n1 + n2) m = WCAT (NBWORD n1 (m div 2 ^ n2), NBWORD n2 m)"
   by (import bword_num NBWORD_SPLIT)
 
-lemma WSEG_NBWORD: "ALL (m::nat) (k::nat) n::nat.
-   m + k <= n -->
-   (ALL l::nat. WSEG m k (NBWORD n l) = NBWORD m (l div 2 ^ k))"
+lemma WSEG_NBWORD: "m + k <= n ==> WSEG m k (NBWORD n l) = NBWORD m (l div 2 ^ k)"
   by (import bword_num WSEG_NBWORD)
 
-lemma NBWORD_SUC_FST: "ALL (n::nat) x::nat.
-   NBWORD (Suc n) x = WCAT (WORD [VB (x div 2 ^ n mod 2)], NBWORD n x)"
+lemma NBWORD_SUC_FST: "NBWORD (Suc n) x = WCAT (WORD [VB (x div 2 ^ n mod 2)], NBWORD n x)"
   by (import bword_num NBWORD_SUC_FST)
 
-lemma BIT_NBWORD0: "ALL (k::nat) n::nat. k < n --> bit k (NBWORD n 0) = False"
+lemma BIT_NBWORD0: "k < n ==> bit k (NBWORD n 0) = False"
   by (import bword_num BIT_NBWORD0)
 
-lemma ADD_BNVAL_LEFT: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN (Suc n))
-         (%w2::bool word.
-             BNVAL w1 + BNVAL w2 =
-             (BV (bit n w1) + BV (bit n w2)) * 2 ^ n +
-             (BNVAL (WSEG n 0 w1) + BNVAL (WSEG n 0 w2))))"
+lemma ADD_BNVAL_LEFT: "RES_FORALL (PWORDLEN (Suc n))
+ (%w1. RES_FORALL (PWORDLEN (Suc n))
+        (%w2. BNVAL w1 + BNVAL w2 =
+              (BV (bit n w1) + BV (bit n w2)) * 2 ^ n +
+              (BNVAL (WSEG n 0 w1) + BNVAL (WSEG n 0 w2))))"
   by (import bword_num ADD_BNVAL_LEFT)
 
-lemma ADD_BNVAL_RIGHT: "ALL n::nat.
-   RES_FORALL (PWORDLEN (Suc n))
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN (Suc n))
-         (%w2::bool word.
-             BNVAL w1 + BNVAL w2 =
-             (BNVAL (WSEG n 1 w1) + BNVAL (WSEG n 1 w2)) * 2 +
-             (BV (bit 0 w1) + BV (bit 0 w2))))"
+lemma ADD_BNVAL_RIGHT: "RES_FORALL (PWORDLEN (Suc n))
+ (%w1. RES_FORALL (PWORDLEN (Suc n))
+        (%w2. BNVAL w1 + BNVAL w2 =
+              (BNVAL (WSEG n 1 w1) + BNVAL (WSEG n 1 w2)) * 2 +
+              (BV (bit 0 w1) + BV (bit 0 w2))))"
   by (import bword_num ADD_BNVAL_RIGHT)
 
-lemma ADD_BNVAL_SPLIT: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN (n1 + n2))
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN (n1 + n2))
-         (%w2::bool word.
-             BNVAL w1 + BNVAL w2 =
-             (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2)) * 2 ^ n2 +
-             (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2))))"
+lemma ADD_BNVAL_SPLIT: "RES_FORALL (PWORDLEN (n1 + n2))
+ (%w1. RES_FORALL (PWORDLEN (n1 + n2))
+        (%w2. BNVAL w1 + BNVAL w2 =
+              (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2)) * 2 ^ n2 +
+              (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2))))"
   by (import bword_num ADD_BNVAL_SPLIT)
 
 ;end_setup
@@ -1396,8 +924,8 @@
 consts
   ACARRY :: "nat => bool word => bool word => bool => bool" 
 
-specification (ACARRY) ACARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. ACARRY 0 w1 w2 cin = cin) &
-(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool.
+specification (ACARRY) ACARRY_DEF: "(ALL w1 w2 cin. ACARRY 0 w1 w2 cin = cin) &
+(ALL n w1 w2 cin.
     ACARRY (Suc n) w1 w2 cin =
     VB ((BV (bit n w1) + BV (bit n w2) + BV (ACARRY n w1 w2 cin)) div 2))"
   by (import bword_arith ACARRY_DEF)
@@ -1405,185 +933,122 @@
 consts
   ICARRY :: "nat => bool word => bool word => bool => bool" 
 
-specification (ICARRY) ICARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. ICARRY 0 w1 w2 cin = cin) &
-(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool.
+specification (ICARRY) ICARRY_DEF: "(ALL w1 w2 cin. ICARRY 0 w1 w2 cin = cin) &
+(ALL n w1 w2 cin.
     ICARRY (Suc n) w1 w2 cin =
     (bit n w1 & bit n w2 | (bit n w1 | bit n w2) & ICARRY n w1 w2 cin))"
   by (import bword_arith ICARRY_DEF)
 
-lemma ACARRY_EQ_ICARRY: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL (cin::bool) k::nat.
-                k <= n --> ACARRY k w1 w2 cin = ICARRY k w1 w2 cin))"
+lemma ACARRY_EQ_ICARRY: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL cin k.
+                 k <= n --> ACARRY k w1 w2 cin = ICARRY k w1 w2 cin))"
   by (import bword_arith ACARRY_EQ_ICARRY)
 
-lemma BNVAL_LESS_EQ: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. BNVAL w <= 2 ^ n - 1)"
+lemma BNVAL_LESS_EQ: "RES_FORALL (PWORDLEN n) (%w. BNVAL w <= 2 ^ n - 1)"
   by (import bword_arith BNVAL_LESS_EQ)
 
-lemma ADD_BNVAL_LESS_EQ1: "ALL (n::nat) cin::bool.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n <= Suc 0))"
+lemma ADD_BNVAL_LESS_EQ1: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n <= Suc 0))"
   by (import bword_arith ADD_BNVAL_LESS_EQ1)
 
-lemma ADD_BV_BNVAL_DIV_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             (BV x1 + BV x2 +
-              (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n) div
-             2
-             <= 1))"
+lemma ADD_BV_BNVAL_DIV_LESS_EQ1: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. (BV x1 + BV x2 +
+               (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n) div
+              2
+              <= 1))"
   by (import bword_arith ADD_BV_BNVAL_DIV_LESS_EQ1)
 
-lemma ADD_BV_BNVAL_LESS_EQ: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))
-             <= Suc (2 ^ Suc n)))"
+lemma ADD_BV_BNVAL_LESS_EQ: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))
+              <= Suc (2 ^ Suc n)))"
   by (import bword_arith ADD_BV_BNVAL_LESS_EQ)
 
-lemma ADD_BV_BNVAL_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             (BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))) div
-             2 ^ Suc n
-             <= 1))"
+lemma ADD_BV_BNVAL_LESS_EQ1: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. (BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))) div
+              2 ^ Suc n
+              <= 1))"
   by (import bword_arith ADD_BV_BNVAL_LESS_EQ1)
 
-lemma ACARRY_EQ_ADD_DIV: "(All::(nat => bool) => bool)
- (%n::nat.
-     (RES_FORALL::(bool word => bool) => (bool word => bool) => bool)
-      ((PWORDLEN::nat => bool word => bool) n)
-      (%w1::bool word.
-          (RES_FORALL::(bool word => bool) => (bool word => bool) => bool)
-           ((PWORDLEN::nat => bool word => bool) n)
-           (%w2::bool word.
-               (All::(nat => bool) => bool)
-                (%k::nat.
-                    (op -->::bool => bool => bool)
-                     ((op <::nat => nat => bool) k n)
-                     ((op =::nat => nat => bool)
-                       ((BV::bool => nat)
-                         ((ACARRY::nat
-                                   => bool word
-=> bool word => bool => bool)
-                           k w1 w2 (cin::bool)))
-                       ((op div::nat => nat => nat)
-                         ((op +::nat => nat => nat)
-                           ((op +::nat => nat => nat)
-                             ((BNVAL::bool word => nat)
-                               ((WSEG::nat => nat => bool word => bool word)
-                                 k (0::nat) w1))
-                             ((BNVAL::bool word => nat)
-                               ((WSEG::nat => nat => bool word => bool word)
-                                 k (0::nat) w2)))
-                           ((BV::bool => nat) cin))
-                         ((op ^::nat => nat => nat)
-                           ((number_of \<Colon> int => nat)
-                             ((Int.Bit0 \<Colon> int => int)
-                               ((Int.Bit1 \<Colon> int => int)
-                                 (Int.Pls \<Colon> int))))
-                           k)))))))"
+lemma ACARRY_EQ_ADD_DIV: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL k<n.
+                 BV (ACARRY k w1 w2 cin) =
+                 (BNVAL (WSEG k 0 w1) + BNVAL (WSEG k 0 w2) + BV cin) div
+                 2 ^ k))"
   by (import bword_arith ACARRY_EQ_ADD_DIV)
 
-lemma ADD_WORD_SPLIT: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN (n1 + n2))
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN (n1 + n2))
-         (%w2::bool word.
-             ALL cin::bool.
-                NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) =
-                WCAT
-                 (NBWORD n1
-                   (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) +
-                    BV (ACARRY n2 w1 w2 cin)),
-                  NBWORD n2
-                   (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + BV cin))))"
+lemma ADD_WORD_SPLIT: "RES_FORALL (PWORDLEN (n1 + n2))
+ (%w1. RES_FORALL (PWORDLEN (n1 + n2))
+        (%w2. ALL cin.
+                 NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) =
+                 WCAT
+                  (NBWORD n1
+                    (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) +
+                     BV (ACARRY n2 w1 w2 cin)),
+                   NBWORD n2
+                    (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) +
+                     BV cin))))"
   by (import bword_arith ADD_WORD_SPLIT)
 
-lemma WSEG_NBWORD_ADD: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL (m::nat) (k::nat) cin::bool.
-                m + k <= n -->
-                WSEG m k (NBWORD n (BNVAL w1 + BNVAL w2 + BV cin)) =
-                NBWORD m
-                 (BNVAL (WSEG m k w1) + BNVAL (WSEG m k w2) +
-                  BV (ACARRY k w1 w2 cin))))"
+lemma WSEG_NBWORD_ADD: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL m k cin.
+                 m + k <= n -->
+                 WSEG m k (NBWORD n (BNVAL w1 + BNVAL w2 + BV cin)) =
+                 NBWORD m
+                  (BNVAL (WSEG m k w1) + BNVAL (WSEG m k w2) +
+                   BV (ACARRY k w1 w2 cin))))"
   by (import bword_arith WSEG_NBWORD_ADD)
 
-lemma ADD_NBWORD_EQ0_SPLIT: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN (n1 + n2))
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN (n1 + n2))
-         (%w2::bool word.
-             ALL cin::bool.
-                (NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) =
-                 NBWORD (n1 + n2) 0) =
-                (NBWORD n1
-                  (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) +
-                   BV (ACARRY n2 w1 w2 cin)) =
-                 NBWORD n1 0 &
-                 NBWORD n2
-                  (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + BV cin) =
-                 NBWORD n2 0)))"
+lemma ADD_NBWORD_EQ0_SPLIT: "RES_FORALL (PWORDLEN (n1 + n2))
+ (%w1. RES_FORALL (PWORDLEN (n1 + n2))
+        (%w2. ALL cin.
+                 (NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) =
+                  NBWORD (n1 + n2) 0) =
+                 (NBWORD n1
+                   (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) +
+                    BV (ACARRY n2 w1 w2 cin)) =
+                  NBWORD n1 0 &
+                  NBWORD n2
+                   (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + BV cin) =
+                  NBWORD n2 0)))"
   by (import bword_arith ADD_NBWORD_EQ0_SPLIT)
 
-lemma ACARRY_MSB: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL cin::bool.
-                ACARRY n w1 w2 cin =
-                bit n (NBWORD (Suc n) (BNVAL w1 + BNVAL w2 + BV cin))))"
+lemma ACARRY_MSB: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL cin.
+                 ACARRY n w1 w2 cin =
+                 bit n (NBWORD (Suc n) (BNVAL w1 + BNVAL w2 + BV cin))))"
   by (import bword_arith ACARRY_MSB)
 
-lemma ACARRY_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL (cin::bool) (k::nat) m::nat.
-                k < m & m <= n -->
-                ACARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin =
-                ACARRY k w1 w2 cin))"
+lemma ACARRY_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL cin k m.
+                 k < m & m <= n -->
+                 ACARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin =
+                 ACARRY k w1 w2 cin))"
   by (import bword_arith ACARRY_WSEG)
 
-lemma ICARRY_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL (cin::bool) (k::nat) m::nat.
-                k < m & m <= n -->
-                ICARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin =
-                ICARRY k w1 w2 cin))"
+lemma ICARRY_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL cin k m.
+                 k < m & m <= n -->
+                 ICARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin =
+                 ICARRY k w1 w2 cin))"
   by (import bword_arith ICARRY_WSEG)
 
-lemma ACARRY_ACARRY_WSEG: "ALL n::nat.
-   RES_FORALL (PWORDLEN n)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n)
-         (%w2::bool word.
-             ALL (cin::bool) (m::nat) (k1::nat) k2::nat.
-                k1 < m & k2 < n & m + k2 <= n -->
-                ACARRY k1 (WSEG m k2 w1) (WSEG m k2 w2)
-                 (ACARRY k2 w1 w2 cin) =
-                ACARRY (k1 + k2) w1 w2 cin))"
+lemma ACARRY_ACARRY_WSEG: "RES_FORALL (PWORDLEN n)
+ (%w1. RES_FORALL (PWORDLEN n)
+        (%w2. ALL cin m k1 k2.
+                 k1 < m & k2 < n & m + k2 <= n -->
+                 ACARRY k1 (WSEG m k2 w1) (WSEG m k2 w2)
+                  (ACARRY k2 w1 w2 cin) =
+                 ACARRY (k1 + k2) w1 w2 cin))"
   by (import bword_arith ACARRY_ACARRY_WSEG)
 
 ;end_setup
@@ -1593,27 +1058,24 @@
 consts
   WNOT :: "bool word => bool word" 
 
-specification (WNOT) WNOT_DEF: "ALL l::bool list. WNOT (WORD l) = WORD (map Not l)"
+specification (WNOT) WNOT_DEF: "ALL l. WNOT (WORD l) = WORD (map Not l)"
   by (import bword_bitop WNOT_DEF)
 
 lemma PBITOP_WNOT: "IN WNOT PBITOP"
   by (import bword_bitop PBITOP_WNOT)
 
-lemma WNOT_WNOT: "ALL w::bool word. WNOT (WNOT w) = w"
+lemma WNOT_WNOT: "WNOT (WNOT w) = w"
   by (import bword_bitop WNOT_WNOT)
 
-lemma WCAT_WNOT: "ALL (n1::nat) n2::nat.
-   RES_FORALL (PWORDLEN n1)
-    (%w1::bool word.
-        RES_FORALL (PWORDLEN n2)
-         (%w2::bool word. WCAT (WNOT w1, WNOT w2) = WNOT (WCAT (w1, w2))))"
+lemma WCAT_WNOT: "RES_FORALL (PWORDLEN n1)
+ (%w1. RES_FORALL (PWORDLEN n2)
+        (%w2. WCAT (WNOT w1, WNOT w2) = WNOT (WCAT (w1, w2))))"
   by (import bword_bitop WCAT_WNOT)
 
 consts
   WAND :: "bool word => bool word => bool word" 
 
-specification (WAND) WAND_DEF: "ALL (l1::bool list) l2::bool list.
-   WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)"
+specification (WAND) WAND_DEF: "ALL l1 l2. WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)"
   by (import bword_bitop WAND_DEF)
 
 lemma PBITBOP_WAND: "IN WAND PBITBOP"
@@ -1622,8 +1084,7 @@
 consts
   WOR :: "bool word => bool word => bool word" 
 
-specification (WOR) WOR_DEF: "ALL (l1::bool list) l2::bool list.
-   WOR (WORD l1) (WORD l2) = WORD (map2 op | l1 l2)"
+specification (WOR) WOR_DEF: "ALL l1 l2. WOR (WORD l1) (WORD l2) = WORD (map2 op | l1 l2)"
   by (import bword_bitop WOR_DEF)
 
 lemma PBITBOP_WOR: "IN WOR PBITBOP"
@@ -1632,8 +1093,7 @@
 consts
   WXOR :: "bool word => bool word => bool word" 
 
-specification (WXOR) WXOR_DEF: "ALL (l1::bool list) l2::bool list.
-   WXOR (WORD l1) (WORD l2) = WORD (map2 (%(x::bool) y::bool. x ~= y) l1 l2)"
+specification (WXOR) WXOR_DEF: "ALL l1 l2. WXOR (WORD l1) (WORD l2) = WORD (map2 op ~= l1 l2)"
   by (import bword_bitop WXOR_DEF)
 
 lemma PBITBOP_WXOR: "IN WXOR PBITBOP"
--- a/src/HOL/Import/HOL/HOL4Word32.thy	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/HOL4Word32.thy	Tue Sep 06 22:41:35 2011 -0700
@@ -8,80 +8,79 @@
   DIV2 :: "nat => nat" 
 
 defs
-  DIV2_primdef: "DIV2 == %n::nat. n div 2"
+  DIV2_primdef: "DIV2 == %n. n div 2"
 
-lemma DIV2_def: "ALL n::nat. DIV2 n = n div 2"
+lemma DIV2_def: "DIV2 n = n div 2"
   by (import bits DIV2_def)
 
 consts
   TIMES_2EXP :: "nat => nat => nat" 
 
 defs
-  TIMES_2EXP_primdef: "TIMES_2EXP == %(x::nat) n::nat. n * 2 ^ x"
+  TIMES_2EXP_primdef: "TIMES_2EXP == %x n. n * 2 ^ x"
 
-lemma TIMES_2EXP_def: "ALL (x::nat) n::nat. TIMES_2EXP x n = n * 2 ^ x"
+lemma TIMES_2EXP_def: "TIMES_2EXP x n = n * 2 ^ x"
   by (import bits TIMES_2EXP_def)
 
 consts
   DIV_2EXP :: "nat => nat => nat" 
 
 defs
-  DIV_2EXP_primdef: "DIV_2EXP == %(x::nat) n::nat. n div 2 ^ x"
+  DIV_2EXP_primdef: "DIV_2EXP == %x n. n div 2 ^ x"
 
-lemma DIV_2EXP_def: "ALL (x::nat) n::nat. DIV_2EXP x n = n div 2 ^ x"
+lemma DIV_2EXP_def: "DIV_2EXP x n = n div 2 ^ x"
   by (import bits DIV_2EXP_def)
 
 consts
   MOD_2EXP :: "nat => nat => nat" 
 
 defs
-  MOD_2EXP_primdef: "MOD_2EXP == %(x::nat) n::nat. n mod 2 ^ x"
+  MOD_2EXP_primdef: "MOD_2EXP == %x n. n mod 2 ^ x"
 
-lemma MOD_2EXP_def: "ALL (x::nat) n::nat. MOD_2EXP x n = n mod 2 ^ x"
+lemma MOD_2EXP_def: "MOD_2EXP x n = n mod 2 ^ x"
   by (import bits MOD_2EXP_def)
 
 consts
   DIVMOD_2EXP :: "nat => nat => nat * nat" 
 
 defs
-  DIVMOD_2EXP_primdef: "DIVMOD_2EXP == %(x::nat) n::nat. (n div 2 ^ x, n mod 2 ^ x)"
+  DIVMOD_2EXP_primdef: "DIVMOD_2EXP == %x n. (n div 2 ^ x, n mod 2 ^ x)"
 
-lemma DIVMOD_2EXP_def: "ALL (x::nat) n::nat. DIVMOD_2EXP x n = (n div 2 ^ x, n mod 2 ^ x)"
+lemma DIVMOD_2EXP_def: "DIVMOD_2EXP x n = (n div 2 ^ x, n mod 2 ^ x)"
   by (import bits DIVMOD_2EXP_def)
 
 consts
   SBIT :: "bool => nat => nat" 
 
 defs
-  SBIT_primdef: "SBIT == %(b::bool) n::nat. if b then 2 ^ n else 0"
+  SBIT_primdef: "SBIT == %b n. if b then 2 ^ n else 0"
 
-lemma SBIT_def: "ALL (b::bool) n::nat. SBIT b n = (if b then 2 ^ n else 0)"
+lemma SBIT_def: "SBIT b n = (if b then 2 ^ n else 0)"
   by (import bits SBIT_def)
 
 consts
   BITS :: "nat => nat => nat => nat" 
 
 defs
-  BITS_primdef: "BITS == %(h::nat) (l::nat) n::nat. MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
+  BITS_primdef: "BITS == %h l n. MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
 
-lemma BITS_def: "ALL (h::nat) (l::nat) n::nat.
-   BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
+lemma BITS_def: "BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
   by (import bits BITS_def)
 
-definition bit :: "nat => nat => bool" where 
-  "bit == %(b::nat) n::nat. BITS b b n = 1"
+definition
+  bit :: "nat => nat => bool"  where
+  "bit == %b n. BITS b b n = 1"
 
-lemma BIT_def: "ALL (b::nat) n::nat. bit b n = (BITS b b n = 1)"
+lemma BIT_def: "bit b n = (BITS b b n = 1)"
   by (import bits BIT_def)
 
 consts
   SLICE :: "nat => nat => nat => nat" 
 
 defs
-  SLICE_primdef: "SLICE == %(h::nat) (l::nat) n::nat. MOD_2EXP (Suc h) n - MOD_2EXP l n"
+  SLICE_primdef: "SLICE == %h l n. MOD_2EXP (Suc h) n - MOD_2EXP l n"
 
-lemma SLICE_def: "ALL (h::nat) (l::nat) n::nat.
-   SLICE h l n = MOD_2EXP (Suc h) n - MOD_2EXP l n"
+lemma SLICE_def: "SLICE h l n = MOD_2EXP (Suc h) n - MOD_2EXP l n"
   by (import bits SLICE_def)
 
 consts
@@ -96,282 +95,192 @@
 consts
   BITWISE :: "nat => (bool => bool => bool) => nat => nat => nat" 
 
-specification (BITWISE_primdef: BITWISE) BITWISE_def: "(ALL (oper::bool => bool => bool) (x::nat) y::nat. BITWISE 0 oper x y = 0) &
-(ALL (n::nat) (oper::bool => bool => bool) (x::nat) y::nat.
+specification (BITWISE_primdef: BITWISE) BITWISE_def: "(ALL oper x y. BITWISE 0 oper x y = 0) &
+(ALL n oper x y.
     BITWISE (Suc n) oper x y =
     BITWISE n oper x y + SBIT (oper (bit n x) (bit n y)) n)"
   by (import bits BITWISE_def)
 
-lemma DIV1: "ALL x::nat. x div 1 = x"
-  by (import bits DIV1)
-
-lemma SUC_SUB: "Suc (a::nat) - a = 1"
+lemma SUC_SUB: "Suc a - a = 1"
   by (import bits SUC_SUB)
 
-lemma DIV_MULT_1: "ALL (r::nat) n::nat. r < n --> (n + r) div n = 1"
+lemma DIV_MULT_1: "(r::nat) < (n::nat) ==> (n + r) div n = (1::nat)"
   by (import bits DIV_MULT_1)
 
-lemma ZERO_LT_TWOEXP: "(All::(nat => bool) => bool)
- (%n::nat.
-     (op <::nat => nat => bool) (0::nat)
-      ((op ^::nat => nat => nat)
-        ((number_of \<Colon> int => nat)
-          ((Int.Bit0 \<Colon> int => int)
-            ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-        n))"
+lemma ZERO_LT_TWOEXP: "(0::nat) < (2::nat) ^ (n::nat)"
   by (import bits ZERO_LT_TWOEXP)
 
-lemma MOD_2EXP_LT: "ALL (n::nat) k::nat. k mod 2 ^ n < 2 ^ n"
+lemma MOD_2EXP_LT: "(k::nat) mod (2::nat) ^ (n::nat) < (2::nat) ^ n"
   by (import bits MOD_2EXP_LT)
 
-lemma TWOEXP_DIVISION: "ALL (n::nat) k::nat. k = k div 2 ^ n * 2 ^ n + k mod 2 ^ n"
+lemma TWOEXP_DIVISION: "(k::nat) = k div (2::nat) ^ (n::nat) * (2::nat) ^ n + k mod (2::nat) ^ n"
   by (import bits TWOEXP_DIVISION)
 
-lemma TWOEXP_MONO: "(All::(nat => bool) => bool)
- (%a::nat.
-     (All::(nat => bool) => bool)
-      (%b::nat.
-          (op -->::bool => bool => bool) ((op <::nat => nat => bool) a b)
-           ((op <::nat => nat => bool)
-             ((op ^::nat => nat => nat)
-               ((number_of \<Colon> int => nat)
-                 ((Int.Bit0 \<Colon> int => int)
-                   ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-               a)
-             ((op ^::nat => nat => nat)
-               ((number_of \<Colon> int => nat)
-                 ((Int.Bit0 \<Colon> int => int)
-                   ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-               b))))"
+lemma TWOEXP_MONO: "(a::nat) < (b::nat) ==> (2::nat) ^ a < (2::nat) ^ b"
   by (import bits TWOEXP_MONO)
 
-lemma TWOEXP_MONO2: "(All::(nat => bool) => bool)
- (%a::nat.
-     (All::(nat => bool) => bool)
-      (%b::nat.
-          (op -->::bool => bool => bool) ((op <=::nat => nat => bool) a b)
-           ((op <=::nat => nat => bool)
-             ((op ^::nat => nat => nat)
-               ((number_of \<Colon> int => nat)
-                 ((Int.Bit0 \<Colon> int => int)
-                   ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-               a)
-             ((op ^::nat => nat => nat)
-               ((number_of \<Colon> int => nat)
-                 ((Int.Bit0 \<Colon> int => int)
-                   ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-               b))))"
+lemma TWOEXP_MONO2: "(a::nat) <= (b::nat) ==> (2::nat) ^ a <= (2::nat) ^ b"
   by (import bits TWOEXP_MONO2)
 
-lemma EXP_SUB_LESS_EQ: "(All::(nat => bool) => bool)
- (%a::nat.
-     (All::(nat => bool) => bool)
-      (%b::nat.
-          (op <=::nat => nat => bool)
-           ((op ^::nat => nat => nat)
-             ((number_of \<Colon> int => nat)
-               ((Int.Bit0 \<Colon> int => int)
-                 ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-             ((op -::nat => nat => nat) a b))
-           ((op ^::nat => nat => nat)
-             ((number_of \<Colon> int => nat)
-               ((Int.Bit0 \<Colon> int => int)
-                 ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-             a)))"
+lemma EXP_SUB_LESS_EQ: "(2::nat) ^ ((a::nat) - (b::nat)) <= (2::nat) ^ a"
   by (import bits EXP_SUB_LESS_EQ)
 
-lemma BITS_THM: "ALL (x::nat) (xa::nat) xb::nat.
-   BITS x xa xb = xb div 2 ^ xa mod 2 ^ (Suc x - xa)"
+lemma BITS_THM: "BITS x xa xb = xb div 2 ^ xa mod 2 ^ (Suc x - xa)"
   by (import bits BITS_THM)
 
-lemma BITSLT_THM: "ALL (h::nat) (l::nat) n::nat. BITS h l n < 2 ^ (Suc h - l)"
+lemma BITSLT_THM: "BITS h l n < 2 ^ (Suc h - l)"
   by (import bits BITSLT_THM)
 
-lemma DIV_MULT_LEM: "ALL (m::nat) n::nat. 0 < n --> m div n * n <= m"
+lemma DIV_MULT_LEM: "(0::nat) < (n::nat) ==> (m::nat) div n * n <= m"
   by (import bits DIV_MULT_LEM)
 
-lemma MOD_2EXP_LEM: "ALL (n::nat) x::nat. n mod 2 ^ x = n - n div 2 ^ x * 2 ^ x"
+lemma MOD_2EXP_LEM: "(n::nat) mod (2::nat) ^ (x::nat) = n - n div (2::nat) ^ x * (2::nat) ^ x"
   by (import bits MOD_2EXP_LEM)
 
-lemma BITS2_THM: "ALL (h::nat) (l::nat) n::nat. BITS h l n = n mod 2 ^ Suc h div 2 ^ l"
+lemma BITS2_THM: "BITS h l n = n mod 2 ^ Suc h div 2 ^ l"
   by (import bits BITS2_THM)
 
-lemma BITS_COMP_THM: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::nat.
-   h2 + l1 <= h1 --> BITS h2 l2 (BITS h1 l1 n) = BITS (h2 + l1) (l2 + l1) n"
+lemma BITS_COMP_THM: "h2 + l1 <= h1 ==> BITS h2 l2 (BITS h1 l1 n) = BITS (h2 + l1) (l2 + l1) n"
   by (import bits BITS_COMP_THM)
 
-lemma BITS_DIV_THM: "ALL (h::nat) (l::nat) (x::nat) n::nat.
-   BITS h l x div 2 ^ n = BITS h (l + n) x"
+lemma BITS_DIV_THM: "BITS h l x div 2 ^ n = BITS h (l + n) x"
   by (import bits BITS_DIV_THM)
 
-lemma BITS_LT_HIGH: "ALL (h::nat) (l::nat) n::nat. n < 2 ^ Suc h --> BITS h l n = n div 2 ^ l"
+lemma BITS_LT_HIGH: "n < 2 ^ Suc h ==> BITS h l n = n div 2 ^ l"
   by (import bits BITS_LT_HIGH)
 
-lemma BITS_ZERO: "ALL (h::nat) (l::nat) n::nat. h < l --> BITS h l n = 0"
+lemma BITS_ZERO: "h < l ==> BITS h l n = 0"
   by (import bits BITS_ZERO)
 
-lemma BITS_ZERO2: "ALL (h::nat) l::nat. BITS h l 0 = 0"
+lemma BITS_ZERO2: "BITS h l 0 = 0"
   by (import bits BITS_ZERO2)
 
-lemma BITS_ZERO3: "ALL (h::nat) x::nat. BITS h 0 x = x mod 2 ^ Suc h"
+lemma BITS_ZERO3: "BITS h 0 x = x mod 2 ^ Suc h"
   by (import bits BITS_ZERO3)
 
-lemma BITS_COMP_THM2: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::nat.
-   BITS h2 l2 (BITS h1 l1 n) = BITS (min h1 (h2 + l1)) (l2 + l1) n"
+lemma BITS_COMP_THM2: "BITS h2 l2 (BITS h1 l1 n) = BITS (min h1 (h2 + l1)) (l2 + l1) n"
   by (import bits BITS_COMP_THM2)
 
-lemma NOT_MOD2_LEM: "ALL n::nat. (n mod 2 ~= 0) = (n mod 2 = 1)"
+lemma NOT_MOD2_LEM: "((n::nat) mod (2::nat) ~= (0::nat)) = (n mod (2::nat) = (1::nat))"
   by (import bits NOT_MOD2_LEM)
 
-lemma NOT_MOD2_LEM2: "ALL (n::nat) a::'a::type. (n mod 2 ~= 1) = (n mod 2 = 0)"
+lemma NOT_MOD2_LEM2: "((n::nat) mod (2::nat) ~= (1::nat)) = (n mod (2::nat) = (0::nat))"
   by (import bits NOT_MOD2_LEM2)
 
-lemma EVEN_MOD2_LEM: "ALL n::nat. EVEN n = (n mod 2 = 0)"
+lemma EVEN_MOD2_LEM: "EVEN n = (n mod 2 = 0)"
   by (import bits EVEN_MOD2_LEM)
 
-lemma ODD_MOD2_LEM: "ALL n::nat. ODD n = (n mod 2 = 1)"
+lemma ODD_MOD2_LEM: "ODD n = (n mod 2 = 1)"
   by (import bits ODD_MOD2_LEM)
 
 lemma LSB_ODD: "LSBn = ODD"
   by (import bits LSB_ODD)
 
-lemma DIV_MULT_THM: "ALL (x::nat) n::nat. n div 2 ^ x * 2 ^ x = n - n mod 2 ^ x"
+lemma DIV_MULT_THM: "(n::nat) div (2::nat) ^ (x::nat) * (2::nat) ^ x = n - n mod (2::nat) ^ x"
   by (import bits DIV_MULT_THM)
 
-lemma DIV_MULT_THM2: "ALL x::nat. 2 * (x div 2) = x - x mod 2"
+lemma DIV_MULT_THM2: "(2::nat) * ((x::nat) div (2::nat)) = x - x mod (2::nat)"
   by (import bits DIV_MULT_THM2)
 
-lemma LESS_EQ_EXP_MULT: "ALL (a::nat) b::nat. a <= b --> (EX x::nat. 2 ^ b = x * 2 ^ a)"
+lemma LESS_EQ_EXP_MULT: "(a::nat) <= (b::nat) ==> EX x::nat. (2::nat) ^ b = x * (2::nat) ^ a"
   by (import bits LESS_EQ_EXP_MULT)
 
-lemma SLICE_LEM1: "ALL (a::nat) (x::nat) y::nat.
-   a div 2 ^ (x + y) * 2 ^ (x + y) =
-   a div 2 ^ x * 2 ^ x - a div 2 ^ x mod 2 ^ y * 2 ^ x"
+lemma SLICE_LEM1: "(a::nat) div (2::nat) ^ ((x::nat) + (y::nat)) * (2::nat) ^ (x + y) =
+a div (2::nat) ^ x * (2::nat) ^ x -
+a div (2::nat) ^ x mod (2::nat) ^ y * (2::nat) ^ x"
   by (import bits SLICE_LEM1)
 
-lemma SLICE_LEM2: "ALL (a::'a::type) (x::nat) y::nat.
-   (n::nat) mod 2 ^ (x + y) = n mod 2 ^ x + n div 2 ^ x mod 2 ^ y * 2 ^ x"
+lemma SLICE_LEM2: "(n::nat) mod (2::nat) ^ ((x::nat) + (y::nat)) =
+n mod (2::nat) ^ x + n div (2::nat) ^ x mod (2::nat) ^ y * (2::nat) ^ x"
   by (import bits SLICE_LEM2)
 
-lemma SLICE_LEM3: "ALL (n::nat) (h::nat) l::nat. l < h --> n mod 2 ^ Suc l <= n mod 2 ^ h"
+lemma SLICE_LEM3: "(l::nat) < (h::nat) ==> (n::nat) mod (2::nat) ^ Suc l <= n mod (2::nat) ^ h"
   by (import bits SLICE_LEM3)
 
-lemma SLICE_THM: "ALL (n::nat) (h::nat) l::nat. SLICE h l n = BITS h l n * 2 ^ l"
+lemma SLICE_THM: "SLICE h l n = BITS h l n * 2 ^ l"
   by (import bits SLICE_THM)
 
-lemma SLICELT_THM: "ALL (h::nat) (l::nat) n::nat. SLICE h l n < 2 ^ Suc h"
+lemma SLICELT_THM: "SLICE h l n < 2 ^ Suc h"
   by (import bits SLICELT_THM)
 
-lemma BITS_SLICE_THM: "ALL (h::nat) (l::nat) n::nat. BITS h l (SLICE h l n) = BITS h l n"
+lemma BITS_SLICE_THM: "BITS h l (SLICE h l n) = BITS h l n"
   by (import bits BITS_SLICE_THM)
 
-lemma BITS_SLICE_THM2: "ALL (h::nat) (l::nat) n::nat.
-   h <= (h2::nat) --> BITS h2 l (SLICE h l n) = BITS h l n"
+lemma BITS_SLICE_THM2: "h <= h2 ==> BITS h2 l (SLICE h l n) = BITS h l n"
   by (import bits BITS_SLICE_THM2)
 
-lemma MOD_2EXP_MONO: "ALL (n::nat) (h::nat) l::nat. l <= h --> n mod 2 ^ l <= n mod 2 ^ Suc h"
+lemma MOD_2EXP_MONO: "(l::nat) <= (h::nat) ==> (n::nat) mod (2::nat) ^ l <= n mod (2::nat) ^ Suc h"
   by (import bits MOD_2EXP_MONO)
 
-lemma SLICE_COMP_THM: "ALL (h::nat) (m::nat) (l::nat) n::nat.
-   Suc m <= h & l <= m --> SLICE h (Suc m) n + SLICE m l n = SLICE h l n"
+lemma SLICE_COMP_THM: "Suc m <= h & l <= m ==> SLICE h (Suc m) n + SLICE m l n = SLICE h l n"
   by (import bits SLICE_COMP_THM)
 
-lemma SLICE_ZERO: "ALL (h::nat) (l::nat) n::nat. h < l --> SLICE h l n = 0"
+lemma SLICE_ZERO: "h < l ==> SLICE h l n = 0"
   by (import bits SLICE_ZERO)
 
-lemma BIT_COMP_THM3: "ALL (h::nat) (m::nat) (l::nat) n::nat.
-   Suc m <= h & l <= m -->
-   BITS h (Suc m) n * 2 ^ (Suc m - l) + BITS m l n = BITS h l n"
+lemma BIT_COMP_THM3: "Suc m <= h & l <= m
+==> BITS h (Suc m) n * 2 ^ (Suc m - l) + BITS m l n = BITS h l n"
   by (import bits BIT_COMP_THM3)
 
-lemma NOT_BIT: "ALL (n::nat) a::nat. (~ bit n a) = (BITS n n a = 0)"
+lemma NOT_BIT: "(~ bit n a) = (BITS n n a = 0)"
   by (import bits NOT_BIT)
 
-lemma NOT_BITS: "ALL (n::nat) a::nat. (BITS n n a ~= 0) = (BITS n n a = 1)"
+lemma NOT_BITS: "(BITS n n a ~= 0) = (BITS n n a = 1)"
   by (import bits NOT_BITS)
 
-lemma NOT_BITS2: "ALL (n::nat) a::nat. (BITS n n a ~= 1) = (BITS n n a = 0)"
+lemma NOT_BITS2: "(BITS n n a ~= 1) = (BITS n n a = 0)"
   by (import bits NOT_BITS2)
 
-lemma BIT_SLICE: "ALL (n::nat) (a::nat) b::nat.
-   (bit n a = bit n b) = (SLICE n n a = SLICE n n b)"
+lemma BIT_SLICE: "(bit n a = bit n b) = (SLICE n n a = SLICE n n b)"
   by (import bits BIT_SLICE)
 
-lemma BIT_SLICE_LEM: "ALL (y::nat) (x::nat) n::nat. SBIT (bit x n) (x + y) = SLICE x x n * 2 ^ y"
+lemma BIT_SLICE_LEM: "SBIT (bit x n) (x + y) = SLICE x x n * 2 ^ y"
   by (import bits BIT_SLICE_LEM)
 
-lemma BIT_SLICE_THM: "ALL (x::nat) xa::nat. SBIT (bit x xa) x = SLICE x x xa"
+lemma BIT_SLICE_THM: "SBIT (bit x xa) x = SLICE x x xa"
   by (import bits BIT_SLICE_THM)
 
-lemma SBIT_DIV: "ALL (b::bool) (m::nat) n::nat. n < m --> SBIT b (m - n) = SBIT b m div 2 ^ n"
+lemma SBIT_DIV: "n < m ==> SBIT b (m - n) = SBIT b m div 2 ^ n"
   by (import bits SBIT_DIV)
 
-lemma BITS_SUC: "ALL (h::nat) (l::nat) n::nat.
-   l <= Suc h -->
-   SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n = BITS (Suc h) l n"
+lemma BITS_SUC: "l <= Suc h
+==> SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n = BITS (Suc h) l n"
   by (import bits BITS_SUC)
 
-lemma BITS_SUC_THM: "ALL (h::nat) (l::nat) n::nat.
-   BITS (Suc h) l n =
-   (if Suc h < l then 0 else SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n)"
+lemma BITS_SUC_THM: "BITS (Suc h) l n =
+(if Suc h < l then 0 else SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n)"
   by (import bits BITS_SUC_THM)
 
-lemma BIT_BITS_THM: "ALL (h::nat) (l::nat) (a::nat) b::nat.
-   (ALL x::nat. l <= x & x <= h --> bit x a = bit x b) =
-   (BITS h l a = BITS h l b)"
+lemma BIT_BITS_THM: "(ALL x. l <= x & x <= h --> bit x a = bit x b) = (BITS h l a = BITS h l b)"
   by (import bits BIT_BITS_THM)
 
-lemma BITWISE_LT_2EXP: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   BITWISE n oper a b < 2 ^ n"
+lemma BITWISE_LT_2EXP: "BITWISE n oper a b < 2 ^ n"
   by (import bits BITWISE_LT_2EXP)
 
-lemma LESS_EXP_MULT2: "(All::(nat => bool) => bool)
- (%a::nat.
-     (All::(nat => bool) => bool)
-      (%b::nat.
-          (op -->::bool => bool => bool) ((op <::nat => nat => bool) a b)
-           ((Ex::(nat => bool) => bool)
-             (%x::nat.
-                 (op =::nat => nat => bool)
-                  ((op ^::nat => nat => nat)
-                    ((number_of \<Colon> int => nat)
-                      ((Int.Bit0 \<Colon> int => int)
-                        ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-                    b)
-                  ((op *::nat => nat => nat)
-                    ((op ^::nat => nat => nat)
-                      ((number_of \<Colon> int => nat)
-                        ((Int.Bit0 \<Colon> int => int)
-                          ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-                      ((op +::nat => nat => nat) x (1::nat)))
-                    ((op ^::nat => nat => nat)
-                      ((number_of \<Colon> int => nat)
-                        ((Int.Bit0 \<Colon> int => int)
-                          ((Int.Bit1 \<Colon> int => int) (Int.Pls \<Colon> int))))
-                      a))))))"
+lemma LESS_EXP_MULT2: "(a::nat) < (b::nat)
+==> EX x::nat. (2::nat) ^ b = (2::nat) ^ (x + (1::nat)) * (2::nat) ^ a"
   by (import bits LESS_EXP_MULT2)
 
-lemma BITWISE_THM: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   x < n --> bit x (BITWISE n oper a b) = oper (bit x a) (bit x b)"
+lemma BITWISE_THM: "x < n ==> bit x (BITWISE n oper a b) = oper (bit x a) (bit x b)"
   by (import bits BITWISE_THM)
 
-lemma BITWISE_COR: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   x < n -->
-   oper (bit x a) (bit x b) --> BITWISE n oper a b div 2 ^ x mod 2 = 1"
+lemma BITWISE_COR: "[| x < n; oper (bit x a) (bit x b) |]
+==> BITWISE n oper a b div 2 ^ x mod 2 = 1"
   by (import bits BITWISE_COR)
 
-lemma BITWISE_NOT_COR: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   x < n -->
-   ~ oper (bit x a) (bit x b) --> BITWISE n oper a b div 2 ^ x mod 2 = 0"
+lemma BITWISE_NOT_COR: "[| x < n; ~ oper (bit x a) (bit x b) |]
+==> BITWISE n oper a b div 2 ^ x mod 2 = 0"
   by (import bits BITWISE_NOT_COR)
 
-lemma MOD_PLUS_RIGHT: "ALL n>0. ALL (j::nat) k::nat. (j + k mod n) mod n = (j + k) mod n"
+lemma MOD_PLUS_RIGHT: "(0::nat) < (n::nat) ==> ((j::nat) + (k::nat) mod n) mod n = (j + k) mod n"
   by (import bits MOD_PLUS_RIGHT)
 
-lemma MOD_PLUS_1: "ALL n>0. ALL x::nat. ((x + 1) mod n = 0) = (x mod n + 1 = n)"
+lemma MOD_PLUS_1: "(0::nat) < (n::nat)
+==> (((x::nat) + (1::nat)) mod n = (0::nat)) = (x mod n + (1::nat) = n)"
   by (import bits MOD_PLUS_1)
 
-lemma MOD_ADD_1: "ALL n>0. ALL x::nat. (x + 1) mod n ~= 0 --> (x + 1) mod n = x mod n + 1"
+lemma MOD_ADD_1: "[| (0::nat) < (n::nat); ((x::nat) + (1::nat)) mod n ~= (0::nat) |]
+==> (x + (1::nat)) mod n = x mod n + (1::nat)"
   by (import bits MOD_ADD_1)
 
 ;end_setup
@@ -406,63 +315,57 @@
   MODw :: "nat => nat" 
 
 defs
-  MODw_primdef: "MODw == %n::nat. n mod 2 ^ WL"
+  MODw_primdef: "MODw == %n. n mod 2 ^ WL"
 
-lemma MODw_def: "ALL n::nat. MODw n = n mod 2 ^ WL"
+lemma MODw_def: "MODw n = n mod 2 ^ WL"
   by (import word32 MODw_def)
 
 consts
   INw :: "nat => bool" 
 
 defs
-  INw_primdef: "INw == %n::nat. n < 2 ^ WL"
+  INw_primdef: "INw == %n. n < 2 ^ WL"
 
-lemma INw_def: "ALL n::nat. INw n = (n < 2 ^ WL)"
+lemma INw_def: "INw n = (n < 2 ^ WL)"
   by (import word32 INw_def)
 
 consts
   EQUIV :: "nat => nat => bool" 
 
 defs
-  EQUIV_primdef: "EQUIV == %(x::nat) y::nat. MODw x = MODw y"
+  EQUIV_primdef: "EQUIV == %x y. MODw x = MODw y"
 
-lemma EQUIV_def: "ALL (x::nat) y::nat. EQUIV x y = (MODw x = MODw y)"
+lemma EQUIV_def: "EQUIV x y = (MODw x = MODw y)"
   by (import word32 EQUIV_def)
 
-lemma EQUIV_QT: "ALL (x::nat) y::nat. EQUIV x y = (EQUIV x = EQUIV y)"
+lemma EQUIV_QT: "EQUIV x y = (EQUIV x = EQUIV y)"
   by (import word32 EQUIV_QT)
 
-lemma FUNPOW_THM: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type.
-   (f ^^ n) (f x) = f ((f ^^ n) x)"
-  by (import word32 FUNPOW_THM)
-
-lemma FUNPOW_THM2: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type.
-   (f ^^ Suc n) x = f ((f ^^ n) x)"
+lemma FUNPOW_THM2: "(f ^^ Suc n) x = f ((f ^^ n) x)"
   by (import word32 FUNPOW_THM2)
 
-lemma FUNPOW_COMP: "ALL (f::'a::type => 'a::type) (m::nat) (n::nat) a::'a::type.
-   (f ^^ m) ((f ^^ n) a) = (f ^^ (m + n)) a"
+lemma FUNPOW_COMP: "(f ^^ m) ((f ^^ n) a) = (f ^^ (m + n)) a"
   by (import word32 FUNPOW_COMP)
 
-lemma INw_MODw: "ALL n::nat. INw (MODw n)"
+lemma INw_MODw: "INw (MODw n)"
   by (import word32 INw_MODw)
 
-lemma TOw_IDEM: "ALL a::nat. INw a --> MODw a = a"
+lemma TOw_IDEM: "INw a ==> MODw a = a"
   by (import word32 TOw_IDEM)
 
-lemma MODw_IDEM2: "ALL a::nat. MODw (MODw a) = MODw a"
+lemma MODw_IDEM2: "MODw (MODw a) = MODw a"
   by (import word32 MODw_IDEM2)
 
-lemma TOw_QT: "ALL a::nat. EQUIV (MODw a) a"
+lemma TOw_QT: "EQUIV (MODw a) a"
   by (import word32 TOw_QT)
 
 lemma MODw_THM: "MODw = BITS HB 0"
   by (import word32 MODw_THM)
 
-lemma MOD_ADD: "ALL (a::nat) b::nat. MODw (a + b) = MODw (MODw a + MODw b)"
+lemma MOD_ADD: "MODw (a + b) = MODw (MODw a + MODw b)"
   by (import word32 MOD_ADD)
 
-lemma MODw_MULT: "ALL (a::nat) b::nat. MODw (a * b) = MODw (MODw a * MODw b)"
+lemma MODw_MULT: "MODw (a * b) = MODw (MODw a * MODw b)"
   by (import word32 MODw_MULT)
 
 consts
@@ -474,65 +377,62 @@
 lemma AONE_def: "AONE = 1"
   by (import word32 AONE_def)
 
-lemma ADD_QT: "(ALL n::nat. EQUIV (0 + n) n) &
-(ALL (m::nat) n::nat. EQUIV (Suc m + n) (Suc (m + n)))"
+lemma ADD_QT: "(ALL n. EQUIV (0 + n) n) & (ALL m n. EQUIV (Suc m + n) (Suc (m + n)))"
   by (import word32 ADD_QT)
 
-lemma ADD_0_QT: "ALL a::nat. EQUIV (a + 0) a"
+lemma ADD_0_QT: "EQUIV (a + 0) a"
   by (import word32 ADD_0_QT)
 
-lemma ADD_COMM_QT: "ALL (a::nat) b::nat. EQUIV (a + b) (b + a)"
+lemma ADD_COMM_QT: "EQUIV (a + b) (b + a)"
   by (import word32 ADD_COMM_QT)
 
-lemma ADD_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (a + (b + c)) (a + b + c)"
+lemma ADD_ASSOC_QT: "EQUIV (a + (b + c)) (a + b + c)"
   by (import word32 ADD_ASSOC_QT)
 
-lemma MULT_QT: "(ALL n::nat. EQUIV (0 * n) 0) &
-(ALL (m::nat) n::nat. EQUIV (Suc m * n) (m * n + n))"
+lemma MULT_QT: "(ALL n. EQUIV (0 * n) 0) & (ALL m n. EQUIV (Suc m * n) (m * n + n))"
   by (import word32 MULT_QT)
 
-lemma ADD1_QT: "ALL m::nat. EQUIV (Suc m) (m + AONE)"
+lemma ADD1_QT: "EQUIV (Suc m) (m + AONE)"
   by (import word32 ADD1_QT)
 
-lemma ADD_CLAUSES_QT: "(ALL m::nat. EQUIV (0 + m) m) &
-(ALL m::nat. EQUIV (m + 0) m) &
-(ALL (m::nat) n::nat. EQUIV (Suc m + n) (Suc (m + n))) &
-(ALL (m::nat) n::nat. EQUIV (m + Suc n) (Suc (m + n)))"
+lemma ADD_CLAUSES_QT: "(ALL m. EQUIV (0 + m) m) &
+(ALL m. EQUIV (m + 0) m) &
+(ALL m n. EQUIV (Suc m + n) (Suc (m + n))) &
+(ALL m n. EQUIV (m + Suc n) (Suc (m + n)))"
   by (import word32 ADD_CLAUSES_QT)
 
-lemma SUC_EQUIV_COMP: "ALL (a::nat) b::nat. EQUIV (Suc a) b --> EQUIV a (b + (2 ^ WL - 1))"
+lemma SUC_EQUIV_COMP: "EQUIV (Suc a) b ==> EQUIV a (b + (2 ^ WL - 1))"
   by (import word32 SUC_EQUIV_COMP)
 
-lemma INV_SUC_EQ_QT: "ALL (m::nat) n::nat. EQUIV (Suc m) (Suc n) = EQUIV m n"
+lemma INV_SUC_EQ_QT: "EQUIV (Suc m) (Suc n) = EQUIV m n"
   by (import word32 INV_SUC_EQ_QT)
 
-lemma ADD_INV_0_QT: "ALL (m::nat) n::nat. EQUIV (m + n) m --> EQUIV n 0"
+lemma ADD_INV_0_QT: "EQUIV (m + n) m ==> EQUIV n 0"
   by (import word32 ADD_INV_0_QT)
 
-lemma ADD_INV_0_EQ_QT: "ALL (m::nat) n::nat. EQUIV (m + n) m = EQUIV n 0"
+lemma ADD_INV_0_EQ_QT: "EQUIV (m + n) m = EQUIV n 0"
   by (import word32 ADD_INV_0_EQ_QT)
 
-lemma EQ_ADD_LCANCEL_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (m + n) (m + p) = EQUIV n p"
+lemma EQ_ADD_LCANCEL_QT: "EQUIV (m + n) (m + p) = EQUIV n p"
   by (import word32 EQ_ADD_LCANCEL_QT)
 
-lemma EQ_ADD_RCANCEL_QT: "ALL (x::nat) (xa::nat) xb::nat. EQUIV (x + xb) (xa + xb) = EQUIV x xa"
+lemma EQ_ADD_RCANCEL_QT: "EQUIV (x + xb) (xa + xb) = EQUIV x xa"
   by (import word32 EQ_ADD_RCANCEL_QT)
 
-lemma LEFT_ADD_DISTRIB_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (p * (m + n)) (p * m + p * n)"
+lemma LEFT_ADD_DISTRIB_QT: "EQUIV (p * (m + n)) (p * m + p * n)"
   by (import word32 LEFT_ADD_DISTRIB_QT)
 
-lemma MULT_ASSOC_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (m * (n * p)) (m * n * p)"
+lemma MULT_ASSOC_QT: "EQUIV (m * (n * p)) (m * n * p)"
   by (import word32 MULT_ASSOC_QT)
 
-lemma MULT_COMM_QT: "ALL (m::nat) n::nat. EQUIV (m * n) (n * m)"
+lemma MULT_COMM_QT: "EQUIV (m * n) (n * m)"
   by (import word32 MULT_COMM_QT)
 
-lemma MULT_CLAUSES_QT: "ALL (m::nat) n::nat.
-   EQUIV (0 * m) 0 &
-   EQUIV (m * 0) 0 &
-   EQUIV (AONE * m) m &
-   EQUIV (m * AONE) m &
-   EQUIV (Suc m * n) (m * n + n) & EQUIV (m * Suc n) (m + m * n)"
+lemma MULT_CLAUSES_QT: "EQUIV (0 * m) 0 &
+EQUIV (m * 0) 0 &
+EQUIV (AONE * m) m &
+EQUIV (m * AONE) m &
+EQUIV (Suc m * n) (m * n + n) & EQUIV (m * Suc n) (m + m * n)"
   by (import word32 MULT_CLAUSES_QT)
 
 consts
@@ -548,48 +448,36 @@
   ONE_COMP :: "nat => nat" 
 
 defs
-  ONE_COMP_primdef: "ONE_COMP == %x::nat. 2 ^ WL - 1 - MODw x"
+  ONE_COMP_primdef: "ONE_COMP == %x. 2 ^ WL - 1 - MODw x"
 
-lemma ONE_COMP_def: "ALL x::nat. ONE_COMP x = 2 ^ WL - 1 - MODw x"
+lemma ONE_COMP_def: "ONE_COMP x = 2 ^ WL - 1 - MODw x"
   by (import word32 ONE_COMP_def)
 
 consts
   TWO_COMP :: "nat => nat" 
 
 defs
-  TWO_COMP_primdef: "TWO_COMP == %x::nat. 2 ^ WL - MODw x"
+  TWO_COMP_primdef: "TWO_COMP == %x. 2 ^ WL - MODw x"
 
-lemma TWO_COMP_def: "ALL x::nat. TWO_COMP x = 2 ^ WL - MODw x"
+lemma TWO_COMP_def: "TWO_COMP x = 2 ^ WL - MODw x"
   by (import word32 TWO_COMP_def)
 
-lemma ADD_TWO_COMP_QT: "ALL a::nat. EQUIV (MODw a + TWO_COMP a) 0"
+lemma ADD_TWO_COMP_QT: "EQUIV (MODw a + TWO_COMP a) 0"
   by (import word32 ADD_TWO_COMP_QT)
 
-lemma TWO_COMP_ONE_COMP_QT: "ALL a::nat. EQUIV (TWO_COMP a) (ONE_COMP a + AONE)"
+lemma TWO_COMP_ONE_COMP_QT: "EQUIV (TWO_COMP a) (ONE_COMP a + AONE)"
   by (import word32 TWO_COMP_ONE_COMP_QT)
 
-lemma BIT_EQUIV_THM: "(All::(nat => bool) => bool)
- (%x::nat.
-     (All::(nat => bool) => bool)
-      (%xa::nat.
-          (op =::bool => bool => bool)
-           ((All::(nat => bool) => bool)
-             (%xb::nat.
-                 (op -->::bool => bool => bool)
-                  ((op <::nat => nat => bool) xb (WL::nat))
-                  ((op =::bool => bool => bool)
-                    ((bit::nat => nat => bool) xb x)
-                    ((bit::nat => nat => bool) xb xa))))
-           ((EQUIV::nat => nat => bool) x xa)))"
+lemma BIT_EQUIV_THM: "(ALL xb<WL. bit xb x = bit xb xa) = EQUIV x xa"
   by (import word32 BIT_EQUIV_THM)
 
-lemma BITS_SUC2: "ALL (n::nat) a::nat. BITS (Suc n) 0 a = SLICE (Suc n) (Suc n) a + BITS n 0 a"
+lemma BITS_SUC2: "BITS (Suc n) 0 a = SLICE (Suc n) (Suc n) a + BITS n 0 a"
   by (import word32 BITS_SUC2)
 
-lemma BITWISE_ONE_COMP_THM: "ALL (a::nat) b::nat. BITWISE WL (%(x::bool) y::bool. ~ x) a b = ONE_COMP a"
+lemma BITWISE_ONE_COMP_THM: "BITWISE WL (%x y. ~ x) a b = ONE_COMP a"
   by (import word32 BITWISE_ONE_COMP_THM)
 
-lemma ONE_COMP_THM: "ALL (x::nat) xa::nat. xa < WL --> bit xa (ONE_COMP x) = (~ bit xa x)"
+lemma ONE_COMP_THM: "xa < WL ==> bit xa (ONE_COMP x) = (~ bit xa x)"
   by (import word32 ONE_COMP_THM)
 
 consts
@@ -614,9 +502,9 @@
   EOR :: "nat => nat => nat" 
 
 defs
-  EOR_primdef: "EOR == BITWISE WL (%(x::bool) y::bool. x ~= y)"
+  EOR_primdef: "EOR == BITWISE WL op ~="
 
-lemma EOR_def: "EOR = BITWISE WL (%(x::bool) y::bool. x ~= y)"
+lemma EOR_def: "EOR = BITWISE WL op ~="
   by (import word32 EOR_def)
 
 consts
@@ -628,177 +516,147 @@
 lemma COMP0_def: "COMP0 = ONE_COMP 0"
   by (import word32 COMP0_def)
 
-lemma BITWISE_THM2: "(All::(nat => bool) => bool)
- (%y::nat.
-     (All::((bool => bool => bool) => bool) => bool)
-      (%oper::bool => bool => bool.
-          (All::(nat => bool) => bool)
-           (%a::nat.
-               (All::(nat => bool) => bool)
-                (%b::nat.
-                    (op =::bool => bool => bool)
-                     ((All::(nat => bool) => bool)
-                       (%x::nat.
-                           (op -->::bool => bool => bool)
-                            ((op <::nat => nat => bool) x (WL::nat))
-                            ((op =::bool => bool => bool)
-                              (oper ((bit::nat => nat => bool) x a)
-                                ((bit::nat => nat => bool) x b))
-                              ((bit::nat => nat => bool) x y))))
-                     ((EQUIV::nat => nat => bool)
-                       ((BITWISE::nat
-                                  => (bool => bool => bool)
-                                     => nat => nat => nat)
-                         (WL::nat) oper a b)
-                       y)))))"
+lemma BITWISE_THM2: "(ALL x<WL. oper (bit x a) (bit x b) = bit x y) =
+EQUIV (BITWISE WL oper a b) y"
   by (import word32 BITWISE_THM2)
 
-lemma OR_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (OR a (OR b c)) (OR (OR a b) c)"
+lemma OR_ASSOC_QT: "EQUIV (OR a (OR b c)) (OR (OR a b) c)"
   by (import word32 OR_ASSOC_QT)
 
-lemma OR_COMM_QT: "ALL (a::nat) b::nat. EQUIV (OR a b) (OR b a)"
+lemma OR_COMM_QT: "EQUIV (OR a b) (OR b a)"
   by (import word32 OR_COMM_QT)
 
-lemma OR_ABSORB_QT: "ALL (a::nat) b::nat. EQUIV (AND a (OR a b)) a"
+lemma OR_ABSORB_QT: "EQUIV (AND a (OR a b)) a"
   by (import word32 OR_ABSORB_QT)
 
-lemma OR_IDEM_QT: "ALL a::nat. EQUIV (OR a a) a"
+lemma OR_IDEM_QT: "EQUIV (OR a a) a"
   by (import word32 OR_IDEM_QT)
 
-lemma AND_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (AND a (AND b c)) (AND (AND a b) c)"
+lemma AND_ASSOC_QT: "EQUIV (AND a (AND b c)) (AND (AND a b) c)"
   by (import word32 AND_ASSOC_QT)
 
-lemma AND_COMM_QT: "ALL (a::nat) b::nat. EQUIV (AND a b) (AND b a)"
+lemma AND_COMM_QT: "EQUIV (AND a b) (AND b a)"
   by (import word32 AND_COMM_QT)
 
-lemma AND_ABSORB_QT: "ALL (a::nat) b::nat. EQUIV (OR a (AND a b)) a"
+lemma AND_ABSORB_QT: "EQUIV (OR a (AND a b)) a"
   by (import word32 AND_ABSORB_QT)
 
-lemma AND_IDEM_QT: "ALL a::nat. EQUIV (AND a a) a"
+lemma AND_IDEM_QT: "EQUIV (AND a a) a"
   by (import word32 AND_IDEM_QT)
 
-lemma OR_COMP_QT: "ALL a::nat. EQUIV (OR a (ONE_COMP a)) COMP0"
+lemma OR_COMP_QT: "EQUIV (OR a (ONE_COMP a)) COMP0"
   by (import word32 OR_COMP_QT)
 
-lemma AND_COMP_QT: "ALL a::nat. EQUIV (AND a (ONE_COMP a)) 0"
+lemma AND_COMP_QT: "EQUIV (AND a (ONE_COMP a)) 0"
   by (import word32 AND_COMP_QT)
 
-lemma ONE_COMP_QT: "ALL a::nat. EQUIV (ONE_COMP (ONE_COMP a)) a"
+lemma ONE_COMP_QT: "EQUIV (ONE_COMP (ONE_COMP a)) a"
   by (import word32 ONE_COMP_QT)
 
-lemma RIGHT_AND_OVER_OR_QT: "ALL (a::nat) (b::nat) c::nat.
-   EQUIV (AND (OR a b) c) (OR (AND a c) (AND b c))"
+lemma RIGHT_AND_OVER_OR_QT: "EQUIV (AND (OR a b) c) (OR (AND a c) (AND b c))"
   by (import word32 RIGHT_AND_OVER_OR_QT)
 
-lemma RIGHT_OR_OVER_AND_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (OR (AND a b) c) (AND (OR a c) (OR b c))"
+lemma RIGHT_OR_OVER_AND_QT: "EQUIV (OR (AND a b) c) (AND (OR a c) (OR b c))"
   by (import word32 RIGHT_OR_OVER_AND_QT)
 
-lemma DE_MORGAN_THM_QT: "ALL (a::nat) b::nat.
-   EQUIV (ONE_COMP (AND a b)) (OR (ONE_COMP a) (ONE_COMP b)) &
-   EQUIV (ONE_COMP (OR a b)) (AND (ONE_COMP a) (ONE_COMP b))"
+lemma DE_MORGAN_THM_QT: "EQUIV (ONE_COMP (AND a b)) (OR (ONE_COMP a) (ONE_COMP b)) &
+EQUIV (ONE_COMP (OR a b)) (AND (ONE_COMP a) (ONE_COMP b))"
   by (import word32 DE_MORGAN_THM_QT)
 
-lemma BIT_EQUIV: "ALL (n::nat) (a::nat) b::nat. n < WL --> EQUIV a b --> bit n a = bit n b"
+lemma BIT_EQUIV: "[| n < WL; EQUIV a b |] ==> bit n a = bit n b"
   by (import word32 BIT_EQUIV)
 
-lemma LSB_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> LSBn a = LSBn b"
+lemma LSB_WELLDEF: "EQUIV a b ==> LSBn a = LSBn b"
   by (import word32 LSB_WELLDEF)
 
-lemma MSB_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> MSBn a = MSBn b"
+lemma MSB_WELLDEF: "EQUIV a b ==> MSBn a = MSBn b"
   by (import word32 MSB_WELLDEF)
 
-lemma BITWISE_ISTEP: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   0 < n -->
-   BITWISE n oper (a div 2) (b div 2) =
-   BITWISE n oper a b div 2 + SBIT (oper (bit n a) (bit n b)) (n - 1)"
+lemma BITWISE_ISTEP: "0 < n
+==> BITWISE n oper (a div 2) (b div 2) =
+    BITWISE n oper a b div 2 + SBIT (oper (bit n a) (bit n b)) (n - 1)"
   by (import word32 BITWISE_ISTEP)
 
-lemma BITWISE_EVAL: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat.
-   BITWISE (Suc n) oper a b =
-   2 * BITWISE n oper (a div 2) (b div 2) + SBIT (oper (LSBn a) (LSBn b)) 0"
+lemma BITWISE_EVAL: "BITWISE (Suc n) oper a b =
+2 * BITWISE n oper (a div 2) (b div 2) + SBIT (oper (LSBn a) (LSBn b)) 0"
   by (import word32 BITWISE_EVAL)
 
-lemma BITWISE_WELLDEF: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) (b::nat) (c::nat) d::nat.
-   EQUIV a b & EQUIV c d --> EQUIV (BITWISE n oper a c) (BITWISE n oper b d)"
+lemma BITWISE_WELLDEF: "EQUIV a b & EQUIV c d ==> EQUIV (BITWISE n oper a c) (BITWISE n oper b d)"
   by (import word32 BITWISE_WELLDEF)
 
-lemma BITWISEw_WELLDEF: "ALL (oper::bool => bool => bool) (a::nat) (b::nat) (c::nat) d::nat.
-   EQUIV a b & EQUIV c d -->
-   EQUIV (BITWISE WL oper a c) (BITWISE WL oper b d)"
+lemma BITWISEw_WELLDEF: "EQUIV a b & EQUIV c d ==> EQUIV (BITWISE WL oper a c) (BITWISE WL oper b d)"
   by (import word32 BITWISEw_WELLDEF)
 
-lemma SUC_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (Suc a) (Suc b)"
+lemma SUC_WELLDEF: "EQUIV a b ==> EQUIV (Suc a) (Suc b)"
   by (import word32 SUC_WELLDEF)
 
-lemma ADD_WELLDEF: "ALL (a::nat) (b::nat) (c::nat) d::nat.
-   EQUIV a b & EQUIV c d --> EQUIV (a + c) (b + d)"
+lemma ADD_WELLDEF: "EQUIV a b & EQUIV c d ==> EQUIV (a + c) (b + d)"
   by (import word32 ADD_WELLDEF)
 
-lemma MUL_WELLDEF: "ALL (a::nat) (b::nat) (c::nat) d::nat.
-   EQUIV a b & EQUIV c d --> EQUIV (a * c) (b * d)"
+lemma MUL_WELLDEF: "EQUIV a b & EQUIV c d ==> EQUIV (a * c) (b * d)"
   by (import word32 MUL_WELLDEF)
 
-lemma ONE_COMP_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ONE_COMP a) (ONE_COMP b)"
+lemma ONE_COMP_WELLDEF: "EQUIV a b ==> EQUIV (ONE_COMP a) (ONE_COMP b)"
   by (import word32 ONE_COMP_WELLDEF)
 
-lemma TWO_COMP_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (TWO_COMP a) (TWO_COMP b)"
+lemma TWO_COMP_WELLDEF: "EQUIV a b ==> EQUIV (TWO_COMP a) (TWO_COMP b)"
   by (import word32 TWO_COMP_WELLDEF)
 
-lemma TOw_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (MODw a) (MODw b)"
+lemma TOw_WELLDEF: "EQUIV a b ==> EQUIV (MODw a) (MODw b)"
   by (import word32 TOw_WELLDEF)
 
 consts
   LSR_ONE :: "nat => nat" 
 
 defs
-  LSR_ONE_primdef: "LSR_ONE == %a::nat. MODw a div 2"
+  LSR_ONE_primdef: "LSR_ONE == %a. MODw a div 2"
 
-lemma LSR_ONE_def: "ALL a::nat. LSR_ONE a = MODw a div 2"
+lemma LSR_ONE_def: "LSR_ONE a = MODw a div 2"
   by (import word32 LSR_ONE_def)
 
 consts
   ASR_ONE :: "nat => nat" 
 
 defs
-  ASR_ONE_primdef: "ASR_ONE == %a::nat. LSR_ONE a + SBIT (MSBn a) HB"
+  ASR_ONE_primdef: "ASR_ONE == %a. LSR_ONE a + SBIT (MSBn a) HB"
 
-lemma ASR_ONE_def: "ALL a::nat. ASR_ONE a = LSR_ONE a + SBIT (MSBn a) HB"
+lemma ASR_ONE_def: "ASR_ONE a = LSR_ONE a + SBIT (MSBn a) HB"
   by (import word32 ASR_ONE_def)
 
 consts
   ROR_ONE :: "nat => nat" 
 
 defs
-  ROR_ONE_primdef: "ROR_ONE == %a::nat. LSR_ONE a + SBIT (LSBn a) HB"
+  ROR_ONE_primdef: "ROR_ONE == %a. LSR_ONE a + SBIT (LSBn a) HB"
 
-lemma ROR_ONE_def: "ALL a::nat. ROR_ONE a = LSR_ONE a + SBIT (LSBn a) HB"
+lemma ROR_ONE_def: "ROR_ONE a = LSR_ONE a + SBIT (LSBn a) HB"
   by (import word32 ROR_ONE_def)
 
 consts
   RRXn :: "bool => nat => nat" 
 
 defs
-  RRXn_primdef: "RRXn == %(c::bool) a::nat. LSR_ONE a + SBIT c HB"
+  RRXn_primdef: "RRXn == %c a. LSR_ONE a + SBIT c HB"
 
-lemma RRXn_def: "ALL (c::bool) a::nat. RRXn c a = LSR_ONE a + SBIT c HB"
+lemma RRXn_def: "RRXn c a = LSR_ONE a + SBIT c HB"
   by (import word32 RRXn_def)
 
-lemma LSR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (LSR_ONE a) (LSR_ONE b)"
+lemma LSR_ONE_WELLDEF: "EQUIV a b ==> EQUIV (LSR_ONE a) (LSR_ONE b)"
   by (import word32 LSR_ONE_WELLDEF)
 
-lemma ASR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ASR_ONE a) (ASR_ONE b)"
+lemma ASR_ONE_WELLDEF: "EQUIV a b ==> EQUIV (ASR_ONE a) (ASR_ONE b)"
   by (import word32 ASR_ONE_WELLDEF)
 
-lemma ROR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ROR_ONE a) (ROR_ONE b)"
+lemma ROR_ONE_WELLDEF: "EQUIV a b ==> EQUIV (ROR_ONE a) (ROR_ONE b)"
   by (import word32 ROR_ONE_WELLDEF)
 
-lemma RRX_WELLDEF: "ALL (a::nat) (b::nat) c::bool. EQUIV a b --> EQUIV (RRXn c a) (RRXn c b)"
+lemma RRX_WELLDEF: "EQUIV a b ==> EQUIV (RRXn c a) (RRXn c b)"
   by (import word32 RRX_WELLDEF)
 
 lemma LSR_ONE: "LSR_ONE = BITS HB 1"
   by (import word32 LSR_ONE)
 
-typedef (open) word32 = "{x::nat => bool. EX xa::nat. x = EQUIV xa}" 
+typedef (open) word32 = "{x::nat => bool. EX xa. x = EQUIV xa}" 
   by (rule typedef_helper,import word32 word32_TY_DEF)
 
 lemmas word32_TY_DEF = typedef_hol2hol4 [OF type_definition_word32]
@@ -807,9 +665,8 @@
   mk_word32 :: "(nat => bool) => word32" 
   dest_word32 :: "word32 => nat => bool" 
 
-specification (dest_word32 mk_word32) word32_tybij: "(ALL a::word32. mk_word32 (dest_word32 a) = a) &
-(ALL r::nat => bool.
-    (EX x::nat. r = EQUIV x) = (dest_word32 (mk_word32 r) = r))"
+specification (dest_word32 mk_word32) word32_tybij: "(ALL a. mk_word32 (dest_word32 a) = a) &
+(ALL r. (EX x. r = EQUIV x) = (dest_word32 (mk_word32 r) = r))"
   by (import word32 word32_tybij)
 
 consts
@@ -839,258 +696,237 @@
 lemma w_T_def: "w_T = mk_word32 (EQUIV COMP0)"
   by (import word32 w_T_def)
 
-definition word_suc :: "word32 => word32" where 
-  "word_suc == %T1::word32. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
+definition
+  word_suc :: "word32 => word32"  where
+  "word_suc == %T1. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
 
-lemma word_suc: "ALL T1::word32. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
+lemma word_suc: "word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
   by (import word32 word_suc)
 
-definition word_add :: "word32 => word32 => word32" where 
+definition
+  word_add :: "word32 => word32 => word32"  where
   "word_add ==
-%(T1::word32) T2::word32.
-   mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
+%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
 
-lemma word_add: "ALL (T1::word32) T2::word32.
-   word_add T1 T2 =
-   mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
+lemma word_add: "word_add T1 T2 =
+mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
   by (import word32 word_add)
 
-definition word_mul :: "word32 => word32 => word32" where 
+definition
+  word_mul :: "word32 => word32 => word32"  where
   "word_mul ==
-%(T1::word32) T2::word32.
-   mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
+%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
 
-lemma word_mul: "ALL (T1::word32) T2::word32.
-   word_mul T1 T2 =
-   mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
+lemma word_mul: "word_mul T1 T2 =
+mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
   by (import word32 word_mul)
 
-definition word_1comp :: "word32 => word32" where 
-  "word_1comp ==
-%T1::word32. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
+definition
+  word_1comp :: "word32 => word32"  where
+  "word_1comp == %T1. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
 
-lemma word_1comp: "ALL T1::word32.
-   word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
+lemma word_1comp: "word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
   by (import word32 word_1comp)
 
-definition word_2comp :: "word32 => word32" where 
-  "word_2comp ==
-%T1::word32. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
+definition
+  word_2comp :: "word32 => word32"  where
+  "word_2comp == %T1. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
 
-lemma word_2comp: "ALL T1::word32.
-   word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
+lemma word_2comp: "word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
   by (import word32 word_2comp)
 
-definition word_lsr1 :: "word32 => word32" where 
-  "word_lsr1 == %T1::word32. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
+definition
+  word_lsr1 :: "word32 => word32"  where
+  "word_lsr1 == %T1. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
 
-lemma word_lsr1: "ALL T1::word32.
-   word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
+lemma word_lsr1: "word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
   by (import word32 word_lsr1)
 
-definition word_asr1 :: "word32 => word32" where 
-  "word_asr1 == %T1::word32. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
+definition
+  word_asr1 :: "word32 => word32"  where
+  "word_asr1 == %T1. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
 
-lemma word_asr1: "ALL T1::word32.
-   word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
+lemma word_asr1: "word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
   by (import word32 word_asr1)
 
-definition word_ror1 :: "word32 => word32" where 
-  "word_ror1 == %T1::word32. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
+definition
+  word_ror1 :: "word32 => word32"  where
+  "word_ror1 == %T1. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
 
-lemma word_ror1: "ALL T1::word32.
-   word_ror1 T1 = mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
+lemma word_ror1: "word_ror1 T1 = mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
   by (import word32 word_ror1)
 
 consts
   RRX :: "bool => word32 => word32" 
 
 defs
-  RRX_primdef: "RRX ==
-%(T1::bool) T2::word32. mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))"
+  RRX_primdef: "RRX == %T1 T2. mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))"
 
-lemma RRX_def: "ALL (T1::bool) T2::word32.
-   RRX T1 T2 = mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))"
+lemma RRX_def: "RRX T1 T2 = mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))"
   by (import word32 RRX_def)
 
 consts
   LSB :: "word32 => bool" 
 
 defs
-  LSB_primdef: "LSB == %T1::word32. LSBn (Eps (dest_word32 T1))"
+  LSB_primdef: "LSB == %T1. LSBn (Eps (dest_word32 T1))"
 
-lemma LSB_def: "ALL T1::word32. LSB T1 = LSBn (Eps (dest_word32 T1))"
+lemma LSB_def: "LSB T1 = LSBn (Eps (dest_word32 T1))"
   by (import word32 LSB_def)
 
 consts
   MSB :: "word32 => bool" 
 
 defs
-  MSB_primdef: "MSB == %T1::word32. MSBn (Eps (dest_word32 T1))"
+  MSB_primdef: "MSB == %T1. MSBn (Eps (dest_word32 T1))"
 
-lemma MSB_def: "ALL T1::word32. MSB T1 = MSBn (Eps (dest_word32 T1))"
+lemma MSB_def: "MSB T1 = MSBn (Eps (dest_word32 T1))"
   by (import word32 MSB_def)
 
-definition bitwise_or :: "word32 => word32 => word32" where 
+definition
+  bitwise_or :: "word32 => word32 => word32"  where
   "bitwise_or ==
-%(T1::word32) T2::word32.
-   mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+%T1 T2. mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
 
-lemma bitwise_or: "ALL (T1::word32) T2::word32.
-   bitwise_or T1 T2 =
-   mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+lemma bitwise_or: "bitwise_or T1 T2 =
+mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
   by (import word32 bitwise_or)
 
-definition bitwise_eor :: "word32 => word32 => word32" where 
+definition
+  bitwise_eor :: "word32 => word32 => word32"  where
   "bitwise_eor ==
-%(T1::word32) T2::word32.
+%T1 T2.
    mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
 
-lemma bitwise_eor: "ALL (T1::word32) T2::word32.
-   bitwise_eor T1 T2 =
-   mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+lemma bitwise_eor: "bitwise_eor T1 T2 =
+mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
   by (import word32 bitwise_eor)
 
-definition bitwise_and :: "word32 => word32 => word32" where 
+definition
+  bitwise_and :: "word32 => word32 => word32"  where
   "bitwise_and ==
-%(T1::word32) T2::word32.
+%T1 T2.
    mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
 
-lemma bitwise_and: "ALL (T1::word32) T2::word32.
-   bitwise_and T1 T2 =
-   mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+lemma bitwise_and: "bitwise_and T1 T2 =
+mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
   by (import word32 bitwise_and)
 
 consts
   TOw :: "word32 => word32" 
 
 defs
-  TOw_primdef: "TOw == %T1::word32. mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
+  TOw_primdef: "TOw == %T1. mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
 
-lemma TOw_def: "ALL T1::word32. TOw T1 = mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
+lemma TOw_def: "TOw T1 = mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
   by (import word32 TOw_def)
 
 consts
   n2w :: "nat => word32" 
 
 defs
-  n2w_primdef: "n2w == %n::nat. mk_word32 (EQUIV n)"
+  n2w_primdef: "n2w == %n. mk_word32 (EQUIV n)"
 
-lemma n2w_def: "ALL n::nat. n2w n = mk_word32 (EQUIV n)"
+lemma n2w_def: "n2w n = mk_word32 (EQUIV n)"
   by (import word32 n2w_def)
 
 consts
   w2n :: "word32 => nat" 
 
 defs
-  w2n_primdef: "w2n == %w::word32. MODw (Eps (dest_word32 w))"
+  w2n_primdef: "w2n == %w. MODw (Eps (dest_word32 w))"
 
-lemma w2n_def: "ALL w::word32. w2n w = MODw (Eps (dest_word32 w))"
+lemma w2n_def: "w2n w = MODw (Eps (dest_word32 w))"
   by (import word32 w2n_def)
 
-lemma ADDw: "(ALL x::word32. word_add w_0 x = x) &
-(ALL (x::word32) xa::word32.
-    word_add (word_suc x) xa = word_suc (word_add x xa))"
+lemma ADDw: "(ALL x. word_add w_0 x = x) &
+(ALL x xa. word_add (word_suc x) xa = word_suc (word_add x xa))"
   by (import word32 ADDw)
 
-lemma ADD_0w: "ALL x::word32. word_add x w_0 = x"
+lemma ADD_0w: "word_add x w_0 = x"
   by (import word32 ADD_0w)
 
-lemma ADD1w: "ALL x::word32. word_suc x = word_add x w_1"
+lemma ADD1w: "word_suc x = word_add x w_1"
   by (import word32 ADD1w)
 
-lemma ADD_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32.
-   word_add x (word_add xa xb) = word_add (word_add x xa) xb"
+lemma ADD_ASSOCw: "word_add x (word_add xa xb) = word_add (word_add x xa) xb"
   by (import word32 ADD_ASSOCw)
 
-lemma ADD_CLAUSESw: "(ALL x::word32. word_add w_0 x = x) &
-(ALL x::word32. word_add x w_0 = x) &
-(ALL (x::word32) xa::word32.
-    word_add (word_suc x) xa = word_suc (word_add x xa)) &
-(ALL (x::word32) xa::word32.
-    word_add x (word_suc xa) = word_suc (word_add x xa))"
+lemma ADD_CLAUSESw: "(ALL x. word_add w_0 x = x) &
+(ALL x. word_add x w_0 = x) &
+(ALL x xa. word_add (word_suc x) xa = word_suc (word_add x xa)) &
+(ALL x xa. word_add x (word_suc xa) = word_suc (word_add x xa))"
   by (import word32 ADD_CLAUSESw)
 
-lemma ADD_COMMw: "ALL (x::word32) xa::word32. word_add x xa = word_add xa x"
+lemma ADD_COMMw: "word_add x xa = word_add xa x"
   by (import word32 ADD_COMMw)
 
-lemma ADD_INV_0_EQw: "ALL (x::word32) xa::word32. (word_add x xa = x) = (xa = w_0)"
+lemma ADD_INV_0_EQw: "(word_add x xa = x) = (xa = w_0)"
   by (import word32 ADD_INV_0_EQw)
 
-lemma EQ_ADD_LCANCELw: "ALL (x::word32) (xa::word32) xb::word32.
-   (word_add x xa = word_add x xb) = (xa = xb)"
+lemma EQ_ADD_LCANCELw: "(word_add x xa = word_add x xb) = (xa = xb)"
   by (import word32 EQ_ADD_LCANCELw)
 
-lemma EQ_ADD_RCANCELw: "ALL (x::word32) (xa::word32) xb::word32.
-   (word_add x xb = word_add xa xb) = (x = xa)"
+lemma EQ_ADD_RCANCELw: "(word_add x xb = word_add xa xb) = (x = xa)"
   by (import word32 EQ_ADD_RCANCELw)
 
-lemma LEFT_ADD_DISTRIBw: "ALL (x::word32) (xa::word32) xb::word32.
-   word_mul xb (word_add x xa) = word_add (word_mul xb x) (word_mul xb xa)"
+lemma LEFT_ADD_DISTRIBw: "word_mul xb (word_add x xa) = word_add (word_mul xb x) (word_mul xb xa)"
   by (import word32 LEFT_ADD_DISTRIBw)
 
-lemma MULT_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32.
-   word_mul x (word_mul xa xb) = word_mul (word_mul x xa) xb"
+lemma MULT_ASSOCw: "word_mul x (word_mul xa xb) = word_mul (word_mul x xa) xb"
   by (import word32 MULT_ASSOCw)
 
-lemma MULT_COMMw: "ALL (x::word32) xa::word32. word_mul x xa = word_mul xa x"
+lemma MULT_COMMw: "word_mul x xa = word_mul xa x"
   by (import word32 MULT_COMMw)
 
-lemma MULT_CLAUSESw: "ALL (x::word32) xa::word32.
-   word_mul w_0 x = w_0 &
-   word_mul x w_0 = w_0 &
-   word_mul w_1 x = x &
-   word_mul x w_1 = x &
-   word_mul (word_suc x) xa = word_add (word_mul x xa) xa &
-   word_mul x (word_suc xa) = word_add x (word_mul x xa)"
+lemma MULT_CLAUSESw: "word_mul w_0 x = w_0 &
+word_mul x w_0 = w_0 &
+word_mul w_1 x = x &
+word_mul x w_1 = x &
+word_mul (word_suc x) xa = word_add (word_mul x xa) xa &
+word_mul x (word_suc xa) = word_add x (word_mul x xa)"
   by (import word32 MULT_CLAUSESw)
 
-lemma TWO_COMP_ONE_COMP: "ALL x::word32. word_2comp x = word_add (word_1comp x) w_1"
+lemma TWO_COMP_ONE_COMP: "word_2comp x = word_add (word_1comp x) w_1"
   by (import word32 TWO_COMP_ONE_COMP)
 
-lemma OR_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32.
-   bitwise_or x (bitwise_or xa xb) = bitwise_or (bitwise_or x xa) xb"
+lemma OR_ASSOCw: "bitwise_or x (bitwise_or xa xb) = bitwise_or (bitwise_or x xa) xb"
   by (import word32 OR_ASSOCw)
 
-lemma OR_COMMw: "ALL (x::word32) xa::word32. bitwise_or x xa = bitwise_or xa x"
+lemma OR_COMMw: "bitwise_or x xa = bitwise_or xa x"
   by (import word32 OR_COMMw)
 
-lemma OR_IDEMw: "ALL x::word32. bitwise_or x x = x"
+lemma OR_IDEMw: "bitwise_or x x = x"
   by (import word32 OR_IDEMw)
 
-lemma OR_ABSORBw: "ALL (x::word32) xa::word32. bitwise_and x (bitwise_or x xa) = x"
+lemma OR_ABSORBw: "bitwise_and x (bitwise_or x xa) = x"
   by (import word32 OR_ABSORBw)
 
-lemma AND_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32.
-   bitwise_and x (bitwise_and xa xb) = bitwise_and (bitwise_and x xa) xb"
+lemma AND_ASSOCw: "bitwise_and x (bitwise_and xa xb) = bitwise_and (bitwise_and x xa) xb"
   by (import word32 AND_ASSOCw)
 
-lemma AND_COMMw: "ALL (x::word32) xa::word32. bitwise_and x xa = bitwise_and xa x"
+lemma AND_COMMw: "bitwise_and x xa = bitwise_and xa x"
   by (import word32 AND_COMMw)
 
-lemma AND_IDEMw: "ALL x::word32. bitwise_and x x = x"
+lemma AND_IDEMw: "bitwise_and x x = x"
   by (import word32 AND_IDEMw)
 
-lemma AND_ABSORBw: "ALL (x::word32) xa::word32. bitwise_or x (bitwise_and x xa) = x"
+lemma AND_ABSORBw: "bitwise_or x (bitwise_and x xa) = x"
   by (import word32 AND_ABSORBw)
 
-lemma ONE_COMPw: "ALL x::word32. word_1comp (word_1comp x) = x"
+lemma ONE_COMPw: "word_1comp (word_1comp x) = x"
   by (import word32 ONE_COMPw)
 
-lemma RIGHT_AND_OVER_ORw: "ALL (x::word32) (xa::word32) xb::word32.
-   bitwise_and (bitwise_or x xa) xb =
-   bitwise_or (bitwise_and x xb) (bitwise_and xa xb)"
+lemma RIGHT_AND_OVER_ORw: "bitwise_and (bitwise_or x xa) xb =
+bitwise_or (bitwise_and x xb) (bitwise_and xa xb)"
   by (import word32 RIGHT_AND_OVER_ORw)
 
-lemma RIGHT_OR_OVER_ANDw: "ALL (x::word32) (xa::word32) xb::word32.
-   bitwise_or (bitwise_and x xa) xb =
-   bitwise_and (bitwise_or x xb) (bitwise_or xa xb)"
+lemma RIGHT_OR_OVER_ANDw: "bitwise_or (bitwise_and x xa) xb =
+bitwise_and (bitwise_or x xb) (bitwise_or xa xb)"
   by (import word32 RIGHT_OR_OVER_ANDw)
 
-lemma DE_MORGAN_THMw: "ALL (x::word32) xa::word32.
-   word_1comp (bitwise_and x xa) =
-   bitwise_or (word_1comp x) (word_1comp xa) &
-   word_1comp (bitwise_or x xa) = bitwise_and (word_1comp x) (word_1comp xa)"
+lemma DE_MORGAN_THMw: "word_1comp (bitwise_and x xa) = bitwise_or (word_1comp x) (word_1comp xa) &
+word_1comp (bitwise_or x xa) = bitwise_and (word_1comp x) (word_1comp xa)"
   by (import word32 DE_MORGAN_THMw)
 
 lemma w_0: "w_0 = n2w 0"
@@ -1136,433 +972,411 @@
                                 ALT_ZERO)))))))))))))))))))))))))))))))))"
   by (import word32 w_T)
 
-lemma ADD_TWO_COMP: "ALL x::word32. word_add x (word_2comp x) = w_0"
+lemma ADD_TWO_COMP: "word_add x (word_2comp x) = w_0"
   by (import word32 ADD_TWO_COMP)
 
-lemma ADD_TWO_COMP2: "ALL x::word32. word_add (word_2comp x) x = w_0"
+lemma ADD_TWO_COMP2: "word_add (word_2comp x) x = w_0"
   by (import word32 ADD_TWO_COMP2)
 
-definition word_sub :: "word32 => word32 => word32" where 
-  "word_sub == %(a::word32) b::word32. word_add a (word_2comp b)"
+definition
+  word_sub :: "word32 => word32 => word32"  where
+  "word_sub == %a b. word_add a (word_2comp b)"
 
-lemma word_sub: "ALL (a::word32) b::word32. word_sub a b = word_add a (word_2comp b)"
+lemma word_sub: "word_sub a b = word_add a (word_2comp b)"
   by (import word32 word_sub)
 
-definition word_lsl :: "word32 => nat => word32" where 
-  "word_lsl == %(a::word32) n::nat. word_mul a (n2w (2 ^ n))"
+definition
+  word_lsl :: "word32 => nat => word32"  where
+  "word_lsl == %a n. word_mul a (n2w (2 ^ n))"
 
-lemma word_lsl: "ALL (a::word32) n::nat. word_lsl a n = word_mul a (n2w (2 ^ n))"
+lemma word_lsl: "word_lsl a n = word_mul a (n2w (2 ^ n))"
   by (import word32 word_lsl)
 
-definition word_lsr :: "word32 => nat => word32" where 
-  "word_lsr == %(a::word32) n::nat. (word_lsr1 ^^ n) a"
+definition
+  word_lsr :: "word32 => nat => word32"  where
+  "word_lsr == %a n. (word_lsr1 ^^ n) a"
 
-lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^^ n) a"
+lemma word_lsr: "word_lsr a n = (word_lsr1 ^^ n) a"
   by (import word32 word_lsr)
 
-definition word_asr :: "word32 => nat => word32" where 
-  "word_asr == %(a::word32) n::nat. (word_asr1 ^^ n) a"
+definition
+  word_asr :: "word32 => nat => word32"  where
+  "word_asr == %a n. (word_asr1 ^^ n) a"
 
-lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^^ n) a"
+lemma word_asr: "word_asr a n = (word_asr1 ^^ n) a"
   by (import word32 word_asr)
 
-definition word_ror :: "word32 => nat => word32" where 
-  "word_ror == %(a::word32) n::nat. (word_ror1 ^^ n) a"
+definition
+  word_ror :: "word32 => nat => word32"  where
+  "word_ror == %a n. (word_ror1 ^^ n) a"
 
-lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^^ n) a"
+lemma word_ror: "word_ror a n = (word_ror1 ^^ n) a"
   by (import word32 word_ror)
 
 consts
   BITw :: "nat => word32 => bool" 
 
 defs
-  BITw_primdef: "BITw == %(b::nat) n::word32. bit b (w2n n)"
+  BITw_primdef: "BITw == %b n. bit b (w2n n)"
 
-lemma BITw_def: "ALL (b::nat) n::word32. BITw b n = bit b (w2n n)"
+lemma BITw_def: "BITw b n = bit b (w2n n)"
   by (import word32 BITw_def)
 
 consts
   BITSw :: "nat => nat => word32 => nat" 
 
 defs
-  BITSw_primdef: "BITSw == %(h::nat) (l::nat) n::word32. BITS h l (w2n n)"
+  BITSw_primdef: "BITSw == %h l n. BITS h l (w2n n)"
 
-lemma BITSw_def: "ALL (h::nat) (l::nat) n::word32. BITSw h l n = BITS h l (w2n n)"
+lemma BITSw_def: "BITSw h l n = BITS h l (w2n n)"
   by (import word32 BITSw_def)
 
 consts
   SLICEw :: "nat => nat => word32 => nat" 
 
 defs
-  SLICEw_primdef: "SLICEw == %(h::nat) (l::nat) n::word32. SLICE h l (w2n n)"
+  SLICEw_primdef: "SLICEw == %h l n. SLICE h l (w2n n)"
 
-lemma SLICEw_def: "ALL (h::nat) (l::nat) n::word32. SLICEw h l n = SLICE h l (w2n n)"
+lemma SLICEw_def: "SLICEw h l n = SLICE h l (w2n n)"
   by (import word32 SLICEw_def)
 
-lemma TWO_COMP_ADD: "ALL (a::word32) b::word32.
-   word_2comp (word_add a b) = word_add (word_2comp a) (word_2comp b)"
+lemma TWO_COMP_ADD: "word_2comp (word_add a b) = word_add (word_2comp a) (word_2comp b)"
   by (import word32 TWO_COMP_ADD)
 
-lemma TWO_COMP_ELIM: "ALL a::word32. word_2comp (word_2comp a) = a"
+lemma TWO_COMP_ELIM: "word_2comp (word_2comp a) = a"
   by (import word32 TWO_COMP_ELIM)
 
-lemma ADD_SUB_ASSOC: "ALL (a::word32) (b::word32) c::word32.
-   word_sub (word_add a b) c = word_add a (word_sub b c)"
+lemma ADD_SUB_ASSOC: "word_sub (word_add a b) c = word_add a (word_sub b c)"
   by (import word32 ADD_SUB_ASSOC)
 
-lemma ADD_SUB_SYM: "ALL (a::word32) (b::word32) c::word32.
-   word_sub (word_add a b) c = word_add (word_sub a c) b"
+lemma ADD_SUB_SYM: "word_sub (word_add a b) c = word_add (word_sub a c) b"
   by (import word32 ADD_SUB_SYM)
 
-lemma SUB_EQUALw: "ALL a::word32. word_sub a a = w_0"
+lemma SUB_EQUALw: "word_sub a a = w_0"
   by (import word32 SUB_EQUALw)
 
-lemma ADD_SUBw: "ALL (a::word32) b::word32. word_sub (word_add a b) b = a"
+lemma ADD_SUBw: "word_sub (word_add a b) b = a"
   by (import word32 ADD_SUBw)
 
-lemma SUB_SUBw: "ALL (a::word32) (b::word32) c::word32.
-   word_sub a (word_sub b c) = word_sub (word_add a c) b"
+lemma SUB_SUBw: "word_sub a (word_sub b c) = word_sub (word_add a c) b"
   by (import word32 SUB_SUBw)
 
-lemma ONE_COMP_TWO_COMP: "ALL a::word32. word_1comp a = word_sub (word_2comp a) w_1"
+lemma ONE_COMP_TWO_COMP: "word_1comp a = word_sub (word_2comp a) w_1"
   by (import word32 ONE_COMP_TWO_COMP)
 
-lemma SUBw: "ALL (m::word32) n::word32. word_sub (word_suc m) n = word_suc (word_sub m n)"
+lemma SUBw: "word_sub (word_suc m) n = word_suc (word_sub m n)"
   by (import word32 SUBw)
 
-lemma ADD_EQ_SUBw: "ALL (m::word32) (n::word32) p::word32.
-   (word_add m n = p) = (m = word_sub p n)"
+lemma ADD_EQ_SUBw: "(word_add m n = p) = (m = word_sub p n)"
   by (import word32 ADD_EQ_SUBw)
 
-lemma CANCEL_SUBw: "ALL (m::word32) (n::word32) p::word32.
-   (word_sub n p = word_sub m p) = (n = m)"
+lemma CANCEL_SUBw: "(word_sub n p = word_sub m p) = (n = m)"
   by (import word32 CANCEL_SUBw)
 
-lemma SUB_PLUSw: "ALL (a::word32) (b::word32) c::word32.
-   word_sub a (word_add b c) = word_sub (word_sub a b) c"
+lemma SUB_PLUSw: "word_sub a (word_add b c) = word_sub (word_sub a b) c"
   by (import word32 SUB_PLUSw)
 
-lemma word_nchotomy: "ALL w::word32. EX n::nat. w = n2w n"
+lemma word_nchotomy: "EX n. w = n2w n"
   by (import word32 word_nchotomy)
 
-lemma dest_word_mk_word_eq3: "ALL a::nat. dest_word32 (mk_word32 (EQUIV a)) = EQUIV a"
+lemma dest_word_mk_word_eq3: "dest_word32 (mk_word32 (EQUIV a)) = EQUIV a"
   by (import word32 dest_word_mk_word_eq3)
 
-lemma MODw_ELIM: "ALL n::nat. n2w (MODw n) = n2w n"
+lemma MODw_ELIM: "n2w (MODw n) = n2w n"
   by (import word32 MODw_ELIM)
 
-lemma w2n_EVAL: "ALL n::nat. w2n (n2w n) = MODw n"
+lemma w2n_EVAL: "w2n (n2w n) = MODw n"
   by (import word32 w2n_EVAL)
 
-lemma w2n_ELIM: "ALL a::word32. n2w (w2n a) = a"
+lemma w2n_ELIM: "n2w (w2n a) = a"
   by (import word32 w2n_ELIM)
 
-lemma n2w_11: "ALL (a::nat) b::nat. (n2w a = n2w b) = (MODw a = MODw b)"
+lemma n2w_11: "(n2w a = n2w b) = (MODw a = MODw b)"
   by (import word32 n2w_11)
 
-lemma ADD_EVAL: "word_add (n2w (a::nat)) (n2w (b::nat)) = n2w (a + b)"
+lemma ADD_EVAL: "word_add (n2w a) (n2w b) = n2w (a + b)"
   by (import word32 ADD_EVAL)
 
-lemma MUL_EVAL: "word_mul (n2w (a::nat)) (n2w (b::nat)) = n2w (a * b)"
+lemma MUL_EVAL: "word_mul (n2w a) (n2w b) = n2w (a * b)"
   by (import word32 MUL_EVAL)
 
-lemma ONE_COMP_EVAL: "word_1comp (n2w (a::nat)) = n2w (ONE_COMP a)"
+lemma ONE_COMP_EVAL: "word_1comp (n2w a) = n2w (ONE_COMP a)"
   by (import word32 ONE_COMP_EVAL)
 
-lemma TWO_COMP_EVAL: "word_2comp (n2w (a::nat)) = n2w (TWO_COMP a)"
+lemma TWO_COMP_EVAL: "word_2comp (n2w a) = n2w (TWO_COMP a)"
   by (import word32 TWO_COMP_EVAL)
 
-lemma LSR_ONE_EVAL: "word_lsr1 (n2w (a::nat)) = n2w (LSR_ONE a)"
+lemma LSR_ONE_EVAL: "word_lsr1 (n2w a) = n2w (LSR_ONE a)"
   by (import word32 LSR_ONE_EVAL)
 
-lemma ASR_ONE_EVAL: "word_asr1 (n2w (a::nat)) = n2w (ASR_ONE a)"
+lemma ASR_ONE_EVAL: "word_asr1 (n2w a) = n2w (ASR_ONE a)"
   by (import word32 ASR_ONE_EVAL)
 
-lemma ROR_ONE_EVAL: "word_ror1 (n2w (a::nat)) = n2w (ROR_ONE a)"
+lemma ROR_ONE_EVAL: "word_ror1 (n2w a) = n2w (ROR_ONE a)"
   by (import word32 ROR_ONE_EVAL)
 
-lemma RRX_EVAL: "RRX (c::bool) (n2w (a::nat)) = n2w (RRXn c a)"
+lemma RRX_EVAL: "RRX c (n2w a) = n2w (RRXn c a)"
   by (import word32 RRX_EVAL)
 
-lemma LSB_EVAL: "LSB (n2w (a::nat)) = LSBn a"
+lemma LSB_EVAL: "LSB (n2w a) = LSBn a"
   by (import word32 LSB_EVAL)
 
-lemma MSB_EVAL: "MSB (n2w (a::nat)) = MSBn a"
+lemma MSB_EVAL: "MSB (n2w a) = MSBn a"
   by (import word32 MSB_EVAL)
 
-lemma OR_EVAL: "bitwise_or (n2w (a::nat)) (n2w (b::nat)) = n2w (OR a b)"
+lemma OR_EVAL: "bitwise_or (n2w a) (n2w b) = n2w (OR a b)"
   by (import word32 OR_EVAL)
 
-lemma EOR_EVAL: "bitwise_eor (n2w (a::nat)) (n2w (b::nat)) = n2w (EOR a b)"
+lemma EOR_EVAL: "bitwise_eor (n2w a) (n2w b) = n2w (EOR a b)"
   by (import word32 EOR_EVAL)
 
-lemma AND_EVAL: "bitwise_and (n2w (a::nat)) (n2w (b::nat)) = n2w (AND a b)"
+lemma AND_EVAL: "bitwise_and (n2w a) (n2w b) = n2w (AND a b)"
   by (import word32 AND_EVAL)
 
-lemma BITS_EVAL: "ALL (h::nat) (l::nat) a::nat. BITSw h l (n2w a) = BITS h l (MODw a)"
+lemma BITS_EVAL: "BITSw h l (n2w a) = BITS h l (MODw a)"
   by (import word32 BITS_EVAL)
 
-lemma BIT_EVAL: "ALL (b::nat) a::nat. BITw b (n2w a) = bit b (MODw a)"
+lemma BIT_EVAL: "BITw b (n2w a) = bit b (MODw a)"
   by (import word32 BIT_EVAL)
 
-lemma SLICE_EVAL: "ALL (h::nat) (l::nat) a::nat. SLICEw h l (n2w a) = SLICE h l (MODw a)"
+lemma SLICE_EVAL: "SLICEw h l (n2w a) = SLICE h l (MODw a)"
   by (import word32 SLICE_EVAL)
 
-lemma LSL_ADD: "ALL (a::word32) (m::nat) n::nat.
-   word_lsl (word_lsl a m) n = word_lsl a (m + n)"
+lemma LSL_ADD: "word_lsl (word_lsl a m) n = word_lsl a (m + n)"
   by (import word32 LSL_ADD)
 
-lemma LSR_ADD: "ALL (x::word32) (xa::nat) xb::nat.
-   word_lsr (word_lsr x xa) xb = word_lsr x (xa + xb)"
+lemma LSR_ADD: "word_lsr (word_lsr x xa) xb = word_lsr x (xa + xb)"
   by (import word32 LSR_ADD)
 
-lemma ASR_ADD: "ALL (x::word32) (xa::nat) xb::nat.
-   word_asr (word_asr x xa) xb = word_asr x (xa + xb)"
+lemma ASR_ADD: "word_asr (word_asr x xa) xb = word_asr x (xa + xb)"
   by (import word32 ASR_ADD)
 
-lemma ROR_ADD: "ALL (x::word32) (xa::nat) xb::nat.
-   word_ror (word_ror x xa) xb = word_ror x (xa + xb)"
+lemma ROR_ADD: "word_ror (word_ror x xa) xb = word_ror x (xa + xb)"
   by (import word32 ROR_ADD)
 
-lemma LSL_LIMIT: "ALL (w::word32) n::nat. HB < n --> word_lsl w n = w_0"
+lemma LSL_LIMIT: "HB < n ==> word_lsl w n = w_0"
   by (import word32 LSL_LIMIT)
 
-lemma MOD_MOD_DIV: "ALL (a::nat) b::nat. INw (MODw a div 2 ^ b)"
+lemma MOD_MOD_DIV: "INw (MODw a div 2 ^ b)"
   by (import word32 MOD_MOD_DIV)
 
-lemma MOD_MOD_DIV_2EXP: "ALL (a::nat) n::nat. MODw (MODw a div 2 ^ n) div 2 = MODw a div 2 ^ Suc n"
+lemma MOD_MOD_DIV_2EXP: "MODw (MODw a div 2 ^ n) div 2 = MODw a div 2 ^ Suc n"
   by (import word32 MOD_MOD_DIV_2EXP)
 
-lemma LSR_EVAL: "ALL n::nat. word_lsr (n2w (a::nat)) n = n2w (MODw a div 2 ^ n)"
+lemma LSR_EVAL: "word_lsr (n2w a) n = n2w (MODw a div 2 ^ n)"
   by (import word32 LSR_EVAL)
 
-lemma LSR_THM: "ALL (x::nat) n::nat. word_lsr (n2w n) x = n2w (BITS HB (min WL x) n)"
+lemma LSR_THM: "word_lsr (n2w n) x = n2w (BITS HB (min WL x) n)"
   by (import word32 LSR_THM)
 
-lemma LSR_LIMIT: "ALL (x::nat) w::word32. HB < x --> word_lsr w x = w_0"
+lemma LSR_LIMIT: "HB < x ==> word_lsr w x = w_0"
   by (import word32 LSR_LIMIT)
 
-lemma LEFT_SHIFT_LESS: "ALL (n::nat) (m::nat) a::nat. a < 2 ^ m --> 2 ^ n + a * 2 ^ n <= 2 ^ (m + n)"
+lemma LEFT_SHIFT_LESS: "(a::nat) < (2::nat) ^ (m::nat)
+==> (2::nat) ^ (n::nat) + a * (2::nat) ^ n <= (2::nat) ^ (m + n)"
   by (import word32 LEFT_SHIFT_LESS)
 
-lemma ROR_THM: "ALL (x::nat) n::nat.
-   word_ror (n2w n) x =
-   (let x'::nat = x mod WL
-    in n2w (BITS HB x' n + BITS (x' - 1) 0 n * 2 ^ (WL - x')))"
+lemma ROR_THM: "word_ror (n2w n) x =
+(let x' = x mod WL
+ in n2w (BITS HB x' n + BITS (x' - 1) 0 n * 2 ^ (WL - x')))"
   by (import word32 ROR_THM)
 
-lemma ROR_CYCLE: "ALL (x::nat) w::word32. word_ror w (x * WL) = w"
+lemma ROR_CYCLE: "word_ror w (x * WL) = w"
   by (import word32 ROR_CYCLE)
 
-lemma ASR_THM: "ALL (x::nat) n::nat.
-   word_asr (n2w n) x =
-   (let x'::nat = min HB x; s::nat = BITS HB x' n
-    in n2w (if MSBn n then 2 ^ WL - 2 ^ (WL - x') + s else s))"
+lemma ASR_THM: "word_asr (n2w n) x =
+(let x' = min HB x; s = BITS HB x' n
+ in n2w (if MSBn n then 2 ^ WL - 2 ^ (WL - x') + s else s))"
   by (import word32 ASR_THM)
 
-lemma ASR_LIMIT: "ALL (x::nat) w::word32.
-   HB <= x --> word_asr w x = (if MSB w then w_T else w_0)"
+lemma ASR_LIMIT: "HB <= x ==> word_asr w x = (if MSB w then w_T else w_0)"
   by (import word32 ASR_LIMIT)
 
-lemma ZERO_SHIFT: "(ALL n::nat. word_lsl w_0 n = w_0) &
-(ALL n::nat. word_asr w_0 n = w_0) &
-(ALL n::nat. word_lsr w_0 n = w_0) & (ALL n::nat. word_ror w_0 n = w_0)"
+lemma ZERO_SHIFT: "(ALL n. word_lsl w_0 n = w_0) &
+(ALL n. word_asr w_0 n = w_0) &
+(ALL n. word_lsr w_0 n = w_0) & (ALL n. word_ror w_0 n = w_0)"
   by (import word32 ZERO_SHIFT)
 
-lemma ZERO_SHIFT2: "(ALL a::word32. word_lsl a 0 = a) &
-(ALL a::word32. word_asr a 0 = a) &
-(ALL a::word32. word_lsr a 0 = a) & (ALL a::word32. word_ror a 0 = a)"
+lemma ZERO_SHIFT2: "(ALL a. word_lsl a 0 = a) &
+(ALL a. word_asr a 0 = a) &
+(ALL a. word_lsr a 0 = a) & (ALL a. word_ror a 0 = a)"
   by (import word32 ZERO_SHIFT2)
 
-lemma ASR_w_T: "ALL n::nat. word_asr w_T n = w_T"
+lemma ASR_w_T: "word_asr w_T n = w_T"
   by (import word32 ASR_w_T)
 
-lemma ROR_w_T: "ALL n::nat. word_ror w_T n = w_T"
+lemma ROR_w_T: "word_ror w_T n = w_T"
   by (import word32 ROR_w_T)
 
-lemma MODw_EVAL: "ALL x::nat.
-   MODw x =
-   x mod
-   NUMERAL
-    (NUMERAL_BIT2
-      (NUMERAL_BIT1
-        (NUMERAL_BIT1
-          (NUMERAL_BIT1
-            (NUMERAL_BIT1
-              (NUMERAL_BIT1
-                (NUMERAL_BIT1
-                  (NUMERAL_BIT1
-                    (NUMERAL_BIT1
-                      (NUMERAL_BIT1
-                        (NUMERAL_BIT1
-                          (NUMERAL_BIT1
-                            (NUMERAL_BIT1
-                              (NUMERAL_BIT1
-                                (NUMERAL_BIT1
-                                  (NUMERAL_BIT1
-                                    (NUMERAL_BIT1
-(NUMERAL_BIT1
-  (NUMERAL_BIT1
-    (NUMERAL_BIT1
-      (NUMERAL_BIT1
-        (NUMERAL_BIT1
-          (NUMERAL_BIT1
-            (NUMERAL_BIT1
-              (NUMERAL_BIT1
-                (NUMERAL_BIT1
-                  (NUMERAL_BIT1
-                    (NUMERAL_BIT1
-                      (NUMERAL_BIT1
-                        (NUMERAL_BIT1
-                          (NUMERAL_BIT1
-                            (NUMERAL_BIT1
-                              ALT_ZERO))))))))))))))))))))))))))))))))"
-  by (import word32 MODw_EVAL)
-
-lemma ADD_EVAL2: "ALL (b::nat) a::nat. word_add (n2w a) (n2w b) = n2w (MODw (a + b))"
-  by (import word32 ADD_EVAL2)
-
-lemma MUL_EVAL2: "ALL (b::nat) a::nat. word_mul (n2w a) (n2w b) = n2w (MODw (a * b))"
-  by (import word32 MUL_EVAL2)
-
-lemma ONE_COMP_EVAL2: "ALL a::nat.
-   word_1comp (n2w a) =
-   n2w (2 ^
-        NUMERAL
-         (NUMERAL_BIT2
-           (NUMERAL_BIT1
-             (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))) -
-        1 -
-        MODw a)"
-  by (import word32 ONE_COMP_EVAL2)
-
-lemma TWO_COMP_EVAL2: "ALL a::nat.
-   word_2comp (n2w a) =
-   n2w (MODw
-         (2 ^
-          NUMERAL
-           (NUMERAL_BIT2
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))) -
-          MODw a))"
-  by (import word32 TWO_COMP_EVAL2)
-
-lemma LSR_ONE_EVAL2: "ALL a::nat. word_lsr1 (n2w a) = n2w (MODw a div 2)"
-  by (import word32 LSR_ONE_EVAL2)
-
-lemma ASR_ONE_EVAL2: "ALL a::nat.
-   word_asr1 (n2w a) =
-   n2w (MODw a div 2 +
-        SBIT (MSBn a)
-         (NUMERAL
+lemma MODw_EVAL: "MODw x =
+x mod
+NUMERAL
+ (NUMERAL_BIT2
+   (NUMERAL_BIT1
+     (NUMERAL_BIT1
+       (NUMERAL_BIT1
+         (NUMERAL_BIT1
            (NUMERAL_BIT1
              (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
-  by (import word32 ASR_ONE_EVAL2)
-
-lemma ROR_ONE_EVAL2: "ALL a::nat.
-   word_ror1 (n2w a) =
-   n2w (MODw a div 2 +
-        SBIT (LSBn a)
-         (NUMERAL
+               (NUMERAL_BIT1
+                 (NUMERAL_BIT1
+                   (NUMERAL_BIT1
+                     (NUMERAL_BIT1
+                       (NUMERAL_BIT1
+                         (NUMERAL_BIT1
+                           (NUMERAL_BIT1
+                             (NUMERAL_BIT1
+                               (NUMERAL_BIT1
+                                 (NUMERAL_BIT1
+                                   (NUMERAL_BIT1
+                                     (NUMERAL_BIT1
+ (NUMERAL_BIT1
+   (NUMERAL_BIT1
+     (NUMERAL_BIT1
+       (NUMERAL_BIT1
+         (NUMERAL_BIT1
            (NUMERAL_BIT1
              (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
+               (NUMERAL_BIT1
+                 (NUMERAL_BIT1
+                   (NUMERAL_BIT1
+                     (NUMERAL_BIT1
+                       (NUMERAL_BIT1
+                         (NUMERAL_BIT1
+                           ALT_ZERO))))))))))))))))))))))))))))))))"
+  by (import word32 MODw_EVAL)
+
+lemma ADD_EVAL2: "word_add (n2w a) (n2w b) = n2w (MODw (a + b))"
+  by (import word32 ADD_EVAL2)
+
+lemma MUL_EVAL2: "word_mul (n2w a) (n2w b) = n2w (MODw (a * b))"
+  by (import word32 MUL_EVAL2)
+
+lemma ONE_COMP_EVAL2: "word_1comp (n2w a) =
+n2w (2 ^
+     NUMERAL
+      (NUMERAL_BIT2
+        (NUMERAL_BIT1
+          (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))) -
+     1 -
+     MODw a)"
+  by (import word32 ONE_COMP_EVAL2)
+
+lemma TWO_COMP_EVAL2: "word_2comp (n2w a) =
+n2w (MODw
+      (2 ^
+       NUMERAL
+        (NUMERAL_BIT2
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))) -
+       MODw a))"
+  by (import word32 TWO_COMP_EVAL2)
+
+lemma LSR_ONE_EVAL2: "word_lsr1 (n2w a) = n2w (MODw a div 2)"
+  by (import word32 LSR_ONE_EVAL2)
+
+lemma ASR_ONE_EVAL2: "word_asr1 (n2w a) =
+n2w (MODw a div 2 +
+     SBIT (MSBn a)
+      (NUMERAL
+        (NUMERAL_BIT1
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
+  by (import word32 ASR_ONE_EVAL2)
+
+lemma ROR_ONE_EVAL2: "word_ror1 (n2w a) =
+n2w (MODw a div 2 +
+     SBIT (LSBn a)
+      (NUMERAL
+        (NUMERAL_BIT1
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
   by (import word32 ROR_ONE_EVAL2)
 
-lemma RRX_EVAL2: "ALL (c::bool) a::nat.
-   RRX c (n2w a) =
-   n2w (MODw a div 2 +
-        SBIT c
-         (NUMERAL
-           (NUMERAL_BIT1
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
+lemma RRX_EVAL2: "RRX c (n2w a) =
+n2w (MODw a div 2 +
+     SBIT c
+      (NUMERAL
+        (NUMERAL_BIT1
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))"
   by (import word32 RRX_EVAL2)
 
-lemma LSB_EVAL2: "ALL a::nat. LSB (n2w a) = ODD a"
+lemma LSB_EVAL2: "LSB (n2w a) = ODD a"
   by (import word32 LSB_EVAL2)
 
-lemma MSB_EVAL2: "ALL a::nat.
-   MSB (n2w a) =
-   bit (NUMERAL
-         (NUMERAL_BIT1
-           (NUMERAL_BIT1
-             (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
-    a"
+lemma MSB_EVAL2: "MSB (n2w a) =
+bit (NUMERAL
+      (NUMERAL_BIT1
+        (NUMERAL_BIT1
+          (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
+ a"
   by (import word32 MSB_EVAL2)
 
-lemma OR_EVAL2: "ALL (b::nat) a::nat.
-   bitwise_or (n2w a) (n2w b) =
-   n2w (BITWISE
-         (NUMERAL
-           (NUMERAL_BIT2
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
-         op | a b)"
+lemma OR_EVAL2: "bitwise_or (n2w a) (n2w b) =
+n2w (BITWISE
+      (NUMERAL
+        (NUMERAL_BIT2
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
+      op | a b)"
   by (import word32 OR_EVAL2)
 
-lemma AND_EVAL2: "ALL (b::nat) a::nat.
-   bitwise_and (n2w a) (n2w b) =
-   n2w (BITWISE
-         (NUMERAL
-           (NUMERAL_BIT2
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
-         op & a b)"
+lemma AND_EVAL2: "bitwise_and (n2w a) (n2w b) =
+n2w (BITWISE
+      (NUMERAL
+        (NUMERAL_BIT2
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
+      op & a b)"
   by (import word32 AND_EVAL2)
 
-lemma EOR_EVAL2: "ALL (b::nat) a::nat.
-   bitwise_eor (n2w a) (n2w b) =
-   n2w (BITWISE
-         (NUMERAL
-           (NUMERAL_BIT2
-             (NUMERAL_BIT1
-               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
-         (%(x::bool) y::bool. x ~= y) a b)"
+lemma EOR_EVAL2: "bitwise_eor (n2w a) (n2w b) =
+n2w (BITWISE
+      (NUMERAL
+        (NUMERAL_BIT2
+          (NUMERAL_BIT1
+            (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
+      op ~= a b)"
   by (import word32 EOR_EVAL2)
 
-lemma BITWISE_EVAL2: "ALL (n::nat) (oper::bool => bool => bool) (x::nat) y::nat.
-   BITWISE n oper x y =
-   (if n = 0 then 0
-    else 2 * BITWISE (n - 1) oper (x div 2) (y div 2) +
-         (if oper (ODD x) (ODD y) then 1 else 0))"
+lemma BITWISE_EVAL2: "BITWISE n oper x y =
+(if n = 0 then 0
+ else 2 * BITWISE (n - 1) oper (x div 2) (y div 2) +
+      (if oper (ODD x) (ODD y) then 1 else 0))"
   by (import word32 BITWISE_EVAL2)
 
-lemma BITSwLT_THM: "ALL (h::nat) (l::nat) n::word32. BITSw h l n < 2 ^ (Suc h - l)"
+lemma BITSwLT_THM: "BITSw h l n < 2 ^ (Suc h - l)"
   by (import word32 BITSwLT_THM)
 
-lemma BITSw_COMP_THM: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::word32.
-   h2 + l1 <= h1 -->
-   BITS h2 l2 (BITSw h1 l1 n) = BITSw (h2 + l1) (l2 + l1) n"
+lemma BITSw_COMP_THM: "h2 + l1 <= h1 ==> BITS h2 l2 (BITSw h1 l1 n) = BITSw (h2 + l1) (l2 + l1) n"
   by (import word32 BITSw_COMP_THM)
 
-lemma BITSw_DIV_THM: "ALL (h::nat) (l::nat) (n::nat) x::word32.
-   BITSw h l x div 2 ^ n = BITSw h (l + n) x"
+lemma BITSw_DIV_THM: "BITSw h l x div 2 ^ n = BITSw h (l + n) x"
   by (import word32 BITSw_DIV_THM)
 
-lemma BITw_THM: "ALL (b::nat) n::word32. BITw b n = (BITSw b b n = 1)"
+lemma BITw_THM: "BITw b n = (BITSw b b n = 1)"
   by (import word32 BITw_THM)
 
-lemma SLICEw_THM: "ALL (n::word32) (h::nat) l::nat. SLICEw h l n = BITSw h l n * 2 ^ l"
+lemma SLICEw_THM: "SLICEw h l n = BITSw h l n * 2 ^ l"
   by (import word32 SLICEw_THM)
 
-lemma BITS_SLICEw_THM: "ALL (h::nat) (l::nat) n::word32. BITS h l (SLICEw h l n) = BITSw h l n"
+lemma BITS_SLICEw_THM: "BITS h l (SLICEw h l n) = BITSw h l n"
   by (import word32 BITS_SLICEw_THM)
 
-lemma SLICEw_ZERO_THM: "ALL (n::word32) h::nat. SLICEw h 0 n = BITSw h 0 n"
+lemma SLICEw_ZERO_THM: "SLICEw h 0 n = BITSw h 0 n"
   by (import word32 SLICEw_ZERO_THM)
 
-lemma SLICEw_COMP_THM: "ALL (h::nat) (m::nat) (l::nat) a::word32.
-   Suc m <= h & l <= m --> SLICEw h (Suc m) a + SLICEw m l a = SLICEw h l a"
+lemma SLICEw_COMP_THM: "Suc m <= h & l <= m ==> SLICEw h (Suc m) a + SLICEw m l a = SLICEw h l a"
   by (import word32 SLICEw_COMP_THM)
 
-lemma BITSw_ZERO: "ALL (h::nat) (l::nat) n::word32. h < l --> BITSw h l n = 0"
+lemma BITSw_ZERO: "h < l ==> BITSw h l n = 0"
   by (import word32 BITSw_ZERO)
 
-lemma SLICEw_ZERO: "ALL (h::nat) (l::nat) n::word32. h < l --> SLICEw h l n = 0"
+lemma SLICEw_ZERO: "h < l ==> SLICEw h l n = 0"
   by (import word32 SLICEw_ZERO)
 
 ;end_setup
 
 end
+
--- a/src/HOL/Import/HOL/arithmetic.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/arithmetic.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -14,19 +14,19 @@
   "NUMERAL_BIT2" > "HOL4Compat.NUMERAL_BIT2"
   "NUMERAL_BIT1" > "HOL4Compat.NUMERAL_BIT1"
   "NUMERAL" > "HOL4Compat.NUMERAL"
+  "MOD" > "Divides.div_class.mod" :: "nat => nat => nat"
   "MIN" > "Orderings.ord_class.min" :: "nat => nat => nat"
   "MAX" > "Orderings.ord_class.max" :: "nat => nat => nat"
   "FUNPOW" > "HOL4Compat.FUNPOW"
   "EXP" > "Power.power_class.power" :: "nat => nat => nat"
   "DIV" > "Divides.div_class.div" :: "nat => nat => nat"
-  "MOD" > "Divides.div_class.mod" :: "nat => nat => nat"
   "ALT_ZERO" > "HOL4Compat.ALT_ZERO"
   ">=" > "HOL4Compat.nat_ge"
   ">" > "HOL4Compat.nat_gt"
-  "<=" > "Orderings.less_eq" :: "nat => nat => bool"
-  "-" > "Groups.minus" :: "nat => nat => nat"
-  "+" > "Groups.plus" :: "nat => nat => nat"
-  "*" > "Groups.times" :: "nat => nat => nat"
+  "<=" > "Orderings.ord_class.less_eq" :: "nat => nat => bool"
+  "-" > "Groups.minus_class.minus" :: "nat => nat => nat"
+  "+" > "Groups.plus_class.plus" :: "nat => nat => nat"
+  "*" > "Groups.times_class.times" :: "nat => nat => nat"
 
 thm_maps
   "num_case_def" > "HOL4Compat.num_case_def"
@@ -41,10 +41,10 @@
   "ZERO_DIV" > "HOL4Base.arithmetic.ZERO_DIV"
   "WOP" > "HOL4Base.arithmetic.WOP"
   "TWO" > "HOL4Base.arithmetic.TWO"
-  "TIMES2" > "NatSimprocs.nat_mult_2"
-  "SUC_SUB1" > "HOL4Base.arithmetic.SUC_SUB1"
+  "TIMES2" > "Int.semiring_mult_2"
+  "SUC_SUB1" > "Nat.diff_Suc_1"
   "SUC_ONE_ADD" > "Nat_Numeral.Suc_eq_plus1_left"
-  "SUC_NOT" > "Nat.nat.simps_2"
+  "SUC_NOT" > "Nat.Zero_not_Suc"
   "SUC_ELIM_THM" > "HOL4Base.arithmetic.SUC_ELIM_THM"
   "SUC_ADD_SYM" > "HOL4Base.arithmetic.SUC_ADD_SYM"
   "SUB_SUB" > "Nat.diff_diff_right"
@@ -77,8 +77,8 @@
   "SUB_ADD" > "Nat.le_add_diff_inverse2"
   "SUB_0" > "HOL4Base.arithmetic.SUB_0"
   "SUB" > "HOL4Compat.SUB"
-  "RIGHT_SUB_DISTRIB" > "Nat.nat_distrib_3"
-  "RIGHT_ADD_DISTRIB" > "Nat.nat_distrib_1"
+  "RIGHT_SUB_DISTRIB" > "Nat.diff_mult_distrib"
+  "RIGHT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_26"
   "PRE_SUC_EQ" > "HOL4Base.arithmetic.PRE_SUC_EQ"
   "PRE_SUB1" > "HOL4Base.arithmetic.PRE_SUB1"
   "PRE_SUB" > "HOL4Base.arithmetic.PRE_SUB"
@@ -101,29 +101,29 @@
   "NOT_SUC_ADD_LESS_EQ" > "HOL4Base.arithmetic.NOT_SUC_ADD_LESS_EQ"
   "NOT_ODD_EQ_EVEN" > "HOL4Base.arithmetic.NOT_ODD_EQ_EVEN"
   "NOT_NUM_EQ" > "HOL4Base.arithmetic.NOT_NUM_EQ"
-  "NOT_LESS_EQUAL" > "Orderings.linorder_not_le"
-  "NOT_LESS" > "Nat.le_def"
-  "NOT_LEQ" > "HOL4Base.arithmetic.NOT_LEQ"
-  "NOT_GREATER_EQ" > "HOL4Base.arithmetic.NOT_GREATER_EQ"
-  "NOT_GREATER" > "Nat.le_def"
+  "NOT_LESS_EQUAL" > "Orderings.linorder_class.not_le"
+  "NOT_LESS" > "Orderings.linorder_class.not_less"
+  "NOT_LEQ" > "Nat.not_less_eq_eq"
+  "NOT_GREATER_EQ" > "Nat.not_less_eq_eq"
+  "NOT_GREATER" > "Orderings.linorder_class.not_less"
   "NOT_EXP_0" > "HOL4Base.arithmetic.NOT_EXP_0"
   "NORM_0" > "HOL4Base.arithmetic.NORM_0"
-  "MULT_SYM" > "Int.zmult_ac_2"
+  "MULT_SYM" > "Fields.linordered_field_class.sign_simps_40"
   "MULT_SUC_EQ" > "HOL4Base.arithmetic.MULT_SUC_EQ"
   "MULT_SUC" > "Nat.mult_Suc_right"
-  "MULT_RIGHT_1" > "Finite_Set.AC_mult.f_e.ident"
+  "MULT_RIGHT_1" > "Divides.arithmetic_simps_43"
   "MULT_MONO_EQ" > "Nat.Suc_mult_cancel1"
   "MULT_LESS_EQ_SUC" > "Nat.Suc_mult_le_cancel1"
-  "MULT_LEFT_1" > "Finite_Set.AC_mult.f_e.left_ident"
+  "MULT_LEFT_1" > "Divides.arithmetic_simps_42"
   "MULT_INCREASES" > "HOL4Base.arithmetic.MULT_INCREASES"
   "MULT_EXP_MONO" > "HOL4Base.arithmetic.MULT_EXP_MONO"
-  "MULT_EQ_1" > "HOL4Base.arithmetic.MULT_EQ_1"
+  "MULT_EQ_1" > "Nat.nat_mult_eq_1_iff"
   "MULT_EQ_0" > "Nat.mult_is_0"
   "MULT_DIV" > "Divides.div_mult_self_is_m"
-  "MULT_COMM" > "Int.zmult_ac_2"
+  "MULT_COMM" > "Fields.linordered_field_class.sign_simps_40"
   "MULT_CLAUSES" > "HOL4Base.arithmetic.MULT_CLAUSES"
-  "MULT_ASSOC" > "Int.zmult_ac_1"
-  "MULT_0" > "Nat.mult_0_right"
+  "MULT_ASSOC" > "Fields.linordered_field_class.sign_simps_41"
+  "MULT_0" > "Divides.arithmetic_simps_41"
   "MULT" > "HOL4Compat.MULT"
   "MOD_UNIQUE" > "HOL4Base.arithmetic.MOD_UNIQUE"
   "MOD_TIMES2" > "HOL4Base.arithmetic.MOD_TIMES2"
@@ -141,19 +141,19 @@
   "MIN_MAX_EQ" > "HOL4Base.arithmetic.MIN_MAX_EQ"
   "MIN_LT" > "HOL4Base.arithmetic.MIN_LT"
   "MIN_LE" > "HOL4Base.arithmetic.MIN_LE"
-  "MIN_IDEM" > "Finite_Set.min.f.ACI_4"
+  "MIN_IDEM" > "Big_Operators.linorder_class.Min.idem"
   "MIN_DEF" > "HOL4Compat.MIN_DEF"
-  "MIN_COMM" > "Finite_Set.min.f.ACI_2"
-  "MIN_ASSOC" > "Finite_Set.min.f.ACI_1"
+  "MIN_COMM" > "Lattices.linorder_class.min_max.inf.commute"
+  "MIN_ASSOC" > "Lattices.linorder_class.min_max.inf.assoc"
   "MIN_0" > "HOL4Base.arithmetic.MIN_0"
   "MAX_LT" > "HOL4Base.arithmetic.MAX_LT"
   "MAX_LE" > "HOL4Base.arithmetic.MAX_LE"
-  "MAX_IDEM" > "Finite_Set.max.f.ACI_4"
+  "MAX_IDEM" > "Big_Operators.linorder_class.Max.idem"
   "MAX_DEF" > "HOL4Compat.MAX_DEF"
-  "MAX_COMM" > "Finite_Set.max.f.ACI_2"
-  "MAX_ASSOC" > "Finite_Set.max.f.ACI_1"
+  "MAX_COMM" > "Lattices.linorder_class.min_max.inf_sup_aci_5"
+  "MAX_ASSOC" > "Lattices.linorder_class.min_max.inf_sup_aci_6"
   "MAX_0" > "HOL4Base.arithmetic.MAX_0"
-  "LESS_TRANS" > "Nat.less_trans"
+  "LESS_TRANS" > "Orderings.order_less_trans"
   "LESS_SUC_NOT" > "HOL4Base.arithmetic.LESS_SUC_NOT"
   "LESS_SUC_EQ_COR" > "Nat.Suc_lessI"
   "LESS_SUB_ADD_LESS" > "HOL4Base.arithmetic.LESS_SUB_ADD_LESS"
@@ -162,28 +162,28 @@
   "LESS_OR" > "Nat.Suc_leI"
   "LESS_NOT_SUC" > "HOL4Base.arithmetic.LESS_NOT_SUC"
   "LESS_MULT_MONO" > "Nat.Suc_mult_less_cancel1"
-  "LESS_MULT2" > "Rings.mult_pos_pos"
+  "LESS_MULT2" > "Rings.linordered_semiring_strict_class.mult_pos_pos"
   "LESS_MONO_REV" > "Nat.Suc_less_SucD"
   "LESS_MONO_MULT" > "Nat.mult_le_mono1"
   "LESS_MONO_EQ" > "Nat.Suc_less_eq"
-  "LESS_MONO_ADD_INV" > "Groups.add_less_imp_less_right"
-  "LESS_MONO_ADD_EQ" > "Groups.add_less_cancel_right"
-  "LESS_MONO_ADD" > "Nat.add_less_mono1"
+  "LESS_MONO_ADD_INV" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_imp_less_right"
+  "LESS_MONO_ADD_EQ" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_right"
+  "LESS_MONO_ADD" > "Groups.ordered_cancel_ab_semigroup_add_class.add_strict_right_mono"
   "LESS_MOD" > "Divides.mod_less"
   "LESS_LESS_SUC" > "HOL4Base.arithmetic.LESS_LESS_SUC"
-  "LESS_LESS_EQ_TRANS" > "Nat.less_le_trans"
+  "LESS_LESS_EQ_TRANS" > "Orderings.order_less_le_trans"
   "LESS_LESS_CASES" > "HOL4Base.arithmetic.LESS_LESS_CASES"
-  "LESS_IMP_LESS_OR_EQ" > "Nat.le_simps_1"
-  "LESS_IMP_LESS_ADD" > "Nat.trans_less_add1"
+  "LESS_IMP_LESS_OR_EQ" > "FunDef.termination_basic_simps_5"
+  "LESS_IMP_LESS_ADD" > "FunDef.termination_basic_simps_1"
   "LESS_EXP_SUC_MONO" > "HOL4Base.arithmetic.LESS_EXP_SUC_MONO"
   "LESS_EQ_TRANS" > "Nat.le_trans"
   "LESS_EQ_SUC_REFL" > "HOL4Base.arithmetic.LESS_EQ_SUC_REFL"
   "LESS_EQ_SUB_LESS" > "HOL4Base.arithmetic.LESS_EQ_SUB_LESS"
-  "LESS_EQ_REFL" > "Finite_Set.max.f_below.below_refl"
-  "LESS_EQ_MONO_ADD_EQ" > "Groups.add_le_cancel_right"
+  "LESS_EQ_REFL" > "Nat.le_refl"
+  "LESS_EQ_MONO_ADD_EQ" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_right"
   "LESS_EQ_MONO" > "Nat.Suc_le_mono"
-  "LESS_EQ_LESS_TRANS" > "Nat.le_less_trans"
-  "LESS_EQ_LESS_EQ_MONO" > "Nat.add_le_mono"
+  "LESS_EQ_LESS_TRANS" > "Orderings.order_le_less_trans"
+  "LESS_EQ_LESS_EQ_MONO" > "Groups.add_mono_thms_linordered_semiring_1"
   "LESS_EQ_IMP_LESS_SUC" > "Nat.le_imp_less_Suc"
   "LESS_EQ_EXISTS" > "Nat.le_iff_add"
   "LESS_EQ_CASES" > "Nat.nat_le_linear"
@@ -192,8 +192,8 @@
   "LESS_EQ_ADD" > "Nat.le_add1"
   "LESS_EQ_0" > "Nat.le_0_eq"
   "LESS_EQUAL_ANTISYM" > "Nat.le_antisym"
-  "LESS_EQUAL_ADD" > "HOL4Base.arithmetic.LESS_EQUAL_ADD"
-  "LESS_EQ" > "Nat.le_simps_3"
+  "LESS_EQUAL_ADD" > "Series.le_Suc_ex"
+  "LESS_EQ" > "Nat.Suc_le_eq"
   "LESS_DIV_EQ_ZERO" > "Divides.div_less"
   "LESS_CASES_IMP" > "HOL4Base.arithmetic.LESS_CASES_IMP"
   "LESS_CASES" > "HOL4Base.arithmetic.LESS_CASES"
@@ -203,8 +203,8 @@
   "LESS_ADD_1" > "HOL4Base.arithmetic.LESS_ADD_1"
   "LESS_ADD" > "HOL4Base.arithmetic.LESS_ADD"
   "LESS_0_CASES" > "HOL4Base.arithmetic.LESS_0_CASES"
-  "LEFT_SUB_DISTRIB" > "Nat.nat_distrib_4"
-  "LEFT_ADD_DISTRIB" > "Nat.nat_distrib_2"
+  "LEFT_SUB_DISTRIB" > "Nat.diff_mult_distrib2"
+  "LEFT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_25"
   "LE" > "HOL4Base.arithmetic.LE"
   "INV_PRE_LESS_EQ" > "HOL4Base.arithmetic.INV_PRE_LESS_EQ"
   "INV_PRE_LESS" > "HOL4Base.arithmetic.INV_PRE_LESS"
@@ -216,11 +216,11 @@
   "FUNPOW" > "HOL4Compat.FUNPOW"
   "FACT_LESS" > "HOL4Base.arithmetic.FACT_LESS"
   "FACT" > "HOL4Base.arithmetic.FACT"
-  "EXP_INJECTIVE" > "Power.power_inject_exp"
-  "EXP_EQ_1" > "HOL4Base.arithmetic.EXP_EQ_1"
+  "EXP_INJECTIVE" > "Power.linordered_semidom_class.power_inject_exp"
+  "EXP_EQ_1" > "Primes.exp_eq_1"
   "EXP_EQ_0" > "HOL4Base.arithmetic.EXP_EQ_0"
   "EXP_ALWAYS_BIG_ENOUGH" > "HOL4Base.arithmetic.EXP_ALWAYS_BIG_ENOUGH"
-  "EXP_ADD" > "Power.power_add"
+  "EXP_ADD" > "Power.monoid_mult_class.power_add"
   "EXP_1" > "HOL4Base.arithmetic.EXP_1"
   "EXP" > "HOL4Compat.EXP"
   "EXISTS_GREATEST" > "HOL4Base.arithmetic.EXISTS_GREATEST"
@@ -233,11 +233,11 @@
   "EVEN_AND_ODD" > "HOL4Base.arithmetic.EVEN_AND_ODD"
   "EVEN_ADD" > "HOL4Base.arithmetic.EVEN_ADD"
   "EVEN" > "HOL4Base.arithmetic.EVEN"
-  "EQ_MULT_LCANCEL" > "Nat_Numeral.nat_mult_eq_cancel_disj"
-  "EQ_MONO_ADD_EQ" > "Nat.nat_add_right_cancel"
-  "EQ_LESS_EQ" > "Orderings.order_eq_iff"
-  "EQ_ADD_RCANCEL" > "Nat.nat_add_right_cancel"
-  "EQ_ADD_LCANCEL" > "Nat.nat_add_left_cancel"
+  "EQ_MULT_LCANCEL" > "Numeral_Simprocs.nat_mult_eq_cancel_disj"
+  "EQ_MONO_ADD_EQ" > "Groups.cancel_semigroup_add_class.add_right_cancel"
+  "EQ_LESS_EQ" > "Orderings.order_class.eq_iff"
+  "EQ_ADD_RCANCEL" > "Groups.cancel_semigroup_add_class.add_right_cancel"
+  "EQ_ADD_LCANCEL" > "Groups.cancel_semigroup_add_class.add_left_cancel"
   "DIV_UNIQUE" > "HOL4Base.arithmetic.DIV_UNIQUE"
   "DIV_P" > "HOL4Base.arithmetic.DIV_P"
   "DIV_ONE" > "Divides.div_1"
@@ -248,23 +248,23 @@
   "DIVMOD_ID" > "HOL4Base.arithmetic.DIVMOD_ID"
   "DIVISION" > "HOL4Compat.DIVISION"
   "DA" > "HOL4Base.arithmetic.DA"
-  "COMPLETE_INDUCTION" > "Nat.less_induct"
+  "COMPLETE_INDUCTION" > "Nat.nat_less_induct"
   "CANCEL_SUB" > "Nat.eq_diff_iff"
   "ALT_ZERO" > "HOL4Compat.ALT_ZERO_def"
-  "ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
+  "ADD_SYM" > "Fields.linordered_field_class.sign_simps_43"
   "ADD_SUC" > "Nat.add_Suc_right"
   "ADD_SUB" > "Nat.diff_add_inverse2"
-  "ADD_MONO_LESS_EQ" > "Nat.nat_add_left_cancel_le"
+  "ADD_MONO_LESS_EQ" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left"
   "ADD_INV_0_EQ" > "HOL4Base.arithmetic.ADD_INV_0_EQ"
   "ADD_INV_0" > "Nat.add_eq_self_zero"
   "ADD_EQ_SUB" > "HOL4Base.arithmetic.ADD_EQ_SUB"
   "ADD_EQ_1" > "HOL4Base.arithmetic.ADD_EQ_1"
   "ADD_EQ_0" > "Nat.add_is_0"
   "ADD_DIV_ADD_DIV" > "HOL4Base.arithmetic.ADD_DIV_ADD_DIV"
-  "ADD_COMM" > "Finite_Set.AC_add.f.AC_2"
+  "ADD_COMM" > "Fields.linordered_field_class.sign_simps_43"
   "ADD_CLAUSES" > "HOL4Base.arithmetic.ADD_CLAUSES"
-  "ADD_ASSOC" > "Finite_Set.AC_add.f.AC_1"
-  "ADD_0" > "Finite_Set.AC_add.f_e.ident"
+  "ADD_ASSOC" > "Fields.linordered_field_class.sign_simps_44"
+  "ADD_0" > "Divides.arithmetic_simps_39"
   "ADD1" > "Nat_Numeral.Suc_eq_plus1"
   "ADD" > "HOL4Compat.ADD"
 
--- a/src/HOL/Import/HOL/bits.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/bits.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -83,7 +83,7 @@
   "DIVMOD_2EXP_def" > "HOL4Word32.bits.DIVMOD_2EXP_def"
   "DIV2_primdef" > "HOL4Word32.bits.DIV2_primdef"
   "DIV2_def" > "HOL4Word32.bits.DIV2_def"
-  "DIV1" > "HOL4Word32.bits.DIV1"
+  "DIV1" > "Divides.semiring_div_class.div_by_1"
   "BIT_def" > "HOL4Word32.bits.BIT_def"
   "BIT_SLICE_THM" > "HOL4Word32.bits.BIT_SLICE_THM"
   "BIT_SLICE_LEM" > "HOL4Word32.bits.BIT_SLICE_LEM"
--- a/src/HOL/Import/HOL/bool.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/bool.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -11,10 +11,13 @@
   "IN" > "IN_def"
   "ARB" > "ARB_def"
 
+type_maps
+  "bool" > "HOL.bool"
+
 const_maps
   "~" > "HOL.Not"
-  "bool_case" > "Datatype.bool.bool_case"
-  "\\/" > HOL.disj
+  "bool_case" > "Product_Type.bool.bool_case"
+  "\\/" > "HOL.disj"
   "TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
   "T" > "HOL.True"
   "RES_SELECT" > "HOL4Base.bool.RES_SELECT"
@@ -30,14 +33,14 @@
   "ARB" > "HOL4Base.bool.ARB"
   "?!" > "HOL.Ex1"
   "?" > "HOL.Ex"
-  "/\\" > HOL.conj
+  "/\\" > "HOL.conj"
   "!" > "HOL.All"
 
 thm_maps
   "bool_case_thm" > "HOL4Base.bool.bool_case_thm"
   "bool_case_ID" > "HOL4Base.bool.bool_case_ID"
   "bool_case_DEF" > "HOL4Compat.bool_case_DEF"
-  "bool_INDUCT" > "Datatype.bool.induct_correctness"
+  "bool_INDUCT" > "Product_Type.bool.induct"
   "boolAxiom" > "HOL4Base.bool.boolAxiom"
   "UNWIND_THM2" > "HOL.simp_thms_39"
   "UNWIND_THM1" > "HOL.simp_thms_40"
@@ -49,21 +52,21 @@
   "TYPE_DEFINITION_THM" > "HOL4Setup.TYPE_DEFINITION"
   "TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
   "TRUTH" > "HOL.TrueI"
-  "SWAP_FORALL_THM" > "HOL4Base.bool.SWAP_FORALL_THM"
-  "SWAP_EXISTS_THM" > "HOL4Base.bool.SWAP_EXISTS_THM"
+  "SWAP_FORALL_THM" > "HOL.all_comm"
+  "SWAP_EXISTS_THM" > "HOL.ex_comm"
   "SKOLEM_THM" > "HOL4Base.bool.SKOLEM_THM"
   "SELECT_UNIQUE" > "HOL4Base.bool.SELECT_UNIQUE"
   "SELECT_THM" > "HOL4Setup.EXISTS_DEF"
   "SELECT_REFL_2" > "Hilbert_Choice.some_sym_eq_trivial"
   "SELECT_REFL" > "Hilbert_Choice.some_eq_trivial"
-  "SELECT_AX" > "Hilbert_Choice.tfl_some"
+  "SELECT_AX" > "Hilbert_Choice.someI"
   "RIGHT_OR_OVER_AND" > "HOL.disj_conj_distribR"
   "RIGHT_OR_EXISTS_THM" > "HOL.ex_simps_4"
   "RIGHT_FORALL_OR_THM" > "HOL.all_simps_4"
   "RIGHT_FORALL_IMP_THM" > "HOL.all_simps_6"
   "RIGHT_EXISTS_IMP_THM" > "HOL.ex_simps_6"
   "RIGHT_EXISTS_AND_THM" > "HOL.ex_simps_2"
-  "RIGHT_AND_OVER_OR" > "HOL.conj_disj_distribR"
+  "RIGHT_AND_OVER_OR" > "Groebner_Basis.dnf_2"
   "RIGHT_AND_FORALL_THM" > "HOL.all_simps_2"
   "RES_SELECT_def" > "HOL4Base.bool.RES_SELECT_def"
   "RES_SELECT_DEF" > "HOL4Base.bool.RES_SELECT_DEF"
@@ -74,11 +77,11 @@
   "RES_EXISTS_UNIQUE_DEF" > "HOL4Base.bool.RES_EXISTS_UNIQUE_DEF"
   "RES_EXISTS_DEF" > "HOL4Base.bool.RES_EXISTS_DEF"
   "RES_ABSTRACT_DEF" > "HOL4Base.bool.RES_ABSTRACT_DEF"
-  "REFL_CLAUSE" > "HOL.simp_thms_6"
+  "REFL_CLAUSE" > "Groebner_Basis.bool_simps_6"
   "OR_INTRO_THM2" > "HOL.disjI2"
   "OR_INTRO_THM1" > "HOL.disjI1"
   "OR_IMP_THM" > "HOL4Base.bool.OR_IMP_THM"
-  "OR_ELIM_THM" > "Recdef.tfl_disjE"
+  "OR_ELIM_THM" > "HOL.disjE"
   "OR_DEF" > "HOL.or_def"
   "OR_CONG" > "HOL4Base.bool.OR_CONG"
   "OR_CLAUSES" > "HOL4Base.bool.OR_CLAUSES"
@@ -87,14 +90,14 @@
   "ONE_ONE_THM" > "HOL4Base.bool.ONE_ONE_THM"
   "ONE_ONE_DEF" > "HOL4Setup.ONE_ONE_DEF"
   "NOT_IMP" > "HOL.not_imp"
-  "NOT_FORALL_THM" > "Inductive.basic_monos_15"
-  "NOT_F" > "HOL.Eq_FalseI"
-  "NOT_EXISTS_THM" > "Inductive.basic_monos_16"
-  "NOT_DEF" > "HOL.simp_thms_19"
+  "NOT_FORALL_THM" > "HOL.not_all"
+  "NOT_F" > "Groebner_Basis.PFalse_2"
+  "NOT_EXISTS_THM" > "HOL.not_ex"
+  "NOT_DEF" > "Groebner_Basis.bool_simps_19"
   "NOT_CLAUSES" > "HOL4Base.bool.NOT_CLAUSES"
   "NOT_AND" > "HOL4Base.bool.NOT_AND"
   "MONO_OR" > "Inductive.basic_monos_3"
-  "MONO_NOT" > "HOL.rev_contrapos"
+  "MONO_NOT" > "HOL.contrapos_nn"
   "MONO_IMP" > "Set.imp_mono"
   "MONO_EXISTS" > "Inductive.basic_monos_5"
   "MONO_COND" > "HOL4Base.bool.MONO_COND"
@@ -104,27 +107,27 @@
   "LET_RATOR" > "HOL4Base.bool.LET_RATOR"
   "LET_RAND" > "HOL4Base.bool.LET_RAND"
   "LET_DEF" > "HOL4Compat.LET_def"
-  "LET_CONG" > "Recdef.let_cong"
+  "LET_CONG" > "FunDef.let_cong"
   "LEFT_OR_OVER_AND" > "HOL.disj_conj_distribL"
   "LEFT_OR_EXISTS_THM" > "HOL.ex_simps_3"
   "LEFT_FORALL_OR_THM" > "HOL.all_simps_3"
-  "LEFT_FORALL_IMP_THM" > "HOL.imp_ex"
-  "LEFT_EXISTS_IMP_THM" > "HOL.imp_all"
+  "LEFT_FORALL_IMP_THM" > "HOL.all_simps_5"
+  "LEFT_EXISTS_IMP_THM" > "HOL.ex_simps_5"
   "LEFT_EXISTS_AND_THM" > "HOL.ex_simps_1"
-  "LEFT_AND_OVER_OR" > "HOL.conj_disj_distribL"
+  "LEFT_AND_OVER_OR" > "Groebner_Basis.dnf_1"
   "LEFT_AND_FORALL_THM" > "HOL.all_simps_1"
   "IN_def" > "HOL4Base.bool.IN_def"
   "IN_DEF" > "HOL4Base.bool.IN_DEF"
   "INFINITY_AX" > "HOL4Setup.INFINITY_AX"
   "IMP_F_EQ_F" > "HOL4Base.bool.IMP_F_EQ_F"
   "IMP_F" > "HOL.notI"
-  "IMP_DISJ_THM" > "Inductive.basic_monos_11"
+  "IMP_DISJ_THM" > "Groebner_Basis.nnf_simps_3"
   "IMP_CONG" > "HOL.imp_cong"
   "IMP_CLAUSES" > "HOL4Base.bool.IMP_CLAUSES"
-  "IMP_ANTISYM_AX" > "HOL4Setup.light_imp_as"
+  "IMP_ANTISYM_AX" > "HOL.iff"
   "F_IMP" > "HOL4Base.bool.F_IMP"
   "F_DEF" > "HOL.False_def"
-  "FUN_EQ_THM" > "Fun.fun_eq_iff"
+  "FUN_EQ_THM" > "HOL.fun_eq_iff"
   "FORALL_THM" > "HOL4Base.bool.FORALL_THM"
   "FORALL_SIMP" > "HOL.simp_thms_35"
   "FORALL_DEF" > "HOL.All_def"
@@ -139,24 +142,24 @@
   "EXISTS_OR_THM" > "HOL.ex_disj_distrib"
   "EXISTS_DEF" > "HOL4Setup.EXISTS_DEF"
   "EXCLUDED_MIDDLE" > "HOL4Base.bool.EXCLUDED_MIDDLE"
-  "ETA_THM" > "Presburger.fm_modd_pinf"
-  "ETA_AX" > "Presburger.fm_modd_pinf"
-  "EQ_TRANS" > "Set.basic_trans_rules_31"
-  "EQ_SYM_EQ" > "HOL.eq_sym_conv"
-  "EQ_SYM" > "HOL.meta_eq_to_obj_eq"
-  "EQ_REFL" > "Presburger.fm_modd_pinf"
+  "ETA_THM" > "HOL.eta_contract_eq"
+  "ETA_AX" > "HOL.eta_contract_eq"
+  "EQ_TRANS" > "HOL.trans"
+  "EQ_SYM_EQ" > "HOL.eq_ac_1"
+  "EQ_SYM" > "HOL.eq_reflection"
+  "EQ_REFL" > "HOL.refl"
   "EQ_IMP_THM" > "HOL.iff_conv_conj_imp"
-  "EQ_EXT" > "HOL.meta_eq_to_obj_eq"
-  "EQ_EXPAND" > "HOL4Base.bool.EQ_EXPAND"
+  "EQ_EXT" > "HOL.eq_reflection"
+  "EQ_EXPAND" > "Groebner_Basis.nnf_simps_4"
   "EQ_CLAUSES" > "HOL4Base.bool.EQ_CLAUSES"
-  "DISJ_SYM" > "HOL.disj_comms_1"
+  "DISJ_SYM" > "Groebner_Basis.dnf_4"
   "DISJ_IMP_THM" > "HOL.imp_disjL"
-  "DISJ_COMM" > "HOL.disj_comms_1"
-  "DISJ_ASSOC" > "Recdef.tfl_disj_assoc"
+  "DISJ_COMM" > "Groebner_Basis.dnf_4"
+  "DISJ_ASSOC" > "HOL.disj_ac_3"
   "DE_MORGAN_THM" > "HOL4Base.bool.DE_MORGAN_THM"
-  "CONJ_SYM" > "HOL.conj_comms_1"
-  "CONJ_COMM" > "HOL.conj_comms_1"
-  "CONJ_ASSOC" > "HOL.conj_assoc"
+  "CONJ_SYM" > "Groebner_Basis.dnf_3"
+  "CONJ_COMM" > "Groebner_Basis.dnf_3"
+  "CONJ_ASSOC" > "HOL.conj_ac_3"
   "COND_RATOR" > "HOL4Base.bool.COND_RATOR"
   "COND_RAND" > "HOL.if_distrib"
   "COND_ID" > "HOL.if_cancel"
@@ -172,8 +175,8 @@
   "BOOL_FUN_INDUCT" > "HOL4Base.bool.BOOL_FUN_INDUCT"
   "BOOL_FUN_CASES_THM" > "HOL4Base.bool.BOOL_FUN_CASES_THM"
   "BOOL_EQ_DISTINCT" > "HOL4Base.bool.BOOL_EQ_DISTINCT"
-  "BOOL_CASES_AX" > "Datatype.bool.nchotomy"
-  "BETA_THM" > "Presburger.fm_modd_pinf"
+  "BOOL_CASES_AX" > "HOL.True_or_False"
+  "BETA_THM" > "HOL.eta_contract_eq"
   "ARB_def" > "HOL4Base.bool.ARB_def"
   "ARB_DEF" > "HOL4Base.bool.ARB_DEF"
   "AND_INTRO_THM" > "HOL.conjI"
@@ -183,7 +186,7 @@
   "AND_CLAUSES" > "HOL4Base.bool.AND_CLAUSES"
   "AND2_THM" > "HOL.conjunct2"
   "AND1_THM" > "HOL.conjunct1"
-  "ABS_SIMP" > "Presburger.fm_modd_pinf"
+  "ABS_SIMP" > "HOL.refl"
   "ABS_REP_THM" > "HOL4Base.bool.ABS_REP_THM"
 
 ignore_thms
--- a/src/HOL/Import/HOL/combin.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/combin.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -18,8 +18,8 @@
   "C" > "HOL4Base.combin.C"
 
 thm_maps
-  "o_THM" > "Fun.o_apply"
-  "o_DEF" > "Fun.o_apply"
+  "o_THM" > "Fun.comp_def"
+  "o_DEF" > "Fun.comp_def"
   "o_ASSOC" > "Fun.o_assoc"
   "W_def" > "HOL4Base.combin.W_def"
   "W_THM" > "HOL4Base.combin.W_def"
--- a/src/HOL/Import/HOL/divides.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/divides.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -3,22 +3,22 @@
 import_segment "hol4"
 
 const_maps
-  "divides" > Divides.times_class.dvd :: "nat => nat => bool"
+  "divides" > "Rings.dvd_class.dvd" :: "nat => nat => bool"
 
 thm_maps
   "divides_def" > "HOL4Compat.divides_def"
-  "ONE_DIVIDES_ALL" > "HOL4Base.divides.ONE_DIVIDES_ALL"
-  "NOT_LT_DIV" > "NatSimprocs.nat_dvd_not_less"
-  "DIVIDES_TRANS" > "Rings.dvd_trans"
-  "DIVIDES_SUB" > "Rings.dvd_diff"
-  "DIVIDES_REFL" > "Rings.dvd_refl"
+  "ONE_DIVIDES_ALL" > "GCD.gcd_lcm_complete_lattice_nat.bot_least"
+  "NOT_LT_DIV" > "Nat.nat_dvd_not_less"
+  "DIVIDES_TRANS" > "Nat.dvd.order_trans"
+  "DIVIDES_SUB" > "Nat.dvd_diff_nat"
+  "DIVIDES_REFL" > "Nat.dvd.order_refl"
   "DIVIDES_MULT_LEFT" > "HOL4Base.divides.DIVIDES_MULT_LEFT"
-  "DIVIDES_MULT" > "Divides.dvd_mult2"
-  "DIVIDES_LE" > "Divides.dvd_imp_le"
+  "DIVIDES_MULT" > "Rings.comm_semiring_1_class.dvd_mult2"
+  "DIVIDES_LE" > "Nat.dvd_imp_le"
   "DIVIDES_FACT" > "HOL4Base.divides.DIVIDES_FACT"
-  "DIVIDES_ANTISYM" > "Divides.dvd_antisym"
-  "DIVIDES_ADD_2" > "HOL4Base.divides.DIVIDES_ADD_2"
-  "DIVIDES_ADD_1" > "Rings.dvd_add"
-  "ALL_DIVIDES_0" > "Divides.dvd_0_right"
+  "DIVIDES_ANTISYM" > "Nat.dvd.antisym"
+  "DIVIDES_ADD_2" > "Primes.divides_add_revr"
+  "DIVIDES_ADD_1" > "Rings.comm_semiring_1_class.dvd_add"
+  "ALL_DIVIDES_0" > "GCD.gcd_lcm_complete_lattice_nat.top_greatest"
 
 end
--- a/src/HOL/Import/HOL/lim.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/lim.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -24,7 +24,7 @@
   "contl_def" > "HOL4Real.lim.contl_def"
   "contl" > "HOL4Real.lim.contl"
   "ROLLE" > "HOL4Real.lim.ROLLE"
-  "MVT_LEMMA" > "HOL4Real.lim.MVT_LEMMA"
+  "MVT_LEMMA" > "Deriv.lemma_MVT"
   "MVT" > "HOL4Real.lim.MVT"
   "LIM_X" > "HOL4Real.lim.LIM_X"
   "LIM_UNIQ" > "HOL4Real.lim.LIM_UNIQ"
@@ -41,7 +41,7 @@
   "LIM" > "HOL4Real.lim.LIM"
   "IVT2" > "HOL4Real.lim.IVT2"
   "IVT" > "HOL4Real.lim.IVT"
-  "INTERVAL_LEMMA" > "HOL4Real.lim.INTERVAL_LEMMA"
+  "INTERVAL_LEMMA" > "Deriv.lemma_interval"
   "INTERVAL_CLEMMA" > "HOL4Real.lim.INTERVAL_CLEMMA"
   "INTERVAL_ABS" > "HOL4Real.lim.INTERVAL_ABS"
   "DIFF_XM1" > "HOL4Real.lim.DIFF_XM1"
--- a/src/HOL/Import/HOL/list.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/list.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -19,7 +19,7 @@
   "REPLICATE" > "List.replicate"
   "NULL" > "List.null"
   "NIL" > "List.list.Nil"
-  "MEM" > "List.op mem"
+  "MEM" > "HOL4Compat.list_mem"
   "MAP2" > "HOL4Compat.map2"
   "MAP" > "List.map"
   "LENGTH" > "Nat.size_class.size"
@@ -30,25 +30,25 @@
   "FOLDL" > "List.foldl"
   "FLAT" > "List.concat"
   "FILTER" > "List.filter"
-  "EXISTS" > "HOL4Compat.list_exists"
+  "EXISTS" > "List.list_ex"
   "EVERY" > "List.list_all"
   "CONS" > "List.list.Cons"
   "APPEND" > "List.append"
 
 thm_maps
-  "list_size_def" > "HOL4Compat.list_size_def"
+  "list_size_def" > "HOL4Compat.list_size_def'"
   "list_size_cong" > "HOL4Base.list.list_size_cong"
   "list_nchotomy" > "HOL4Compat.list_CASES"
-  "list_induction" > "HOL4Compat.unzip.induct"
-  "list_distinct" > "List.list.simps_1"
+  "list_induction" > "HOL4Compat.list_INDUCT"
+  "list_distinct" > "List.list.distinct_1"
   "list_case_def" > "HOL4Compat.list_case_def"
   "list_case_cong" > "HOL4Compat.list_case_cong"
   "list_case_compute" > "HOL4Base.list.list_case_compute"
-  "list_INDUCT" > "HOL4Compat.unzip.induct"
+  "list_INDUCT" > "HOL4Compat.list_INDUCT"
   "list_CASES" > "HOL4Compat.list_CASES"
   "list_Axiom_old" > "HOL4Compat.list_Axiom_old"
   "list_Axiom" > "HOL4Compat.list_Axiom"
-  "list_11" > "List.list.simps_3"
+  "list_11" > "List.list.inject"
   "ZIP_UNZIP" > "HOL4Base.list.ZIP_UNZIP"
   "ZIP_MAP" > "HOL4Base.list.ZIP_MAP"
   "ZIP" > "HOL4Compat.ZIP"
@@ -62,11 +62,11 @@
   "REVERSE_APPEND" > "List.rev_append"
   "NULL_DEF" > "HOL4Compat.NULL_DEF"
   "NULL" > "HOL4Base.list.NULL"
-  "NOT_NIL_CONS" > "List.list.simps_1"
+  "NOT_NIL_CONS" > "List.list.distinct_1"
   "NOT_EXISTS" > "HOL4Base.list.NOT_EXISTS"
   "NOT_EVERY" > "HOL4Base.list.NOT_EVERY"
   "NOT_EQ_LIST" > "HOL4Base.list.NOT_EQ_LIST"
-  "NOT_CONS_NIL" > "List.list.simps_2"
+  "NOT_CONS_NIL" > "List.list.distinct_2"
   "MEM_ZIP" > "HOL4Base.list.MEM_ZIP"
   "MEM_MAP" > "HOL4Base.list.MEM_MAP"
   "MEM_EL" > "HOL4Base.list.MEM_EL"
@@ -102,7 +102,7 @@
   "EXISTS_MEM" > "HOL4Base.list.EXISTS_MEM"
   "EXISTS_DEF" > "HOL4Compat.list_exists_DEF"
   "EXISTS_CONG" > "HOL4Base.list.EXISTS_CONG"
-  "EXISTS_APPEND" > "HOL4Base.list.EXISTS_APPEND"
+  "EXISTS_APPEND" > "List.list_ex_append"
   "EVERY_MONOTONIC" > "HOL4Base.list.EVERY_MONOTONIC"
   "EVERY_MEM" > "HOL4Base.list.EVERY_MEM"
   "EVERY_EL" > "HOL4Base.list.EVERY_EL"
@@ -115,7 +115,7 @@
   "EL_ZIP" > "HOL4Base.list.EL_ZIP"
   "EL" > "HOL4Base.list.EL"
   "CONS_ACYCLIC" > "HOL4Base.list.CONS_ACYCLIC"
-  "CONS_11" > "List.list.simps_3"
+  "CONS_11" > "List.list.inject"
   "CONS" > "HOL4Base.list.CONS"
   "APPEND_eq_NIL" > "HOL4Base.list.APPEND_eq_NIL"
   "APPEND_NIL" > "List.append_Nil2"
--- a/src/HOL/Import/HOL/num.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/num.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -7,12 +7,12 @@
 
 const_maps
   "SUC" > "Nat.Suc"
-  "0" > "0" :: "nat"
+  "0" > "Groups.zero_class.zero" :: "nat"
 
 thm_maps
-  "NOT_SUC" > "Nat.nat.simps_1"
+  "NOT_SUC" > "Nat.Suc_not_Zero"
   "INV_SUC" > "Nat.Suc_inject"
-  "INDUCTION" > "List.lexn.induct"
+  "INDUCTION" > "Fact.fact_nat.induct"
 
 ignore_thms
   "num_TY_DEF"
--- a/src/HOL/Import/HOL/option.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/option.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -3,34 +3,34 @@
 import_segment "hol4"
 
 type_maps
-  "option" > "Datatype.option"
+  "option" > "Option.option"
 
 const_maps
-  "option_case" > "Datatype.option.option_case"
-  "THE" > "Datatype.the"
-  "SOME" > "Datatype.option.Some"
-  "OPTION_MAP" > "Datatype.option_map"
+  "option_case" > "Option.option.option_case"
+  "THE" > "Option.the"
+  "SOME" > "Option.option.Some"
+  "OPTION_MAP" > "Option.map"
   "OPTION_JOIN" > "HOL4Compat.OPTION_JOIN"
-  "NONE" > "Datatype.option.None"
+  "NONE" > "Option.option.None"
   "IS_SOME" > "HOL4Compat.IS_SOME"
   "IS_NONE" > "HOL4Compat.IS_NONE"
 
 thm_maps
-  "option_nchotomy" > "Datatype.option.nchotomy"
-  "option_induction" > "HOL4Compat.OPTION_JOIN.induct"
+  "option_nchotomy" > "Option.option.nchotomy"
+  "option_induction" > "Option.option.induct"
   "option_case_def" > "HOL4Compat.option_case_def"
   "option_case_cong" > "HOL4Base.option.option_case_cong"
   "option_case_compute" > "HOL4Base.option.option_case_compute"
   "option_CLAUSES" > "HOL4Base.option.option_CLAUSES"
-  "THE_DEF" > "Datatype.the.simps"
-  "SOME_11" > "Datatype.option.simps_3"
+  "THE_DEF" > "Option.the.simps"
+  "SOME_11" > "Option.option.inject"
   "OPTION_MAP_EQ_SOME" > "HOL4Base.option.OPTION_MAP_EQ_SOME"
-  "OPTION_MAP_EQ_NONE" > "Datatype.option_map_is_None"
+  "OPTION_MAP_EQ_NONE" > "Option.option_map_is_None"
   "OPTION_MAP_DEF" > "HOL4Compat.OPTION_MAP_DEF"
   "OPTION_JOIN_EQ_SOME" > "HOL4Base.option.OPTION_JOIN_EQ_SOME"
   "OPTION_JOIN_DEF" > "HOL4Compat.OPTION_JOIN_DEF"
-  "NOT_SOME_NONE" > "Datatype.option.simps_2"
-  "NOT_NONE_SOME" > "Datatype.option.simps_1"
+  "NOT_SOME_NONE" > "Option.option.distinct_2"
+  "NOT_NONE_SOME" > "Option.option.distinct_1"
   "IS_SOME_DEF" > "HOL4Compat.IS_SOME_DEF"
   "IS_NONE_DEF" > "HOL4Compat.IS_NONE_DEF"
 
--- a/src/HOL/Import/HOL/pair.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/pair.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -10,38 +10,38 @@
   "prod" > "Product_Type.prod"
 
 const_maps
-  "pair_case" > "Product_Type.prod_case"
-  "UNCURRY" > "Product_Type.prod_case"
-  "FST" > "Product_Type.fst"
+  "pair_case" > "Product_Type.prod.prod_case"
+  "UNCURRY" > "Product_Type.prod.prod_case"
   "SND" > "Product_Type.snd"
   "RPROD" > "HOL4Base.pair.RPROD"
   "LEX" > "HOL4Base.pair.LEX"
+  "FST" > "Product_Type.fst"
   "CURRY" > "Product_Type.curry"
   "," > "Product_Type.Pair"
-  "##" > "map_pair"
+  "##" > "Product_Type.map_pair"
 
 thm_maps
-  "pair_induction" > "Datatype.prod.induct_correctness"
-  "pair_case_thm" > "Product_Type.split"
+  "pair_induction" > "Product_Type.prod.induct"
+  "pair_case_thm" > "Product_Type.prod.cases"
   "pair_case_def" > "HOL4Compat.pair_case_def"
   "pair_case_cong" > "HOL4Base.pair.pair_case_cong"
   "pair_Axiom" > "HOL4Base.pair.pair_Axiom"
   "WF_RPROD" > "HOL4Base.pair.WF_RPROD"
   "WF_LEX" > "HOL4Base.pair.WF_LEX"
-  "UNCURRY_VAR" > "Product_Type.split_beta"
+  "UNCURRY_VAR" > "Product_Type.prod_case_beta"
   "UNCURRY_ONE_ONE_THM" > "HOL4Base.pair.UNCURRY_ONE_ONE_THM"
-  "UNCURRY_DEF" > "Product_Type.split"
+  "UNCURRY_DEF" > "Product_Type.prod.cases"
   "UNCURRY_CURRY_THM" > "Product_Type.split_curry"
   "UNCURRY_CONG" > "HOL4Base.pair.UNCURRY_CONG"
-  "UNCURRY" > "Product_Type.split_beta"
+  "UNCURRY" > "Product_Type.prod_case_beta"
   "SND" > "Product_Type.snd_conv"
   "RPROD_def" > "HOL4Base.pair.RPROD_def"
   "RPROD_DEF" > "HOL4Base.pair.RPROD_DEF"
   "PFORALL_THM" > "HOL4Base.pair.PFORALL_THM"
   "PEXISTS_THM" > "HOL4Base.pair.PEXISTS_THM"
-  "PAIR_MAP_THM" > "Product_Type.map_pair"
+  "PAIR_MAP_THM" > "Product_Type.map_pair_simp"
   "PAIR_MAP" > "HOL4Compat.PAIR_MAP"
-  "PAIR_EQ" > "Datatype.prod.simps_1"
+  "PAIR_EQ" > "Product_Type.Pair_eq"
   "PAIR" > "HOL4Compat.PAIR"
   "LEX_def" > "HOL4Base.pair.LEX_def"
   "LEX_DEF" > "HOL4Base.pair.LEX_DEF"
@@ -51,14 +51,14 @@
   "FST" > "Product_Type.fst_conv"
   "FORALL_PROD" > "Product_Type.split_paired_All"
   "EXISTS_PROD" > "Product_Type.split_paired_Ex"
-  "ELIM_UNCURRY" > "Product_Type.split_beta"
+  "ELIM_UNCURRY" > "Product_Type.prod_case_beta"
   "ELIM_PFORALL" > "HOL4Base.pair.ELIM_PFORALL"
   "ELIM_PEXISTS" > "HOL4Base.pair.ELIM_PEXISTS"
   "CURRY_UNCURRY_THM" > "Product_Type.curry_split"
   "CURRY_ONE_ONE_THM" > "HOL4Base.pair.CURRY_ONE_ONE_THM"
   "CURRY_DEF" > "Product_Type.curry_conv"
-  "CLOSED_PAIR_EQ" > "Datatype.prod.simps_1"
-  "ABS_PAIR_THM" > "Datatype.prod.nchotomy"
+  "CLOSED_PAIR_EQ" > "Product_Type.Pair_eq"
+  "ABS_PAIR_THM" > "Product_Type.prod.nchotomy"
 
 ignore_thms
   "prod_TY_DEF"
--- a/src/HOL/Import/HOL/poly.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/poly.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -13,7 +13,7 @@
   "poly_add" > "poly_add_primdef"
   "poly" > "poly_primdef"
   "normalize" > "normalize_def"
-  "diff" > "diff_minus"
+  "diff" > "diff_def"
   "degree" > "degree_def"
   "##" > "##_def"
 
--- a/src/HOL/Import/HOL/prim_rec.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/prim_rec.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -18,7 +18,7 @@
   "PRIM_REC_FUN" > "HOL4Base.prim_rec.PRIM_REC_FUN"
   "PRIM_REC" > "HOL4Base.prim_rec.PRIM_REC"
   "PRE" > "HOL4Base.prim_rec.PRE"
-  "<" > "Orderings.less" :: "nat => nat => bool"
+  "<" > "Orderings.ord_class.less" :: "nat => nat => bool"
 
 thm_maps
   "wellfounded_primdef" > "HOL4Base.prim_rec.wellfounded_primdef"
@@ -65,7 +65,7 @@
   "LESS_DEF" > "HOL4Compat.LESS_DEF"
   "LESS_0_0" > "HOL4Base.prim_rec.LESS_0_0"
   "LESS_0" > "Nat.zero_less_Suc"
-  "INV_SUC_EQ" > "Nat.nat.simps_3"
+  "INV_SUC_EQ" > "Nat.nat.inject"
   "EQ_LESS" > "HOL4Base.prim_rec.EQ_LESS"
   "DC" > "HOL4Base.prim_rec.DC"
 
--- a/src/HOL/Import/HOL/prob_extra.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/prob_extra.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -23,9 +23,9 @@
   "REAL_SUP_LE_X" > "HOL4Prob.prob_extra.REAL_SUP_LE_X"
   "REAL_SUP_EXISTS_UNIQUE" > "HOL4Prob.prob_extra.REAL_SUP_EXISTS_UNIQUE"
   "REAL_POW" > "RealDef.power_real_of_nat"
-  "REAL_LE_INV_LE" > "Rings.le_imp_inverse_le"
-  "REAL_LE_EQ" > "Set.basic_trans_rules_26"
-  "REAL_INVINV_ALL" > "Rings.inverse_inverse_eq"
+  "REAL_LE_INV_LE" > "Fields.linordered_field_class.le_imp_inverse_le"
+  "REAL_LE_EQ" > "Orderings.order_antisym"
+  "REAL_INVINV_ALL" > "Fields.division_ring_inverse_zero_class.inverse_inverse_eq"
   "REAL_INF_MIN" > "HOL4Prob.prob_extra.REAL_INF_MIN"
   "RAND_THM" > "HOL.arg_cong"
   "POW_HALF_TWICE" > "HOL4Prob.prob_extra.POW_HALF_TWICE"
@@ -59,7 +59,7 @@
   "INTER_UNION_COMPL" > "HOL4Base.pred_set.INTER_UNION_COMPL"
   "INTER_IS_EMPTY" > "HOL4Prob.prob_extra.INTER_IS_EMPTY"
   "INF_DEF_ALT" > "HOL4Prob.prob_extra.INF_DEF_ALT"
-  "HALF_POS" > "HOL4Prob.prob_extra.HALF_POS"
+  "HALF_POS" > "Series.half"
   "HALF_LT_1" > "HOL4Prob.prob_extra.HALF_LT_1"
   "HALF_CANCEL" > "HOL4Prob.prob_extra.HALF_CANCEL"
   "GSPEC_DEF_ALT" > "HOL4Prob.prob_extra.GSPEC_DEF_ALT"
@@ -73,7 +73,7 @@
   "EVEN_ODD_EXISTS_EQ" > "HOL4Prob.prob_extra.EVEN_ODD_EXISTS_EQ"
   "EVEN_ODD_BASIC" > "HOL4Prob.prob_extra.EVEN_ODD_BASIC"
   "EVEN_EXP_TWO" > "HOL4Prob.prob_extra.EVEN_EXP_TWO"
-  "EQ_EXT_EQ" > "Fun.fun_eq_iff"
+  "EQ_EXT_EQ" > "HOL.fun_eq_iff"
   "DIV_TWO_UNIQUE" > "HOL4Prob.prob_extra.DIV_TWO_UNIQUE"
   "DIV_TWO_MONO_EVEN" > "HOL4Prob.prob_extra.DIV_TWO_MONO_EVEN"
   "DIV_TWO_MONO" > "HOL4Prob.prob_extra.DIV_TWO_MONO"
--- a/src/HOL/Import/HOL/real.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/real.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -10,14 +10,14 @@
 const_maps
   "sup" > "HOL4Real.real.sup"
   "sum" > "HOL4Real.real.sum"
-  "real_sub" > "Groups.minus" :: "real => real => real"
+  "real_sub" > "Groups.minus_class.minus" :: "real => real => real"
   "real_of_num" > "RealDef.real" :: "nat => real"
-  "real_lte" > "Orderings.less_eq" :: "real => real => bool"
+  "real_lte" > "Orderings.ord_class.less_eq" :: "real => real => bool"
   "real_gt" > "HOL4Compat.real_gt"
   "real_ge" > "HOL4Compat.real_ge"
   "pow" > "Power.power_class.power" :: "real => nat => real"
-  "abs" > "Groups.abs" :: "real => real"
-  "/" > "Ring.divide" :: "real => real => real"
+  "abs" > "Groups.abs_class.abs" :: "real => real"
+  "/" > "Fields.inverse_class.divide" :: "real => real => real"
 
 thm_maps
   "sup_def" > "HOL4Real.real.sup_def"
@@ -25,13 +25,13 @@
   "sumc" > "HOL4Real.real.sumc"
   "sum_def" > "HOL4Real.real.sum_def"
   "sum" > "HOL4Real.real.sum"
-  "real_sub" > "Groups.diff_minus"
+  "real_sub" > "Fields.linordered_field_class.sign_simps_16"
   "real_of_num" > "HOL4Compat.real_of_num"
   "real_lte" > "HOL4Compat.real_lte"
-  "real_lt" > "Orderings.linorder_not_le"
-  "real_gt" > "HOL4Compat.real_gt"
+  "real_lt" > "Orderings.linorder_class.not_le"
+  "real_gt" > "HOL4Compat.GREATER_DEF"
   "real_ge" > "HOL4Compat.real_ge"
-  "real_div" > "Ring.divide_inverse"
+  "real_div" > "Fields.division_ring_class.divide_inverse"
   "pow" > "HOL4Compat.pow"
   "abs" > "HOL4Compat.abs"
   "SUP_LEMMA3" > "HOL4Real.real.SUP_LEMMA3"
@@ -70,40 +70,40 @@
   "REAL_SUP_EXISTS" > "HOL4Real.real.REAL_SUP_EXISTS"
   "REAL_SUP_ALLPOS" > "HOL4Compat.REAL_SUP_ALLPOS"
   "REAL_SUP" > "HOL4Real.real.REAL_SUP"
-  "REAL_SUMSQ" > "HOL4Real.real.REAL_SUMSQ"
+  "REAL_SUMSQ" > "Nat_Numeral.linordered_ring_strict_class.sum_squares_eq_zero_iff"
   "REAL_SUB_TRIANGLE" > "HOL4Real.real.REAL_SUB_TRIANGLE"
   "REAL_SUB_SUB2" > "HOL4Real.real.REAL_SUB_SUB2"
   "REAL_SUB_SUB" > "HOL4Real.real.REAL_SUB_SUB"
-  "REAL_SUB_RZERO" > "Groups.diff_0_right"
-  "REAL_SUB_RNEG" > "Groups.diff_minus_eq_add"
-  "REAL_SUB_REFL" > "Groups.diff_self"
-  "REAL_SUB_RDISTRIB" > "Rings.ring_eq_simps_3"
+  "REAL_SUB_RZERO" > "Groups.group_add_class.diff_0_right"
+  "REAL_SUB_RNEG" > "Groups.group_add_class.diff_minus_eq_add"
+  "REAL_SUB_REFL" > "Groups.group_add_class.diff_self"
+  "REAL_SUB_RDISTRIB" > "Fields.linordered_field_class.sign_simps_5"
   "REAL_SUB_NEG2" > "HOL4Real.real.REAL_SUB_NEG2"
-  "REAL_SUB_LZERO" > "Groups.diff_0"
+  "REAL_SUB_LZERO" > "Groups.group_add_class.diff_0"
   "REAL_SUB_LT" > "HOL4Real.real.REAL_SUB_LT"
   "REAL_SUB_LNEG" > "HOL4Real.real.REAL_SUB_LNEG"
   "REAL_SUB_LE" > "HOL4Real.real.REAL_SUB_LE"
-  "REAL_SUB_LDISTRIB" > "Rings.ring_eq_simps_4"
+  "REAL_SUB_LDISTRIB" > "Fields.linordered_field_class.sign_simps_6"
   "REAL_SUB_INV2" > "HOL4Real.real.REAL_SUB_INV2"
   "REAL_SUB_ADD2" > "HOL4Real.real.REAL_SUB_ADD2"
-  "REAL_SUB_ADD" > "Groups.diff_add_cancel"
-  "REAL_SUB_ABS" > "Groups.abs_triangle_ineq2"
-  "REAL_SUB_0" > "Groups.diff_eq_0_iff_eq"
-  "REAL_RNEG_UNIQ" > "RealDef.real_add_eq_0_iff"
-  "REAL_RINV_UNIQ" > "Rings.inverse_unique"
-  "REAL_RDISTRIB" > "Rings.ring_eq_simps_1"
-  "REAL_POW_POW" > "Power.power_mult"
+  "REAL_SUB_ADD" > "Groups.group_add_class.diff_add_cancel"
+  "REAL_SUB_ABS" > "Groups.ordered_ab_group_add_abs_class.abs_triangle_ineq2"
+  "REAL_SUB_0" > "Groups.ab_group_add_class.diff_eq_0_iff_eq"
+  "REAL_RNEG_UNIQ" > "Groups.group_add_class.add_eq_0_iff"
+  "REAL_RINV_UNIQ" > "Fields.division_ring_class.inverse_unique"
+  "REAL_RDISTRIB" > "Fields.linordered_field_class.sign_simps_8"
+  "REAL_POW_POW" > "Power.monoid_mult_class.power_mult"
   "REAL_POW_MONO_LT" > "HOL4Real.real.REAL_POW_MONO_LT"
   "REAL_POW_LT2" > "HOL4Real.real.REAL_POW_LT2"
-  "REAL_POW_LT" > "Power.zero_less_power"
+  "REAL_POW_LT" > "Power.linordered_semidom_class.zero_less_power"
   "REAL_POW_INV" > "Power.power_inverse"
   "REAL_POW_DIV" > "Power.power_divide"
-  "REAL_POW_ADD" > "Power.power_add"
-  "REAL_POW2_ABS" > "Nat_Numeral.power2_abs"
+  "REAL_POW_ADD" > "Power.monoid_mult_class.power_add"
+  "REAL_POW2_ABS" > "Nat_Numeral.linordered_idom_class.power2_abs"
   "REAL_POS_NZ" > "HOL4Real.real.REAL_POS_NZ"
   "REAL_POS" > "RealDef.real_of_nat_ge_zero"
   "REAL_POASQ" > "HOL4Real.real.REAL_POASQ"
-  "REAL_OVER1" > "Rings.divide_1"
+  "REAL_OVER1" > "Fields.division_ring_class.divide_1"
   "REAL_OF_NUM_SUC" > "RealDef.real_of_nat_Suc"
   "REAL_OF_NUM_POW" > "RealDef.power_real_of_nat"
   "REAL_OF_NUM_MUL" > "RealDef.real_of_nat_mult"
@@ -112,239 +112,239 @@
   "REAL_OF_NUM_ADD" > "RealDef.real_of_nat_add"
   "REAL_NZ_IMP_LT" > "HOL4Real.real.REAL_NZ_IMP_LT"
   "REAL_NOT_LT" > "HOL4Compat.real_lte"
-  "REAL_NOT_LE" > "Orderings.linorder_not_le"
-  "REAL_NEG_SUB" > "Groups.minus_diff_eq"
-  "REAL_NEG_RMUL" > "Rings.mult_minus_right"
-  "REAL_NEG_NEG" > "Groups.minus_minus"
-  "REAL_NEG_MUL2" > "Rings.minus_mult_minus"
-  "REAL_NEG_MINUS1" > "HOL4Real.real.REAL_NEG_MINUS1"
-  "REAL_NEG_LT0" > "Groups.neg_less_0_iff_less"
-  "REAL_NEG_LMUL" > "Rings.mult_minus_left"
-  "REAL_NEG_LE0" > "Groups.neg_le_0_iff_le"
-  "REAL_NEG_INV" > "Rings.nonzero_inverse_minus_eq"
-  "REAL_NEG_GT0" > "Groups.neg_0_less_iff_less"
-  "REAL_NEG_GE0" > "Groups.neg_0_le_iff_le"
-  "REAL_NEG_EQ0" > "Groups.neg_equal_0_iff_equal"
+  "REAL_NOT_LE" > "Orderings.linorder_class.not_le"
+  "REAL_NEG_SUB" > "Groups.ab_group_add_class.minus_diff_eq"
+  "REAL_NEG_RMUL" > "Int.int_arith_rules_14"
+  "REAL_NEG_NEG" > "Groups.group_add_class.minus_minus"
+  "REAL_NEG_MUL2" > "Rings.ring_class.minus_mult_minus"
+  "REAL_NEG_MINUS1" > "Semiring_Normalization.comm_ring_1_class.normalizing_ring_rules_1"
+  "REAL_NEG_LT0" > "Groups.ordered_ab_group_add_class.neg_less_0_iff_less"
+  "REAL_NEG_LMUL" > "Int.int_arith_rules_13"
+  "REAL_NEG_LE0" > "Groups.ordered_ab_group_add_class.neg_le_0_iff_le"
+  "REAL_NEG_INV" > "Fields.division_ring_class.nonzero_inverse_minus_eq"
+  "REAL_NEG_GT0" > "Groups.ordered_ab_group_add_class.neg_0_less_iff_less"
+  "REAL_NEG_GE0" > "Groups.ordered_ab_group_add_class.neg_0_le_iff_le"
+  "REAL_NEG_EQ0" > "Groups.group_add_class.neg_equal_0_iff_equal"
   "REAL_NEG_EQ" > "HOL4Real.real.REAL_NEG_EQ"
-  "REAL_NEG_ADD" > "Groups.minus_add_distrib"
-  "REAL_NEG_0" > "Groups.minus_zero"
-  "REAL_NEGNEG" > "Groups.minus_minus"
-  "REAL_MUL_SYM" > "Int.zmult_ac_2"
-  "REAL_MUL_RZERO" > "Rings.mult_zero_right"
-  "REAL_MUL_RNEG" > "Rings.mult_minus_right"
-  "REAL_MUL_RINV" > "Rings.right_inverse"
-  "REAL_MUL_RID" > "Finite_Set.AC_mult.f_e.ident"
-  "REAL_MUL_LZERO" > "Rings.mult_zero_left"
-  "REAL_MUL_LNEG" > "Rings.mult_minus_left"
-  "REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV"
-  "REAL_MUL_LID" > "Finite_Set.AC_mult.f_e.left_ident"
-  "REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC"
+  "REAL_NEG_ADD" > "Groups.ab_group_add_class.minus_add_distrib"
+  "REAL_NEG_0" > "Groups.group_add_class.minus_zero"
+  "REAL_NEGNEG" > "Groups.group_add_class.minus_minus"
+  "REAL_MUL_SYM" > "Fields.linordered_field_class.sign_simps_21"
+  "REAL_MUL_RZERO" > "Divides.arithmetic_simps_41"
+  "REAL_MUL_RNEG" > "Int.int_arith_rules_14"
+  "REAL_MUL_RINV" > "Fields.division_ring_class.right_inverse"
+  "REAL_MUL_RID" > "Divides.arithmetic_simps_43"
+  "REAL_MUL_LZERO" > "Divides.arithmetic_simps_40"
+  "REAL_MUL_LNEG" > "Int.int_arith_rules_13"
+  "REAL_MUL_LINV" > "Fields.division_ring_class.left_inverse"
+  "REAL_MUL_LID" > "Divides.arithmetic_simps_42"
+  "REAL_MUL_ASSOC" > "Fields.linordered_field_class.sign_simps_22"
   "REAL_MUL" > "RealDef.real_of_nat_mult"
   "REAL_MIDDLE2" > "HOL4Real.real.REAL_MIDDLE2"
   "REAL_MIDDLE1" > "HOL4Real.real.REAL_MIDDLE1"
-  "REAL_MEAN" > "Rings.dense"
-  "REAL_LT_TRANS" > "Set.basic_trans_rules_21"
+  "REAL_MEAN" > "Orderings.dense_linorder_class.dense"
+  "REAL_LT_TRANS" > "Orderings.order_less_trans"
   "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
-  "REAL_LT_SUB_RADD" > "Groups.compare_rls_6"
-  "REAL_LT_SUB_LADD" > "Groups.compare_rls_7"
-  "REAL_LT_RMUL_IMP" > "Rings.mult_strict_right_mono"
+  "REAL_LT_SUB_RADD" > "Fields.linordered_field_class.sign_simps_4"
+  "REAL_LT_SUB_LADD" > "Fields.linordered_field_class.sign_simps_3"
+  "REAL_LT_RMUL_IMP" > "Rings.linordered_semiring_strict_class.mult_strict_right_mono"
   "REAL_LT_RMUL_0" > "HOL4Real.real.REAL_LT_RMUL_0"
   "REAL_LT_RMUL" > "RealDef.real_mult_less_iff1"
   "REAL_LT_REFL" > "Orderings.order_less_irrefl"
-  "REAL_LT_RDIV_EQ" > "Rings.pos_less_divide_eq"
+  "REAL_LT_RDIV_EQ" > "Fields.linordered_field_class.pos_less_divide_eq"
   "REAL_LT_RDIV_0" > "HOL4Real.real.REAL_LT_RDIV_0"
   "REAL_LT_RDIV" > "HOL4Real.real.REAL_LT_RDIV"
-  "REAL_LT_RADD" > "Groups.add_less_cancel_right"
+  "REAL_LT_RADD" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_right"
   "REAL_LT_NZ" > "HOL4Real.real.REAL_LT_NZ"
   "REAL_LT_NEGTOTAL" > "HOL4Real.real.REAL_LT_NEGTOTAL"
-  "REAL_LT_NEG" > "Groups.neg_less_iff_less"
+  "REAL_LT_NEG" > "Groups.ordered_ab_group_add_class.neg_less_iff_less"
   "REAL_LT_MULTIPLE" > "HOL4Real.real.REAL_LT_MULTIPLE"
-  "REAL_LT_MUL2" > "Rings.mult_strict_mono'"
-  "REAL_LT_MUL" > "Rings.mult_pos_pos"
-  "REAL_LT_LMUL_IMP" > "Rings.linordered_comm_semiring_strict_class.mult_strict_mono"
+  "REAL_LT_MUL2" > "Rings.linordered_semiring_strict_class.mult_strict_mono'"
+  "REAL_LT_MUL" > "RealDef.real_mult_order"
+  "REAL_LT_LMUL_IMP" > "RealDef.real_mult_less_mono2"
   "REAL_LT_LMUL_0" > "HOL4Real.real.REAL_LT_LMUL_0"
-  "REAL_LT_LMUL" > "HOL4Real.real.REAL_LT_LMUL"
-  "REAL_LT_LE" > "Orderings.order_class.order_less_le"
-  "REAL_LT_LDIV_EQ" > "Rings.pos_divide_less_eq"
-  "REAL_LT_LADD" > "Groups.add_less_cancel_left"
-  "REAL_LT_INV_EQ" > "Rings.inverse_positive_iff_positive"
-  "REAL_LT_INV" > "Rings.less_imp_inverse_less"
-  "REAL_LT_IMP_NE" > "Orderings.less_imp_neq"
+  "REAL_LT_LMUL" > "Rings.linordered_ring_strict_class.mult_less_cancel_left_pos"
+  "REAL_LT_LE" > "Orderings.order_class.less_le"
+  "REAL_LT_LDIV_EQ" > "Fields.linordered_field_class.pos_divide_less_eq"
+  "REAL_LT_LADD" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_left"
+  "REAL_LT_INV_EQ" > "Fields.linordered_field_inverse_zero_class.inverse_positive_iff_positive"
+  "REAL_LT_INV" > "Fields.linordered_field_class.less_imp_inverse_less"
+  "REAL_LT_IMP_NE" > "Orderings.order_class.less_imp_neq"
   "REAL_LT_IMP_LE" > "Orderings.order_less_imp_le"
-  "REAL_LT_IADD" > "Groups.add_strict_left_mono"
+  "REAL_LT_IADD" > "Groups.ordered_cancel_ab_semigroup_add_class.add_strict_left_mono"
   "REAL_LT_HALF2" > "HOL4Real.real.REAL_LT_HALF2"
-  "REAL_LT_HALF1" > "NatSimprocs.half_gt_zero_iff"
+  "REAL_LT_HALF1" > "Int.half_gt_zero_iff"
   "REAL_LT_GT" > "Orderings.order_less_not_sym"
   "REAL_LT_FRACTION_0" > "HOL4Real.real.REAL_LT_FRACTION_0"
   "REAL_LT_FRACTION" > "HOL4Real.real.REAL_LT_FRACTION"
-  "REAL_LT_DIV" > "Rings.divide_pos_pos"
+  "REAL_LT_DIV" > "Fields.linordered_field_class.divide_pos_pos"
   "REAL_LT_ANTISYM" > "HOL4Real.real.REAL_LT_ANTISYM"
-  "REAL_LT_ADD_SUB" > "Groups.compare_rls_7"
+  "REAL_LT_ADD_SUB" > "Fields.linordered_field_class.sign_simps_3"
   "REAL_LT_ADDR" > "HOL4Real.real.REAL_LT_ADDR"
   "REAL_LT_ADDNEG2" > "HOL4Real.real.REAL_LT_ADDNEG2"
   "REAL_LT_ADDNEG" > "HOL4Real.real.REAL_LT_ADDNEG"
   "REAL_LT_ADDL" > "HOL4Real.real.REAL_LT_ADDL"
-  "REAL_LT_ADD2" > "Groups.add_strict_mono"
+  "REAL_LT_ADD2" > "Groups.add_mono_thms_linordered_field_5"
   "REAL_LT_ADD1" > "HOL4Real.real.REAL_LT_ADD1"
-  "REAL_LT_ADD" > "Groups.add_pos_pos"
+  "REAL_LT_ADD" > "Groups.ordered_comm_monoid_add_class.add_pos_pos"
   "REAL_LT_1" > "HOL4Real.real.REAL_LT_1"
-  "REAL_LT_01" > "Rings.zero_less_one"
-  "REAL_LTE_TRANS" > "Set.basic_trans_rules_24"
+  "REAL_LT_01" > "Rings.linordered_semidom_class.zero_less_one"
+  "REAL_LTE_TRANS" > "Orderings.order_less_le_trans"
   "REAL_LTE_TOTAL" > "HOL4Real.real.REAL_LTE_TOTAL"
   "REAL_LTE_ANTSYM" > "HOL4Real.real.REAL_LTE_ANTSYM"
-  "REAL_LTE_ADD2" > "Groups.add_less_le_mono"
-  "REAL_LTE_ADD" > "Groups.add_pos_nonneg"
+  "REAL_LTE_ADD2" > "Groups.add_mono_thms_linordered_field_3"
+  "REAL_LTE_ADD" > "Groups.ordered_comm_monoid_add_class.add_pos_nonneg"
   "REAL_LT1_POW2" > "HOL4Real.real.REAL_LT1_POW2"
   "REAL_LT" > "RealDef.real_of_nat_less_iff"
-  "REAL_LNEG_UNIQ" > "HOL4Real.real.REAL_LNEG_UNIQ"
+  "REAL_LNEG_UNIQ" > "Groups.group_add_class.eq_neg_iff_add_eq_0"
   "REAL_LINV_UNIQ" > "HOL4Real.real.REAL_LINV_UNIQ"
-  "REAL_LE_TRANS" > "Set.basic_trans_rules_25"
-  "REAL_LE_TOTAL" > "Orderings.linorder_class.linorder_linear"
-  "REAL_LE_SUB_RADD" > "Groups.compare_rls_8"
-  "REAL_LE_SUB_LADD" > "Groups.compare_rls_9"
-  "REAL_LE_SQUARE" > "Rings.zero_le_square"
-  "REAL_LE_RNEG" > "Groups.le_eq_neg"
-  "REAL_LE_RMUL_IMP" > "Rings.mult_right_mono"
+  "REAL_LE_TRANS" > "Orderings.order_trans_rules_23"
+  "REAL_LE_TOTAL" > "Orderings.linorder_class.linear"
+  "REAL_LE_SUB_RADD" > "Fields.linordered_field_class.sign_simps_2"
+  "REAL_LE_SUB_LADD" > "Fields.linordered_field_class.sign_simps_1"
+  "REAL_LE_SQUARE" > "Rings.linordered_ring_class.zero_le_square"
+  "REAL_LE_RNEG" > "HOL4Real.real.REAL_LE_RNEG"
+  "REAL_LE_RMUL_IMP" > "Rings.ordered_semiring_class.mult_right_mono"
   "REAL_LE_RMUL" > "RealDef.real_mult_le_cancel_iff1"
-  "REAL_LE_REFL" > "Finite_Set.max.f_below.below_refl"
-  "REAL_LE_RDIV_EQ" > "Rings.pos_le_divide_eq"
-  "REAL_LE_RDIV" > "Rings.mult_imp_le_div_pos"
-  "REAL_LE_RADD" > "Groups.add_le_cancel_right"
-  "REAL_LE_POW2" > "Nat_Numeral.zero_compare_simps_12"
+  "REAL_LE_REFL" > "Orderings.preorder_class.order_refl"
+  "REAL_LE_RDIV_EQ" > "Fields.linordered_field_class.pos_le_divide_eq"
+  "REAL_LE_RDIV" > "Fields.linordered_field_class.mult_imp_le_div_pos"
+  "REAL_LE_RADD" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_right"
+  "REAL_LE_POW2" > "Nat_Numeral.linordered_idom_class.zero_le_power2"
   "REAL_LE_NEGTOTAL" > "HOL4Real.real.REAL_LE_NEGTOTAL"
-  "REAL_LE_NEGR" > "Groups.le_minus_self_iff"
-  "REAL_LE_NEGL" > "Groups.minus_le_self_iff"
-  "REAL_LE_NEG2" > "Groups.neg_le_iff_le"
-  "REAL_LE_NEG" > "Groups.neg_le_iff_le"
-  "REAL_LE_MUL2" > "HOL4Real.real.REAL_LE_MUL2"
-  "REAL_LE_MUL" > "Rings.mult_nonneg_nonneg"
-  "REAL_LE_LT" > "Orderings.order_le_less"
+  "REAL_LE_NEGR" > "Groups.linordered_ab_group_add_class.le_minus_self_iff"
+  "REAL_LE_NEGL" > "Groups.linordered_ab_group_add_class.minus_le_self_iff"
+  "REAL_LE_NEG2" > "Groups.ordered_ab_group_add_class.neg_le_iff_le"
+  "REAL_LE_NEG" > "Groups.ordered_ab_group_add_class.neg_le_iff_le"
+  "REAL_LE_MUL2" > "Rings.ordered_semiring_class.mult_mono'"
+  "REAL_LE_MUL" > "Rings.mult_sign_intros_1"
+  "REAL_LE_LT" > "Orderings.order_class.le_less"
   "REAL_LE_LNEG" > "RealDef.real_0_le_add_iff"
-  "REAL_LE_LMUL_IMP" > "Rings.mult_mono"
+  "REAL_LE_LMUL_IMP" > "Rings.ordered_comm_semiring_class.comm_mult_left_mono"
   "REAL_LE_LMUL" > "RealDef.real_mult_le_cancel_iff2"
-  "REAL_LE_LDIV_EQ" > "Rings.pos_divide_le_eq"
-  "REAL_LE_LDIV" > "Rings.mult_imp_div_pos_le"
-  "REAL_LE_LADD_IMP" > "Groups.add_left_mono"
-  "REAL_LE_LADD" > "Groups.add_le_cancel_left"
-  "REAL_LE_INV_EQ" > "Rings.inverse_nonnegative_iff_nonnegative"
+  "REAL_LE_LDIV_EQ" > "Fields.linordered_field_class.pos_divide_le_eq"
+  "REAL_LE_LDIV" > "Fields.linordered_field_class.mult_imp_div_pos_le"
+  "REAL_LE_LADD_IMP" > "Groups.ordered_ab_semigroup_add_class.add_left_mono"
+  "REAL_LE_LADD" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left"
+  "REAL_LE_INV_EQ" > "Fields.linordered_field_inverse_zero_class.inverse_nonnegative_iff_nonnegative"
   "REAL_LE_INV" > "HOL4Real.real.REAL_LE_INV"
-  "REAL_LE_DOUBLE" > "Groups.zero_le_double_add_iff_zero_le_single_add"
+  "REAL_LE_DOUBLE" > "Groups.linordered_ab_group_add_class.zero_le_double_add_iff_zero_le_single_add"
   "REAL_LE_DIV" > "HOL4Real.real.REAL_LE_DIV"
-  "REAL_LE_ANTISYM" > "Orderings.order_eq_iff"
+  "REAL_LE_ANTISYM" > "Orderings.order_class.eq_iff"
   "REAL_LE_ADDR" > "HOL4Real.real.REAL_LE_ADDR"
   "REAL_LE_ADDL" > "HOL4Real.real.REAL_LE_ADDL"
-  "REAL_LE_ADD2" > "Groups.add_mono"
-  "REAL_LE_ADD" > "Groups.add_nonneg_nonneg"
-  "REAL_LE_01" > "Rings.zero_le_one"
-  "REAL_LET_TRANS" > "Set.basic_trans_rules_23"
-  "REAL_LET_TOTAL" > "Orderings.linorder_le_less_linear"
+  "REAL_LE_ADD2" > "Groups.add_mono_thms_linordered_semiring_1"
+  "REAL_LE_ADD" > "Groups.ordered_comm_monoid_add_class.add_nonneg_nonneg"
+  "REAL_LE_01" > "Rings.linordered_semidom_class.zero_le_one"
+  "REAL_LET_TRANS" > "Orderings.order_le_less_trans"
+  "REAL_LET_TOTAL" > "Orderings.linorder_class.le_less_linear"
   "REAL_LET_ANTISYM" > "HOL4Real.real.REAL_LET_ANTISYM"
-  "REAL_LET_ADD2" > "Groups.add_le_less_mono"
-  "REAL_LET_ADD" > "Groups.add_nonneg_pos"
+  "REAL_LET_ADD2" > "Groups.add_mono_thms_linordered_field_4"
+  "REAL_LET_ADD" > "Groups.ordered_comm_monoid_add_class.add_nonneg_pos"
   "REAL_LE1_POW2" > "HOL4Real.real.REAL_LE1_POW2"
   "REAL_LE" > "RealDef.real_of_nat_le_iff"
-  "REAL_LDISTRIB" > "Rings.ring_eq_simps_2"
-  "REAL_INV_POS" > "Rings.positive_imp_inverse_positive"
-  "REAL_INV_NZ" > "Rings.nonzero_imp_inverse_nonzero"
+  "REAL_LDISTRIB" > "Fields.linordered_field_class.sign_simps_7"
+  "REAL_INV_POS" > "Fields.linordered_field_class.positive_imp_inverse_positive"
+  "REAL_INV_NZ" > "Fields.division_ring_class.nonzero_imp_inverse_nonzero"
   "REAL_INV_MUL" > "HOL4Real.real.REAL_INV_MUL"
-  "REAL_INV_LT1" > "Fields.one_less_inverse"
-  "REAL_INV_INV" > "Rings.inverse_inverse_eq"
-  "REAL_INV_EQ_0" > "Rings.inverse_nonzero_iff_nonzero"
-  "REAL_INV_1OVER" > "Rings.inverse_eq_divide"
-  "REAL_INV_0" > "Rings.division_ring_inverse_zero_class.inverse_zero"
-  "REAL_INVINV" > "Rings.nonzero_inverse_inverse_eq"
-  "REAL_INV1" > "Rings.inverse_1"
+  "REAL_INV_LT1" > "Fields.linordered_field_class.one_less_inverse"
+  "REAL_INV_INV" > "Fields.division_ring_inverse_zero_class.inverse_inverse_eq"
+  "REAL_INV_EQ_0" > "Fields.division_ring_inverse_zero_class.inverse_nonzero_iff_nonzero"
+  "REAL_INV_1OVER" > "Fields.division_ring_class.inverse_eq_divide"
+  "REAL_INV_0" > "Fields.division_ring_inverse_zero_class.inverse_zero"
+  "REAL_INVINV" > "Fields.division_ring_class.nonzero_inverse_inverse_eq"
+  "REAL_INV1" > "Fields.division_ring_class.inverse_1"
   "REAL_INJ" > "RealDef.real_of_nat_inject"
   "REAL_HALF_DOUBLE" > "RComplete.real_sum_of_halves"
   "REAL_FACT_NZ" > "HOL4Real.real.REAL_FACT_NZ"
-  "REAL_EQ_SUB_RADD" > "Rings.ring_eq_simps_15"
-  "REAL_EQ_SUB_LADD" > "Rings.ring_eq_simps_16"
-  "REAL_EQ_RMUL_IMP" > "Rings.field_mult_cancel_right_lemma"
-  "REAL_EQ_RMUL" > "Rings.field_mult_cancel_right"
+  "REAL_EQ_SUB_RADD" > "Fields.linordered_field_class.sign_simps_12"
+  "REAL_EQ_SUB_LADD" > "Fields.linordered_field_class.sign_simps_11"
+  "REAL_EQ_RMUL_IMP" > "HOL4Real.real.REAL_EQ_RMUL_IMP"
+  "REAL_EQ_RMUL" > "Rings.mult_compare_simps_13"
   "REAL_EQ_RDIV_EQ" > "HOL4Real.real.REAL_EQ_RDIV_EQ"
-  "REAL_EQ_RADD" > "Groups.add_right_cancel"
-  "REAL_EQ_NEG" > "Groups.neg_equal_iff_equal"
-  "REAL_EQ_MUL_LCANCEL" > "Rings.field_mult_cancel_left"
+  "REAL_EQ_RADD" > "Groups.cancel_semigroup_add_class.add_right_cancel"
+  "REAL_EQ_NEG" > "Groups.group_add_class.neg_equal_iff_equal"
+  "REAL_EQ_MUL_LCANCEL" > "Rings.mult_compare_simps_14"
   "REAL_EQ_LMUL_IMP" > "HOL4Real.real.REAL_EQ_LMUL_IMP"
   "REAL_EQ_LMUL2" > "RealDef.real_mult_left_cancel"
-  "REAL_EQ_LMUL" > "Rings.field_mult_cancel_left"
+  "REAL_EQ_LMUL" > "Rings.mult_compare_simps_14"
   "REAL_EQ_LDIV_EQ" > "HOL4Real.real.REAL_EQ_LDIV_EQ"
-  "REAL_EQ_LADD" > "Groups.add_left_cancel"
+  "REAL_EQ_LADD" > "Groups.cancel_semigroup_add_class.add_left_cancel"
   "REAL_EQ_IMP_LE" > "Orderings.order_eq_refl"
-  "REAL_ENTIRE" > "Rings.field_mult_eq_0_iff"
+  "REAL_ENTIRE" > "Rings.ring_no_zero_divisors_class.mult_eq_0_iff"
   "REAL_DOWN2" > "RealDef.real_lbound_gt_zero"
   "REAL_DOWN" > "HOL4Real.real.REAL_DOWN"
   "REAL_DOUBLE" > "Int.mult_2"
   "REAL_DIV_RMUL" > "HOL4Real.real.REAL_DIV_RMUL"
-  "REAL_DIV_REFL" > "Rings.divide_self"
+  "REAL_DIV_REFL" > "Fields.division_ring_class.divide_self"
   "REAL_DIV_MUL2" > "HOL4Real.real.REAL_DIV_MUL2"
-  "REAL_DIV_LZERO" > "Rings.divide_zero_left"
+  "REAL_DIV_LZERO" > "Fields.division_ring_class.divide_zero_left"
   "REAL_DIV_LMUL" > "HOL4Real.real.REAL_DIV_LMUL"
-  "REAL_DIFFSQ" > "HOL4Real.real.REAL_DIFFSQ"
-  "REAL_ARCH_LEAST" > "HOL4Real.real.REAL_ARCH_LEAST"
+  "REAL_DIFFSQ" > "Rings.comm_ring_class.square_diff_square_factored"
+  "REAL_ARCH_LEAST" > "Transcendental.reals_Archimedean4"
   "REAL_ARCH" > "RComplete.reals_Archimedean3"
-  "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
+  "REAL_ADD_SYM" > "Fields.linordered_field_class.sign_simps_18"
   "REAL_ADD_SUB2" > "HOL4Real.real.REAL_ADD_SUB2"
   "REAL_ADD_SUB" > "HOL4Real.real.REAL_ADD_SUB"
-  "REAL_ADD_RINV" > "Groups.right_minus"
+  "REAL_ADD_RINV" > "Groups.group_add_class.right_minus"
   "REAL_ADD_RID_UNIQ" > "HOL4Real.real.REAL_ADD_RID_UNIQ"
-  "REAL_ADD_RID" > "Finite_Set.AC_add.f_e.ident"
-  "REAL_ADD_RDISTRIB" > "Rings.ring_eq_simps_1"
-  "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
+  "REAL_ADD_RID" > "Divides.arithmetic_simps_39"
+  "REAL_ADD_RDISTRIB" > "Fields.linordered_field_class.sign_simps_8"
+  "REAL_ADD_LINV" > "Groups.ab_group_add_class.ab_left_minus"
   "REAL_ADD_LID_UNIQ" > "HOL4Real.real.REAL_ADD_LID_UNIQ"
-  "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident"
-  "REAL_ADD_LDISTRIB" > "Rings.ring_eq_simps_2"
-  "REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC"
-  "REAL_ADD2_SUB2" > "HOL4Real.real.REAL_ADD2_SUB2"
+  "REAL_ADD_LID" > "Divides.arithmetic_simps_38"
+  "REAL_ADD_LDISTRIB" > "Fields.linordered_field_class.sign_simps_7"
+  "REAL_ADD_ASSOC" > "Fields.linordered_field_class.sign_simps_19"
+  "REAL_ADD2_SUB2" > "RealDef.add_diff_add"
   "REAL_ADD" > "RealDef.real_of_nat_add"
-  "REAL_ABS_TRIANGLE" > "Groups.abs_triangle_ineq"
-  "REAL_ABS_POS" > "Groups.abs_ge_zero"
-  "REAL_ABS_MUL" > "Rings.abs_mult"
-  "REAL_ABS_0" > "Int.bin_arith_simps_28"
+  "REAL_ABS_TRIANGLE" > "Groups.ordered_ab_group_add_abs_class.abs_triangle_ineq"
+  "REAL_ABS_POS" > "Groups.ordered_ab_group_add_abs_class.abs_ge_zero"
+  "REAL_ABS_MUL" > "Rings.linordered_idom_class.abs_mult"
+  "REAL_ABS_0" > "Divides.arithmetic_simps_27"
   "REAL_10" > "HOL4Compat.REAL_10"
   "REAL_1" > "HOL4Real.real.REAL_1"
   "REAL_0" > "HOL4Real.real.REAL_0"
   "REAL" > "RealDef.real_of_nat_Suc"
   "POW_ZERO_EQ" > "HOL4Real.real.POW_ZERO_EQ"
-  "POW_ZERO" > "RealPow.realpow_zero_zero"
+  "POW_ZERO" > "HOL4Real.real.POW_ZERO"
   "POW_POS_LT" > "HOL4Real.real.POW_POS_LT"
-  "POW_POS" > "Power.zero_le_power"
+  "POW_POS" > "Power.linordered_semidom_class.zero_le_power"
   "POW_PLUS1" > "HOL4Real.real.POW_PLUS1"
-  "POW_ONE" > "Power.power_one"
-  "POW_NZ" > "Power.field_power_not_zero"
-  "POW_MUL" > "Power.power_mult_distrib"
-  "POW_MINUS1" > "Nat_Numeral.power_minus1_even"
+  "POW_ONE" > "Power.monoid_mult_class.power_one"
+  "POW_NZ" > "Power.ring_1_no_zero_divisors_class.field_power_not_zero"
+  "POW_MUL" > "Power.comm_monoid_mult_class.power_mult_distrib"
+  "POW_MINUS1" > "Nat_Numeral.ring_1_class.power_minus1_even"
   "POW_M1" > "HOL4Real.real.POW_M1"
   "POW_LT" > "HOL4Real.real.POW_LT"
-  "POW_LE" > "Power.power_mono"
-  "POW_INV" > "Power.nonzero_power_inverse"
-  "POW_EQ" > "Power.power_inject_base"
-  "POW_ADD" > "Power.power_add"
-  "POW_ABS" > "Power.power_abs"
-  "POW_2_LT" > "RealPow.two_realpow_gt"
-  "POW_2_LE1" > "RealPow.two_realpow_ge_one"
-  "POW_2" > "Nat_Numeral.power2_eq_square"
-  "POW_1" > "Power.power_one_right"
+  "POW_LE" > "Power.linordered_semidom_class.power_mono"
+  "POW_INV" > "Power.division_ring_class.nonzero_power_inverse"
+  "POW_EQ" > "Power.linordered_semidom_class.power_inject_base"
+  "POW_ADD" > "Power.monoid_mult_class.power_add"
+  "POW_ABS" > "Power.linordered_idom_class.power_abs"
+  "POW_2_LT" > "RealDef.two_realpow_gt"
+  "POW_2_LE1" > "RealDef.two_realpow_ge_one"
+  "POW_2" > "Nat_Numeral.monoid_mult_class.power2_eq_square"
+  "POW_1" > "Power.monoid_mult_class.power_one_right"
   "POW_0" > "Power.power_0_Suc"
-  "ABS_ZERO" > "Groups.abs_eq_0"
-  "ABS_TRIANGLE" > "Groups.abs_triangle_ineq"
+  "ABS_ZERO" > "Groups.ordered_ab_group_add_abs_class.abs_eq_0"
+  "ABS_TRIANGLE" > "Groups.ordered_ab_group_add_abs_class.abs_triangle_ineq"
   "ABS_SUM" > "HOL4Real.real.ABS_SUM"
-  "ABS_SUB_ABS" > "Groups.abs_triangle_ineq3"
-  "ABS_SUB" > "Groups.abs_minus_commute"
+  "ABS_SUB_ABS" > "Groups.ordered_ab_group_add_abs_class.abs_triangle_ineq3"
+  "ABS_SUB" > "Groups.ordered_ab_group_add_abs_class.abs_minus_commute"
   "ABS_STILLNZ" > "HOL4Real.real.ABS_STILLNZ"
   "ABS_SIGN2" > "HOL4Real.real.ABS_SIGN2"
   "ABS_SIGN" > "HOL4Real.real.ABS_SIGN"
   "ABS_REFL" > "HOL4Real.real.ABS_REFL"
-  "ABS_POW2" > "Nat_Numeral.abs_power2"
-  "ABS_POS" > "Groups.abs_ge_zero"
-  "ABS_NZ" > "Groups.zero_less_abs_iff"
-  "ABS_NEG" > "Groups.abs_minus_cancel"
+  "ABS_POW2" > "Nat_Numeral.linordered_idom_class.abs_power2"
+  "ABS_POS" > "Groups.ordered_ab_group_add_abs_class.abs_ge_zero"
+  "ABS_NZ" > "Groups.ordered_ab_group_add_abs_class.zero_less_abs_iff"
+  "ABS_NEG" > "Groups.ordered_ab_group_add_abs_class.abs_minus_cancel"
   "ABS_N" > "RealDef.abs_real_of_nat_cancel"
-  "ABS_MUL" > "Rings.abs_mult"
+  "ABS_MUL" > "Rings.linordered_idom_class.abs_mult"
   "ABS_LT_MUL2" > "HOL4Real.real.ABS_LT_MUL2"
-  "ABS_LE" > "Groups.abs_ge_self"
-  "ABS_INV" > "Rings.nonzero_abs_inverse"
-  "ABS_DIV" > "Rings.nonzero_abs_divide"
+  "ABS_LE" > "Groups.ordered_ab_group_add_abs_class.abs_ge_self"
+  "ABS_INV" > "Fields.linordered_field_class.nonzero_abs_inverse"
+  "ABS_DIV" > "Fields.linordered_field_class.nonzero_abs_divide"
   "ABS_CIRCLE" > "HOL4Real.real.ABS_CIRCLE"
   "ABS_CASES" > "HOL4Real.real.ABS_CASES"
   "ABS_BOUNDS" > "RealDef.abs_le_interval_iff"
@@ -352,8 +352,8 @@
   "ABS_BETWEEN2" > "HOL4Real.real.ABS_BETWEEN2"
   "ABS_BETWEEN1" > "HOL4Real.real.ABS_BETWEEN1"
   "ABS_BETWEEN" > "HOL4Real.real.ABS_BETWEEN"
-  "ABS_ABS" > "Groups.abs_idempotent"
-  "ABS_1" > "Int.bin_arith_simps_29"
-  "ABS_0" > "Int.bin_arith_simps_28"
+  "ABS_ABS" > "Groups.ordered_ab_group_add_abs_class.abs_idempotent"
+  "ABS_1" > "Divides.arithmetic_simps_28"
+  "ABS_0" > "Divides.arithmetic_simps_27"
 
 end
--- a/src/HOL/Import/HOL/realax.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/realax.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -27,14 +27,18 @@
   "treal_add" > "HOL4Real.realax.treal_add"
   "treal_1" > "HOL4Real.realax.treal_1"
   "treal_0" > "HOL4Real.realax.treal_0"
-  "real_neg" > "Groups.uminus" :: "real => real"
-  "real_mul" > "Groups.times" :: "real => real => real"
-  "real_lt" > "Orderings.less" :: "real => real => bool"
-  "real_add" > "Groups.plus" :: "real => real => real"
-  "real_1" > "Groups.one" :: "real"
-  "real_0" > "Groups.zero" :: "real"
-  "inv" > "Ring.inverse" :: "real => real"
+  "real_sub" > "Groups.minus_class.minus" :: "real => real => real"
+  "real_neg" > "Groups.uminus_class.uminus" :: "real => real"
+  "real_mul" > "Groups.times_class.times" :: "real => real => real"
+  "real_lt" > "Orderings.ord_class.less" :: "real => real => bool"
+  "real_div" > "Fields.inverse_class.divide" :: "real => real => real"
+  "real_add" > "Groups.plus_class.plus" :: "real => real => real"
+  "real_1" > "Groups.one_class.one" :: "real"
+  "real_0" > "Groups.zero_class.zero" :: "real"
+  "mk_real" > "HOL.undefined"
+  "inv" > "Fields.inverse_class.inverse" :: "real => real"
   "hreal_of_treal" > "HOL4Real.realax.hreal_of_treal"
+  "dest_real" > "HOL.undefined"
 
 thm_maps
   "treal_of_hreal_def" > "HOL4Real.realax.treal_of_hreal_def"
@@ -91,21 +95,21 @@
   "TREAL_ADD_ASSOC" > "HOL4Real.realax.TREAL_ADD_ASSOC"
   "TREAL_10" > "HOL4Real.realax.TREAL_10"
   "REAL_SUP_ALLPOS" > "HOL4Compat.REAL_SUP_ALLPOS"
-  "REAL_MUL_SYM" > "Int.zmult_ac_2"
-  "REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV"
-  "REAL_MUL_LID" > "Finite_Set.AC_mult.f_e.left_ident"
-  "REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC"
-  "REAL_LT_TRANS" > "Set.basic_trans_rules_21"
+  "REAL_MUL_SYM" > "Fields.linordered_field_class.sign_simps_21"
+  "REAL_MUL_LINV" > "Fields.division_ring_class.left_inverse"
+  "REAL_MUL_LID" > "Divides.arithmetic_simps_42"
+  "REAL_MUL_ASSOC" > "Fields.linordered_field_class.sign_simps_22"
+  "REAL_LT_TRANS" > "Orderings.order_less_trans"
   "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
   "REAL_LT_REFL" > "Orderings.order_less_irrefl"
-  "REAL_LT_MUL" > "Rings.mult_pos_pos"
-  "REAL_LT_IADD" > "Groups.add_strict_left_mono"
-  "REAL_LDISTRIB" > "Rings.ring_eq_simps_2"
-  "REAL_INV_0" > "Rings.division_ring_inverse_zero_class.inverse_zero"
-  "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
-  "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
-  "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident"
-  "REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC"
+  "REAL_LT_MUL" > "RealDef.real_mult_order"
+  "REAL_LT_IADD" > "Groups.ordered_cancel_ab_semigroup_add_class.add_strict_left_mono"
+  "REAL_LDISTRIB" > "Fields.linordered_field_class.sign_simps_7"
+  "REAL_INV_0" > "Fields.division_ring_inverse_zero_class.inverse_zero"
+  "REAL_ADD_SYM" > "Fields.linordered_field_class.sign_simps_18"
+  "REAL_ADD_LINV" > "Groups.ab_group_add_class.ab_left_minus"
+  "REAL_ADD_LID" > "Divides.arithmetic_simps_38"
+  "REAL_ADD_ASSOC" > "Fields.linordered_field_class.sign_simps_19"
   "REAL_10" > "HOL4Compat.REAL_10"
   "HREAL_RDISTRIB" > "HOL4Real.realax.HREAL_RDISTRIB"
   "HREAL_LT_REFL" > "HOL4Real.realax.HREAL_LT_REFL"
--- a/src/HOL/Import/HOL/rich_list.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/rich_list.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -34,7 +34,7 @@
   "AND_EL" > "HOL4Base.rich_list.AND_EL"
 
 thm_maps
-  "list_INDUCT" > "HOL4Compat.unzip.induct"
+  "list_INDUCT" > "HOL4Compat.list_INDUCT"
   "list_CASES" > "HOL4Compat.list_CASES"
   "ZIP_UNZIP" > "HOL4Base.list.ZIP_UNZIP"
   "ZIP_SNOC" > "HOL4Base.rich_list.ZIP_SNOC"
@@ -60,7 +60,7 @@
   "SPLITP" > "HOL4Base.rich_list.SPLITP"
   "SOME_EL_SNOC" > "HOL4Base.rich_list.SOME_EL_SNOC"
   "SOME_EL_SEG" > "HOL4Base.rich_list.SOME_EL_SEG"
-  "SOME_EL_REVERSE" > "HOL4Base.rich_list.SOME_EL_REVERSE"
+  "SOME_EL_REVERSE" > "List.list_ex_rev"
   "SOME_EL_MAP" > "HOL4Base.rich_list.SOME_EL_MAP"
   "SOME_EL_LASTN" > "HOL4Base.rich_list.SOME_EL_LASTN"
   "SOME_EL_FOLDR_MAP" > "HOL4Base.rich_list.SOME_EL_FOLDR_MAP"
@@ -71,7 +71,7 @@
   "SOME_EL_DISJ" > "HOL4Base.rich_list.SOME_EL_DISJ"
   "SOME_EL_BUTLASTN" > "HOL4Base.rich_list.SOME_EL_BUTLASTN"
   "SOME_EL_BUTFIRSTN" > "HOL4Base.rich_list.SOME_EL_BUTFIRSTN"
-  "SOME_EL_APPEND" > "HOL4Base.list.EXISTS_APPEND"
+  "SOME_EL_APPEND" > "List.list_ex_append"
   "SOME_EL" > "HOL4Compat.list_exists_DEF"
   "SNOC_REVERSE_CONS" > "HOL4Base.rich_list.SNOC_REVERSE_CONS"
   "SNOC_INDUCT" > "HOL4Base.rich_list.SNOC_INDUCT"
@@ -100,7 +100,7 @@
   "REVERSE_SNOC" > "HOL4Base.rich_list.REVERSE_SNOC"
   "REVERSE_REVERSE" > "List.rev_rev_ident"
   "REVERSE_FOLDR" > "HOL4Base.rich_list.REVERSE_FOLDR"
-  "REVERSE_FOLDL" > "HOL4Base.rich_list.REVERSE_FOLDL"
+  "REVERSE_FOLDL" > "List.rev_foldl_cons"
   "REVERSE_FLAT" > "HOL4Base.rich_list.REVERSE_FLAT"
   "REVERSE_EQ_NIL" > "List.rev_is_Nil_conv"
   "REVERSE_APPEND" > "List.rev_append"
@@ -116,21 +116,21 @@
   "OR_EL_DEF" > "HOL4Base.rich_list.OR_EL_DEF"
   "NULL_FOLDR" > "HOL4Base.rich_list.NULL_FOLDR"
   "NULL_FOLDL" > "HOL4Base.rich_list.NULL_FOLDL"
-  "NULL_EQ_NIL" > "HOL4Base.rich_list.NULL_EQ_NIL"
+  "NULL_EQ_NIL" > "List.eq_Nil_null"
   "NULL_DEF" > "HOL4Compat.NULL_DEF"
   "NULL" > "HOL4Base.list.NULL"
   "NOT_SOME_EL_ALL_EL" > "HOL4Base.list.NOT_EXISTS"
   "NOT_SNOC_NIL" > "HOL4Base.rich_list.NOT_SNOC_NIL"
   "NOT_NIL_SNOC" > "HOL4Base.rich_list.NOT_NIL_SNOC"
-  "NOT_NIL_CONS" > "List.list.simps_1"
+  "NOT_NIL_CONS" > "List.list.distinct_1"
   "NOT_EQ_LIST" > "HOL4Base.list.NOT_EQ_LIST"
-  "NOT_CONS_NIL" > "List.list.simps_2"
+  "NOT_CONS_NIL" > "List.list.distinct_2"
   "NOT_ALL_EL_SOME_EL" > "HOL4Base.list.NOT_EVERY"
   "MONOID_APPEND_NIL" > "HOL4Base.rich_list.MONOID_APPEND_NIL"
-  "MAP_o" > "HOL4Base.rich_list.MAP_o"
+  "MAP_o" > "List.map.comp"
   "MAP_SNOC" > "HOL4Base.rich_list.MAP_SNOC"
   "MAP_REVERSE" > "List.rev_map"
-  "MAP_MAP_o" > "List.map_map"
+  "MAP_MAP_o" > "List.map.compositionality"
   "MAP_FOLDR" > "HOL4Base.rich_list.MAP_FOLDR"
   "MAP_FOLDL" > "HOL4Base.rich_list.MAP_FOLDL"
   "MAP_FLAT" > "List.map_concat"
@@ -230,7 +230,7 @@
   "FOLDL_SNOC" > "HOL4Base.rich_list.FOLDL_SNOC"
   "FOLDL_SINGLE" > "HOL4Base.rich_list.FOLDL_SINGLE"
   "FOLDL_REVERSE" > "HOL4Base.rich_list.FOLDL_REVERSE"
-  "FOLDL_MAP" > "HOL4Base.rich_list.FOLDL_MAP"
+  "FOLDL_MAP" > "List.foldl_map"
   "FOLDL_FOLDR_REVERSE" > "List.foldl_foldr"
   "FOLDL_FILTER" > "HOL4Base.rich_list.FOLDL_FILTER"
   "FOLDL_APPEND" > "List.foldl_append"
@@ -238,7 +238,7 @@
   "FLAT_SNOC" > "HOL4Base.rich_list.FLAT_SNOC"
   "FLAT_REVERSE" > "HOL4Base.rich_list.FLAT_REVERSE"
   "FLAT_FOLDR" > "HOL4Base.rich_list.FLAT_FOLDR"
-  "FLAT_FOLDL" > "HOL4Base.rich_list.FLAT_FOLDL"
+  "FLAT_FOLDL" > "List.concat_conv_foldl"
   "FLAT_FLAT" > "HOL4Base.rich_list.FLAT_FLAT"
   "FLAT_APPEND" > "List.concat_append"
   "FLAT" > "HOL4Compat.FLAT"
@@ -301,7 +301,7 @@
   "ELL" > "HOL4Base.rich_list.ELL"
   "EL" > "HOL4Base.rich_list.EL"
   "CONS_APPEND" > "HOL4Base.rich_list.CONS_APPEND"
-  "CONS_11" > "List.list.simps_3"
+  "CONS_11" > "List.list.inject"
   "CONS" > "HOL4Base.list.CONS"
   "COMM_MONOID_FOLDR" > "HOL4Base.rich_list.COMM_MONOID_FOLDR"
   "COMM_MONOID_FOLDL" > "HOL4Base.rich_list.COMM_MONOID_FOLDL"
--- a/src/HOL/Import/HOL/seq.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/seq.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -106,7 +106,7 @@
   "MAX_LEMMA" > "HOL4Real.seq.MAX_LEMMA"
   "GP_FINITE" > "HOL4Real.seq.GP_FINITE"
   "GP" > "HOL4Real.seq.GP"
-  "BOLZANO_LEMMA" > "HOL4Real.seq.BOLZANO_LEMMA"
-  "ABS_NEG_LEMMA" > "HOL4Real.seq.ABS_NEG_LEMMA"
+  "BOLZANO_LEMMA" > "Deriv.lemma_BOLZANO"
+  "ABS_NEG_LEMMA" > "Series.rabs_ratiotest_lemma"
 
 end
--- a/src/HOL/Import/HOL/sum.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/sum.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -3,10 +3,10 @@
 import_segment "hol4"
 
 type_maps
-  "sum" > "+"
+  "sum" > "Sum_Type.sum"
 
 const_maps
-  "sum_case" > "Datatype.sum.sum_case"
+  "sum_case" > "Sum_Type.sum.sum_case"
   "OUTR" > "HOL4Compat.OUTR"
   "OUTL" > "HOL4Compat.OUTL"
   "ISR" > "HOL4Compat.ISR"
@@ -15,24 +15,24 @@
   "INL" > "Sum_Type.Inl"
 
 thm_maps
-  "sum_distinct1" > "Datatype.sum.simps_2"
-  "sum_distinct" > "Datatype.sum.simps_1"
+  "sum_distinct1" > "Sum_Type.Inr_not_Inl"
+  "sum_distinct" > "Sum_Type.Inl_not_Inr"
   "sum_case_def" > "HOL4Compat.sum_case_def"
   "sum_case_cong" > "HOL4Base.sum.sum_case_cong"
-  "sum_INDUCT" > "HOL4Compat.OUTR.induct"
-  "sum_CASES" > "Datatype.sum.nchotomy"
+  "sum_axiom" > "HOL4Compat.sum_axiom"
+  "sum_INDUCT" > "Sum_Type.sum.induct"
+  "sum_CASES" > "Sum_Type.sum.nchotomy"
   "OUTR" > "HOL4Compat.OUTR"
   "OUTL" > "HOL4Compat.OUTL"
   "ISR" > "HOL4Compat.ISR"
   "ISL_OR_ISR" > "HOL4Base.sum.ISL_OR_ISR"
   "ISL" > "HOL4Compat.ISL"
-  "INR_neq_INL" > "Datatype.sum.simps_2"
+  "INR_neq_INL" > "Sum_Type.Inr_not_Inl"
   "INR_INL_11" > "HOL4Compat.INR_INL_11"
   "INR" > "HOL4Base.sum.INR"
   "INL" > "HOL4Base.sum.INL"
 
 ignore_thms
-  "sum_axiom"
   "sum_TY_DEF"
   "sum_ISO_DEF"
   "sum_Axiom"
--- a/src/HOL/Import/HOL/word32.imp	Tue Sep 06 19:03:41 2011 -0700
+++ b/src/HOL/Import/HOL/word32.imp	Tue Sep 06 22:41:35 2011 -0700
@@ -284,7 +284,7 @@
   "HB_primdef" > "HOL4Word32.word32.HB_primdef"
   "HB_def" > "HOL4Word32.word32.HB_def"
   "FUNPOW_THM2" > "HOL4Word32.word32.FUNPOW_THM2"
-  "FUNPOW_THM" > "HOL4Word32.word32.FUNPOW_THM"
+  "FUNPOW_THM" > "Nat.funpow_swap1"
   "FUNPOW_COMP" > "HOL4Word32.word32.FUNPOW_COMP"
   "EQ_ADD_RCANCELw" > "HOL4Word32.word32.EQ_ADD_RCANCELw"
   "EQ_ADD_RCANCEL_QT" > "HOL4Word32.word32.EQ_ADD_RCANCEL_QT"