Added HOL proof importer.
authorskalberg
Fri, 02 Apr 2004 17:37:45 +0200
changeset 14516 a183dec876ab
parent 14515 86f2daf48a3c
child 14517 7ae3b247c6e9
Added HOL proof importer.
src/HOL/Import/Generate-HOL/GenHOL4Base.thy
src/HOL/Import/Generate-HOL/GenHOL4Prob.thy
src/HOL/Import/Generate-HOL/GenHOL4Real.thy
src/HOL/Import/Generate-HOL/GenHOL4Vec.thy
src/HOL/Import/Generate-HOL/GenHOL4Word32.thy
src/HOL/Import/Generate-HOL/ROOT.ML
src/HOL/Import/HOL/HOL4.thy
src/HOL/Import/HOL/HOL4Base.thy
src/HOL/Import/HOL/HOL4Prob.thy
src/HOL/Import/HOL/HOL4Real.thy
src/HOL/Import/HOL/HOL4Vec.thy
src/HOL/Import/HOL/HOL4Word32.thy
src/HOL/Import/HOL/README
src/HOL/Import/HOL/ROOT.ML
src/HOL/Import/HOL/arithmetic.imp
src/HOL/Import/HOL/bits.imp
src/HOL/Import/HOL/bool.imp
src/HOL/Import/HOL/boolean_sequence.imp
src/HOL/Import/HOL/bword_arith.imp
src/HOL/Import/HOL/bword_bitop.imp
src/HOL/Import/HOL/bword_num.imp
src/HOL/Import/HOL/combin.imp
src/HOL/Import/HOL/divides.imp
src/HOL/Import/HOL/hrat.imp
src/HOL/Import/HOL/hreal.imp
src/HOL/Import/HOL/ind_type.imp
src/HOL/Import/HOL/lim.imp
src/HOL/Import/HOL/list.imp
src/HOL/Import/HOL/marker.imp
src/HOL/Import/HOL/nets.imp
src/HOL/Import/HOL/num.imp
src/HOL/Import/HOL/numeral.imp
src/HOL/Import/HOL/one.imp
src/HOL/Import/HOL/operator.imp
src/HOL/Import/HOL/option.imp
src/HOL/Import/HOL/pair.imp
src/HOL/Import/HOL/poly.imp
src/HOL/Import/HOL/powser.imp
src/HOL/Import/HOL/pred_set.imp
src/HOL/Import/HOL/prim_rec.imp
src/HOL/Import/HOL/prime.imp
src/HOL/Import/HOL/prob.imp
src/HOL/Import/HOL/prob_algebra.imp
src/HOL/Import/HOL/prob_canon.imp
src/HOL/Import/HOL/prob_extra.imp
src/HOL/Import/HOL/prob_indep.imp
src/HOL/Import/HOL/prob_pseudo.imp
src/HOL/Import/HOL/prob_uniform.imp
src/HOL/Import/HOL/real.imp
src/HOL/Import/HOL/realax.imp
src/HOL/Import/HOL/relation.imp
src/HOL/Import/HOL/res_quan.imp
src/HOL/Import/HOL/rich_list.imp
src/HOL/Import/HOL/seq.imp
src/HOL/Import/HOL/state_transformer.imp
src/HOL/Import/HOL/sum.imp
src/HOL/Import/HOL/topology.imp
src/HOL/Import/HOL/transc.imp
src/HOL/Import/HOL/word32.imp
src/HOL/Import/HOL/word_base.imp
src/HOL/Import/HOL/word_bitop.imp
src/HOL/Import/HOL/word_num.imp
src/HOL/Import/HOL4Compat.thy
src/HOL/Import/HOL4Setup.thy
src/HOL/Import/HOL4Syntax.thy
src/HOL/Import/MakeEqual.thy
src/HOL/Import/ROOT.ML
src/HOL/Import/hol4rews.ML
src/HOL/Import/import_package.ML
src/HOL/Import/import_syntax.ML
src/HOL/Import/proof_kernel.ML
src/HOL/Import/replay.ML
src/HOL/Import/shuffler.ML
src/HOL/IsaMakefile
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,273 @@
+theory GenHOL4Base = HOL4Compat + HOL4Syntax:;
+
+import_segment "hol4";
+
+setup_dump "../HOL" "HOL4Base";
+
+append_dump "theory HOL4Base = HOL4Compat + HOL4Syntax:";
+
+import_theory bool;
+
+const_maps
+  T               > True
+  F               > False
+  "!"             > All
+  "/\\"           > "op &"
+  "\\/"           > "op |"
+  "?"             > Ex
+  "?!"            > Ex1
+  "~"             > Not
+  COND            > If
+  bool_case       > Datatype.bool.bool_case
+  ONE_ONE         > HOL4Setup.ONE_ONE
+  ONTO            > HOL4Setup.ONTO
+  TYPE_DEFINITION > HOL4Setup.TYPE_DEFINITION
+  LET             > HOL4Compat.LET;
+
+ignore_thms
+  BOUNDED_DEF
+  BOUNDED_THM
+  UNBOUNDED_DEF
+  UNBOUNDED_THM;
+
+end_import;
+
+import_theory combin;
+
+const_maps
+  o > Fun.comp;
+
+end_import;
+
+import_theory sum;
+
+type_maps
+  sum > "+";
+
+const_maps
+  INL      > Inl
+  INR      > Inr
+  ISL      > HOL4Compat.ISL
+  ISR      > HOL4Compat.ISR
+  OUTL     > HOL4Compat.OUTL
+  OUTR     > HOL4Compat.OUTR
+  sum_case > Datatype.sum.sum_case;
+
+ignore_thms
+  sum_TY_DEF
+  sum_ISO_DEF
+  IS_SUM_REP
+  INL_DEF
+  INR_DEF
+  sum_axiom
+  sum_Axiom;
+
+end_import;
+
+import_theory one;
+
+type_maps
+  one > Product_Type.unit;
+
+const_maps
+  one > Product_Type.Unity;
+
+ignore_thms
+    one_TY_DEF
+    one_axiom
+    one_Axiom
+    one_DEF;
+
+end_import;
+
+import_theory option;
+
+type_maps
+    option > Datatype.option;
+
+const_maps
+    NONE        > Datatype.option.None
+    SOME        > Datatype.option.Some
+    option_case > Datatype.option.option_case
+    OPTION_MAP  > Datatype.option_map
+    THE         > Datatype.the
+    IS_SOME     > HOL4Compat.IS_SOME
+    IS_NONE     > HOL4Compat.IS_NONE
+    OPTION_JOIN > HOL4Compat.OPTION_JOIN;
+
+ignore_thms
+    option_axiom
+    option_Axiom
+    option_TY_DEF
+    option_REP_ABS_DEF
+    SOME_DEF
+    NONE_DEF;
+
+end_import;
+
+import_theory marker;
+end_import;
+
+import_theory relation;
+
+const_renames
+  reflexive > pred_reflexive;
+
+end_import;
+
+import_theory pair;
+
+type_maps
+    prod > "*";
+
+const_maps
+    ","       > Pair
+    FST       > fst
+    SND       > snd
+    CURRY     > curry
+    UNCURRY   > split
+    "##"      > prod_fun
+    pair_case > split;
+
+ignore_thms
+    prod_TY_DEF
+    MK_PAIR_DEF
+    IS_PAIR_DEF
+    ABS_REP_prod
+    COMMA_DEF;
+
+end_import;
+
+import_theory num;
+
+type_maps
+  num > nat;
+
+const_maps
+  SUC > Suc
+  0   > 0 :: nat;
+
+ignore_thms
+    num_TY_DEF
+    num_ISO_DEF
+    IS_NUM_REP
+    ZERO_REP_DEF
+    SUC_REP_DEF
+    ZERO_DEF
+    SUC_DEF;
+
+end_import;
+
+import_theory prim_rec;
+
+const_maps
+    "<" > "op <" :: "[nat,nat]\<Rightarrow>bool";
+
+end_import;
+
+import_theory arithmetic;
+
+const_maps
+  ALT_ZERO     > HOL4Compat.ALT_ZERO
+  NUMERAL_BIT1 > HOL4Compat.NUMERAL_BIT1
+  NUMERAL_BIT2 > HOL4Compat.NUMERAL_BIT2
+  NUMERAL      > HOL4Compat.NUMERAL
+  num_case     > Nat.nat.nat_case
+  ">"          > HOL4Compat.nat_gt
+  ">="         > HOL4Compat.nat_ge
+  FUNPOW       > HOL4Compat.FUNPOW
+  "<="         > "op <="          :: "[nat,nat]\<Rightarrow>bool"
+  "+"          > "op +"           :: "[nat,nat]\<Rightarrow>nat"
+  "*"          > "op *"           :: "[nat,nat]\<Rightarrow>nat"
+  "-"          > "op -"           :: "[nat,nat]\<Rightarrow>nat"
+  MIN          > HOL.min          :: "[nat,nat]\<Rightarrow>nat"
+  MAX          > HOL.max          :: "[nat,nat]\<Rightarrow>nat"
+  DIV          > "Divides.op div" :: "[nat,nat]\<Rightarrow>nat"
+  MOD          > "Divides.op mod" :: "[nat,nat]\<Rightarrow>nat"
+  EXP          > Nat.power        :: "[nat,nat]\<Rightarrow>nat";
+
+end_import;
+
+import_theory hrat;
+end_import;
+
+import_theory hreal;
+end_import;
+
+import_theory numeral;
+end_import;
+
+import_theory ind_type;
+end_import;
+
+import_theory divides;
+
+const_maps
+    divides > "Divides.op dvd" :: "[nat,nat]\<Rightarrow>bool";
+
+end_import;
+
+import_theory prime;
+end_import;
+
+import_theory list;
+
+type_maps
+    list > List.list;
+
+const_maps
+  CONS      > List.list.Cons
+  NIL       > List.list.Nil
+  list_case > List.list.list_case
+  NULL      > List.null
+  HD        > List.hd
+  TL        > List.tl
+  MAP       > List.map
+  MEM       > "List.op mem"
+  FILTER    > List.filter
+  FOLDL     > List.foldl
+  EVERY     > List.list_all
+  REVERSE   > List.rev
+  LAST      > List.last
+  FRONT     > List.butlast
+  APPEND    > "List.op @"
+  FLAT      > List.concat
+  LENGTH    > Nat.size
+  REPLICATE > List.replicate
+  list_size > HOL4Compat.list_size
+  SUM       > HOL4Compat.sum
+  FOLDR     > HOL4Compat.FOLDR
+  EXISTS    > HOL4Compat.list_exists
+  MAP2      > HOL4Compat.map2
+  ZIP       > HOL4Compat.ZIP
+  UNZIP     > HOL4Compat.unzip;
+
+ignore_thms
+  list_TY_DEF
+  list_repfns
+  list0_def
+  list1_def
+  NIL
+  CONS_def;
+
+end_import;
+
+import_theory pred_set;
+end_import;
+
+import_theory operator;
+end_import;
+
+import_theory rich_list;
+end_import;
+
+import_theory state_transformer;
+end_import;
+
+append_dump "end";
+
+flush_dump;
+
+import_segment "";
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Prob.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,43 @@
+theory GenHOL4Prob = GenHOL4Real:
+
+import_segment "hol4";
+
+setup_dump "../HOL" "HOL4Prob";
+
+append_dump "theory HOL4Prob = HOL4Real:";
+
+import_theory prob_extra;
+
+const_moves
+  COMPL > GenHOL4Base.pred_set.COMPL;
+
+end_import;
+
+import_theory prob_canon;
+end_import;
+
+import_theory boolean_sequence;
+end_import;
+
+import_theory prob_algebra;
+end_import;
+
+import_theory prob;
+end_import;
+
+import_theory prob_pseudo;
+end_import;
+
+import_theory prob_indep;
+end_import;
+
+import_theory prob_uniform;
+end_import;
+
+append_dump "end";
+
+flush_dump;
+
+import_segment "";
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Real.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,85 @@
+theory GenHOL4Real = GenHOL4Base:
+
+import_segment "hol4";
+
+setup_dump "../HOL" "HOL4Real";
+
+append_dump "theory HOL4Real = HOL4Base:";
+
+import_theory realax;
+
+type_maps
+  real > RealDef.real;
+
+const_maps
+  real_0   > 0           :: real
+  real_1   > 1           :: real
+  real_neg > uminus      :: "real \<Rightarrow> real"
+  inv      > HOL.inverse :: "real \<Rightarrow> real"
+  real_add > "op +"      :: "[real,real] \<Rightarrow> real"
+  real_mul > "op *"      :: "[real,real] \<Rightarrow> real"
+  real_lt  > "op <"      :: "[real,real] \<Rightarrow> bool";
+
+ignore_thms
+    real_TY_DEF
+    real_tybij
+    real_0
+    real_1
+    real_neg
+    real_inv
+    real_add
+    real_mul
+    real_lt
+    real_of_hreal
+    hreal_of_real
+    REAL_ISO_EQ
+    REAL_POS
+    SUP_ALLPOS_LEMMA1
+    SUP_ALLPOS_LEMMA2
+    SUP_ALLPOS_LEMMA3
+    SUP_ALLPOS_LEMMA4;
+
+end_import;
+
+import_theory real;
+
+const_maps
+  real_gt     > HOL4Compat.real_gt
+  real_ge     > HOL4Compat.real_ge
+  real_lte    > "op <="      :: "[real,real] \<Rightarrow> bool"
+  real_sub    > "op -"       :: "[real,real] \<Rightarrow> real"
+  "/"         > HOL.divide   :: "[real,real] \<Rightarrow> real"
+  pow         > Nat.power    :: "[real,nat] \<Rightarrow> real"
+  abs         > HOL.abs      :: "real \<Rightarrow> real"
+  real_of_num > RealDef.real :: "nat \<Rightarrow> real";
+
+end_import;
+
+import_theory topology;
+end_import;
+
+import_theory nets;
+end_import;
+
+import_theory seq;
+end_import;
+
+import_theory lim;
+end_import;
+
+import_theory powser;
+end_import;
+
+import_theory transc;
+end_import;
+
+import_theory poly;
+end_import;
+
+append_dump "end";
+
+flush_dump;
+
+import_segment "";
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Vec.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,40 @@
+theory GenHOL4Vec = GenHOL4Base:
+
+import_segment "hol4";
+
+setup_dump "../HOL" "HOL4Vec";
+
+append_dump "theory HOL4Vec = HOL4Base:";
+
+import_theory res_quan;
+end_import;
+
+import_theory word_base;
+
+const_renames
+  BIT > bit;
+
+end_import;
+
+import_theory word_num;
+end_import;
+
+import_theory word_bitop;
+end_import;
+
+import_theory bword_num;
+end_import;
+
+import_theory bword_arith;
+end_import;
+
+import_theory bword_bitop;
+end_import;
+
+append_dump "end";
+
+flush_dump;
+
+import_segment "";
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Word32.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,29 @@
+theory GenHOL4Word32 = GenHOL4Base:;
+
+import_segment "hol4";
+
+setup_dump "../HOL" "HOL4Word32";
+
+append_dump "theory HOL4Word32 = HOL4Base:";
+
+import_theory bits;
+
+const_renames
+  BIT > bit
+
+end_import;
+
+import_theory word32;
+
+const_renames
+  "==" > EQUIV;
+
+end_import;
+
+append_dump "end";
+
+flush_dump;
+
+import_segment "";
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Generate-HOL/ROOT.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,5 @@
+with_path ".." use_thy "HOL4Compat";
+with_path ".." use_thy "HOL4Syntax";
+use_thy "GenHOL4Prob";
+use_thy "GenHOL4Vec";
+use_thy "GenHOL4Word32";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,3 @@
+theory HOL4 = HOL4Vec + HOL4Word32 + HOL4Real:
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4Base.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,4754 @@
+theory HOL4Base = HOL4Compat + HOL4Syntax:
+
+;setup_theory bool
+
+constdefs
+  ARB :: "'a" 
+  "ARB == SOME x. True"
+
+lemma ARB_DEF: "ARB = (SOME x. True)"
+  by (import bool ARB_DEF)
+
+constdefs
+  IN :: "'a => ('a => bool) => bool" 
+  "IN == %x f. f x"
+
+lemma IN_DEF: "IN = (%x f. f x)"
+  by (import bool IN_DEF)
+
+constdefs
+  RES_FORALL :: "('a => bool) => ('a => bool) => bool" 
+  "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)"
+  by (import bool RES_FORALL_DEF)
+
+constdefs
+  RES_EXISTS :: "('a => bool) => ('a => bool) => bool" 
+  "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)"
+  by (import bool RES_EXISTS_DEF)
+
+constdefs
+  RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" 
+  "RES_EXISTS_UNIQUE ==
+%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 m. RES_EXISTS p m &
+       RES_FORALL p (%x. RES_FORALL p (%y. m x & m y --> x = y)))"
+  by (import bool RES_EXISTS_UNIQUE_DEF)
+
+constdefs
+  RES_SELECT :: "('a => bool) => ('a => bool) => 'a" 
+  "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)"
+  by (import bool RES_SELECT_DEF)
+
+lemma EXCLUDED_MIDDLE: "ALL t. t | ~ t"
+  by (import bool EXCLUDED_MIDDLE)
+
+lemma FORALL_THM: "All f = All f"
+  by (import bool FORALL_THM)
+
+lemma EXISTS_THM: "Ex f = Ex f"
+  by (import bool EXISTS_THM)
+
+lemma F_IMP: "ALL t. ~ t --> t --> False"
+  by (import bool F_IMP)
+
+lemma NOT_AND: "~ (t & ~ t)"
+  by (import bool NOT_AND)
+
+lemma AND_CLAUSES: "ALL t.
+   (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.
+   (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.
+   (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. (~ ~ t) = t) & (~ True) = False & (~ False) = True"
+  by (import bool NOT_CLAUSES)
+
+lemma BOOL_EQ_DISTINCT: "True ~= False & False ~= True"
+  by (import bool BOOL_EQ_DISTINCT)
+
+lemma EQ_CLAUSES: "ALL t.
+   (True = t) = t &
+   (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
+  by (import bool EQ_CLAUSES)
+
+lemma COND_CLAUSES: "ALL t1 t2. (if True then t1 else t2) = t1 & (if False then t1 else t2) = t2"
+  by (import bool COND_CLAUSES)
+
+lemma SELECT_UNIQUE: "ALL P x. (ALL y. 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. P & Q) = ((EX x::'a. P) & (EX x::'a. Q))"
+  by (import bool BOTH_EXISTS_AND_THM)
+
+lemma BOTH_FORALL_OR_THM: "ALL (P::bool) Q::bool.
+   (ALL x::'a. P | Q) = ((ALL x::'a. P) | (ALL x::'a. Q))"
+  by (import bool BOTH_FORALL_OR_THM)
+
+lemma BOTH_FORALL_IMP_THM: "ALL (P::bool) Q::bool.
+   (ALL x::'a. P --> Q) = ((EX x::'a. P) --> (ALL x::'a. Q))"
+  by (import bool BOTH_FORALL_IMP_THM)
+
+lemma BOTH_EXISTS_IMP_THM: "ALL (P::bool) Q::bool.
+   (EX x::'a. P --> Q) = ((ALL x::'a. P) --> (EX x::'a. Q))"
+  by (import bool BOTH_EXISTS_IMP_THM)
+
+lemma OR_IMP_THM: "ALL A B. (A = (B | A)) = (B --> A)"
+  by (import bool OR_IMP_THM)
+
+lemma DE_MORGAN_THM: "ALL A B. (~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)"
+  by (import bool DE_MORGAN_THM)
+
+lemma IMP_F_EQ_F: "ALL t. (t --> False) = (t = False)"
+  by (import bool IMP_F_EQ_F)
+
+lemma EQ_EXPAND: "ALL t1 t2. (t1 = t2) = (t1 & t2 | ~ t1 & ~ t2)"
+  by (import bool EQ_EXPAND)
+
+lemma COND_RATOR: "ALL b f g x. (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 f g. (%x. if b then f x else g x) = (if b then f else g)"
+  by (import bool COND_ABS)
+
+lemma COND_EXPAND: "ALL b t1 t2. (if b then t1 else t2) = ((~ b | t1) & (b | t2))"
+  by (import bool COND_EXPAND)
+
+lemma ONE_ONE_THM: "ALL f. inj f = (ALL x1 x2. f x1 = f x2 --> x1 = x2)"
+  by (import bool ONE_ONE_THM)
+
+lemma ABS_REP_THM: "(All::(('a => bool) => bool) => bool)
+ (%P::'a => bool.
+     (op -->::bool => bool => bool)
+      ((Ex::(('b => 'a) => bool) => bool)
+        ((TYPE_DEFINITION::('a => bool) => ('b => 'a) => bool) P))
+      ((Ex::(('b => 'a) => bool) => bool)
+        (%x::'b => 'a.
+            (Ex::(('a => 'b) => bool) => bool)
+             (%abs::'a => 'b.
+                 (op &::bool => bool => bool)
+                  ((All::('b => bool) => bool)
+                    (%a::'b. (op =::'b => 'b => bool) (abs (x a)) a))
+                  ((All::('a => bool) => bool)
+                    (%r::'a.
+                        (op =::bool => bool => bool) (P r)
+                         ((op =::'a => 'a => bool) (x (abs r)) r)))))))"
+  by (import bool ABS_REP_THM)
+
+lemma LET_RAND: "(P::'b => bool) (Let (M::'a) (N::'a => 'b)) = (let x::'a = M in P (N x))"
+  by (import bool LET_RAND)
+
+lemma LET_RATOR: "Let (M::'a) (N::'a => 'b => 'c) (b::'b) = (let x::'a = M in N x b)"
+  by (import bool LET_RATOR)
+
+lemma SWAP_FORALL_THM: "ALL P. (ALL x. All (P x)) = (ALL y x. P x y)"
+  by (import bool SWAP_FORALL_THM)
+
+lemma SWAP_EXISTS_THM: "ALL P. (EX x. Ex (P x)) = (EX y x. P x y)"
+  by (import bool SWAP_EXISTS_THM)
+
+lemma AND_CONG: "ALL P P' Q Q'. (Q --> P = P') & (P' --> Q = Q') --> (P & Q) = (P' & Q')"
+  by (import bool AND_CONG)
+
+lemma OR_CONG: "ALL P P' Q Q'. (~ Q --> P = P') & (~ P' --> Q = Q') --> (P | Q) = (P' | Q')"
+  by (import bool OR_CONG)
+
+lemma COND_CONG: "ALL P Q x x' y y'.
+   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 --> y) --> (z --> w) --> (if b then x else z) --> (if b then y else w)"
+  by (import bool MONO_COND)
+
+lemma SKOLEM_THM: "ALL P. (ALL x. Ex (P x)) = (EX f. ALL x. P x (f x))"
+  by (import bool SKOLEM_THM)
+
+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)"
+  by (import bool bool_case_thm)
+
+lemma bool_case_ID: "ALL x b. (case b of True => x | False => x) = x"
+  by (import bool bool_case_ID)
+
+lemma boolAxiom: "ALL e0 e1. EX x. x True = e0 & x False = e1"
+  by (import bool boolAxiom)
+
+lemma UEXISTS_OR_THM: "ALL P Q. (EX! x. P x | Q x) --> Ex1 P | Ex1 Q"
+  by (import bool UEXISTS_OR_THM)
+
+lemma UEXISTS_SIMP: "(EX! x::'a. (t::bool)) = (t & (ALL x::'a. All (op = x)))"
+  by (import bool UEXISTS_SIMP)
+
+consts
+  RES_ABSTRACT :: "('a => bool) => ('a => 'b) => 'a => 'b" 
+
+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 => 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. f = (%b. True) | f = (%b. False) | f = (%b. b) | f = Not"
+  by (import bool BOOL_FUN_CASES_THM)
+
+lemma BOOL_FUN_INDUCT: "ALL P. P (%b. True) & P (%b. False) & P (%b. b) & P Not --> All P"
+  by (import bool BOOL_FUN_INDUCT)
+
+;end_setup
+
+;setup_theory combin
+
+constdefs
+  K :: "'a => 'b => 'a" 
+  "K == %x y. x"
+
+lemma K_DEF: "K = (%x y. x)"
+  by (import combin K_DEF)
+
+constdefs
+  S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" 
+  "S == %f g x. f x (g x)"
+
+lemma S_DEF: "S = (%f g x. f x (g x))"
+  by (import combin S_DEF)
+
+constdefs
+  I :: "'a => 'a" 
+  "(op ==::('a => 'a) => ('a => 'a) => prop) (I::'a => 'a)
+ ((S::('a => ('a => 'a) => 'a) => ('a => 'a => 'a) => 'a => 'a)
+   (K::'a => ('a => 'a) => 'a) (K::'a => 'a => 'a))"
+
+lemma I_DEF: "(op =::('a => 'a) => ('a => 'a) => bool) (I::'a => 'a)
+ ((S::('a => ('a => 'a) => 'a) => ('a => 'a => 'a) => 'a => 'a)
+   (K::'a => ('a => 'a) => 'a) (K::'a => 'a => 'a))"
+  by (import combin I_DEF)
+
+constdefs
+  C :: "('a => 'b => 'c) => 'b => 'a => 'c" 
+  "C == %f x y. f y x"
+
+lemma C_DEF: "C = (%f x y. f y x)"
+  by (import combin C_DEF)
+
+constdefs
+  W :: "('a => 'a => 'b) => 'a => 'b" 
+  "W == %f x. f x x"
+
+lemma W_DEF: "W = (%f x. f x x)"
+  by (import combin W_DEF)
+
+lemma I_THM: "ALL x. I x = x"
+  by (import combin I_THM)
+
+lemma I_o_ID: "ALL f. I o f = f & f o I = f"
+  by (import combin I_o_ID)
+
+;end_setup
+
+;setup_theory sum
+
+lemma ISL_OR_ISR: "ALL x. ISL x | ISR x"
+  by (import sum ISL_OR_ISR)
+
+lemma INL: "ALL x. ISL x --> Inl (OUTL x) = x"
+  by (import sum INL)
+
+lemma INR: "ALL x. ISR x --> Inr (OUTR x) = x"
+  by (import sum INR)
+
+lemma sum_case_cong: "ALL (M::'b + 'c) (M'::'b + 'c) (f::'b => 'a) g::'c => 'a.
+   M = M' &
+   (ALL x::'b. M' = Inl x --> f x = (f'::'b => 'a) x) &
+   (ALL y::'c. M' = Inr y --> g y = (g'::'c => 'a) y) -->
+   sum_case f g M = sum_case f' g' M'"
+  by (import sum sum_case_cong)
+
+;end_setup
+
+;setup_theory one
+
+;end_setup
+
+;setup_theory option
+
+lemma option_CLAUSES: "(op &::bool => bool => bool)
+ ((All::('a => bool) => bool)
+   (%x::'a.
+       (All::('a => bool) => bool)
+        (%y::'a.
+            (op =::bool => bool => bool)
+             ((op =::'a option => 'a option => bool) ((Some::'a ~=> 'a) x)
+               ((Some::'a ~=> 'a) y))
+             ((op =::'a => 'a => bool) x y))))
+ ((op &::bool => bool => bool)
+   ((All::('a => bool) => bool)
+     (%x::'a.
+         (op =::'a => 'a => bool)
+          ((the::'a option => 'a) ((Some::'a ~=> 'a) x)) x))
+   ((op &::bool => bool => bool)
+     ((All::('a => bool) => bool)
+       (%x::'a.
+           (Not::bool => bool)
+            ((op =::'a option => 'a option => bool) (None::'a option)
+              ((Some::'a ~=> 'a) x))))
+     ((op &::bool => bool => bool)
+       ((All::('a => bool) => bool)
+         (%x::'a.
+             (Not::bool => bool)
+              ((op =::'a option => 'a option => bool) ((Some::'a ~=> 'a) x)
+                (None::'a option))))
+       ((op &::bool => bool => bool)
+         ((All::('a => bool) => bool)
+           (%x::'a.
+               (op =::bool => bool => bool)
+                ((IS_SOME::'a option => bool) ((Some::'a ~=> 'a) x))
+                (True::bool)))
+         ((op &::bool => bool => bool)
+           ((op =::bool => bool => bool)
+             ((IS_SOME::'a option => bool) (None::'a option)) (False::bool))
+           ((op &::bool => bool => bool)
+             ((All::('a option => bool) => bool)
+               (%x::'a option.
+                   (op =::bool => bool => bool)
+                    ((IS_NONE::'a option => bool) x)
+                    ((op =::'a option => 'a option => bool) x
+                      (None::'a option))))
+             ((op &::bool => bool => bool)
+               ((All::('a option => bool) => bool)
+                 (%x::'a option.
+                     (op =::bool => bool => bool)
+                      ((Not::bool => bool) ((IS_SOME::'a option => bool) x))
+                      ((op =::'a option => 'a option => bool) x
+                        (None::'a option))))
+               ((op &::bool => bool => bool)
+                 ((All::('a option => bool) => bool)
+                   (%x::'a option.
+                       (op -->::bool => bool => bool)
+                        ((IS_SOME::'a option => bool) x)
+                        ((op =::'a option => 'a option => bool)
+                          ((Some::'a ~=> 'a) ((the::'a option => 'a) x))
+                          x)))
+                 ((op &::bool => bool => bool)
+                   ((All::('a option => bool) => bool)
+                     (%x::'a option.
+                         (op =::'a option => 'a option => bool)
+                          ((option_case::'a option
+   => ('a ~=> 'a) => 'a option ~=> 'a)
+                            (None::'a option) (Some::'a ~=> 'a) x)
+                          x))
+                   ((op &::bool => bool => bool)
+                     ((All::('a option => bool) => bool)
+                       (%x::'a option.
+                           (op =::'a option => 'a option => bool)
+                            ((option_case::'a option
+     => ('a ~=> 'a) => 'a option ~=> 'a)
+                              x (Some::'a ~=> 'a) x)
+                            x))
+                     ((op &::bool => bool => bool)
+                       ((All::('a option => bool) => bool)
+                         (%x::'a option.
+                             (op -->::bool => bool => bool)
+                              ((IS_NONE::'a option => bool) x)
+                              ((op =::'b => 'b => bool)
+                                ((option_case::'b
+         => ('a => 'b) => 'a option => 'b)
+                                  (e::'b) (f::'a => 'b) x)
+                                e)))
+                       ((op &::bool => bool => bool)
+                         ((All::('a option => bool) => bool)
+                           (%x::'a option.
+                               (op -->::bool => bool => bool)
+                                ((IS_SOME::'a option => bool) x)
+                                ((op =::'b => 'b => bool)
+                                  ((option_case::'b
+           => ('a => 'b) => 'a option => 'b)
+                                    e f x)
+                                  (f ((the::'a option => 'a) x)))))
+                         ((op &::bool => bool => bool)
+                           ((All::('a option => bool) => bool)
+                             (%x::'a option.
+                                 (op -->::bool => bool => bool)
+                                  ((IS_SOME::'a option => bool) x)
+                                  ((op =::'a option => 'a option => bool)
+                                    ((option_case::'a option
+             => ('a ~=> 'a) => 'a option ~=> 'a)
+(ea::'a option) (Some::'a ~=> 'a) x)
+                                    x)))
+                           ((op &::bool => bool => bool)
+                             ((All::('b => bool) => bool)
+                               (%u::'b.
+                                   (All::(('a => 'b) => bool) => bool)
+                                    (%f::'a => 'b.
+  (op =::'b => 'b => bool)
+   ((option_case::'b => ('a => 'b) => 'a option => 'b) u f
+     (None::'a option))
+   u)))
+                             ((op &::bool => bool => bool)
+                               ((All::('b => bool) => bool)
+                                 (%u::'b.
+                                     (All::(('a => 'b) => bool) => bool)
+(%f::'a => 'b.
+    (All::('a => bool) => bool)
+     (%x::'a.
+         (op =::'b => 'b => bool)
+          ((option_case::'b => ('a => 'b) => 'a option => 'b) u f
+            ((Some::'a ~=> 'a) x))
+          (f x)))))
+                               ((op &::bool => bool => bool)
+                                 ((All::(('a => 'b) => bool) => bool)
+                                   (%f::'a => 'b.
+ (All::('a => bool) => bool)
+  (%x::'a.
+      (op =::'b option => 'b option => bool)
+       ((option_map::('a => 'b) => 'a option ~=> 'b) f
+         ((Some::'a ~=> 'a) x))
+       ((Some::'b ~=> 'b) (f x)))))
+                                 ((op &::bool => bool => bool)
+                                   ((All::(('a => 'b) => bool) => bool)
+                                     (%f::'a => 'b.
+   (op =::'b option => 'b option => bool)
+    ((option_map::('a => 'b) => 'a option ~=> 'b) f (None::'a option))
+    (None::'b option)))
+                                   ((op &::bool => bool => bool)
+                                     ((op =::'a option => 'a option => bool)
+ ((OPTION_JOIN::'a option option ~=> 'a) (None::'a option option))
+ (None::'a option))
+                                     ((All::('a option => bool) => bool)
+ (%x::'a option.
+     (op =::'a option => 'a option => bool)
+      ((OPTION_JOIN::'a option option ~=> 'a)
+        ((Some::'a option ~=> 'a option) x))
+      x))))))))))))))))))))"
+  by (import option option_CLAUSES)
+
+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 x y. (option_map f x = Some y) = (EX z. x = Some z & y = f z)"
+  by (import option OPTION_MAP_EQ_SOME)
+
+lemma OPTION_JOIN_EQ_SOME: "ALL x xa. (OPTION_JOIN x = Some xa) = (x = Some (Some xa))"
+  by (import option OPTION_JOIN_EQ_SOME)
+
+lemma option_case_cong: "ALL M M' u f.
+   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'"
+  by (import option option_case_cong)
+
+;end_setup
+
+;setup_theory marker
+
+consts
+  stmarker :: "'a => 'a" 
+
+defs
+  stmarker_primdef: "stmarker == %x. x"
+
+lemma stmarker_def: "ALL x. stmarker x = x"
+  by (import marker stmarker_def)
+
+lemma move_left_conj: "ALL x xa xb.
+   (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 xa xb.
+   (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 xa xb.
+   (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 xa xb.
+   (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)
+
+;end_setup
+
+;setup_theory relation
+
+constdefs
+  TC :: "('a => 'a => bool) => 'a => 'a => bool" 
+  "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"
+
+lemma TC_DEF: "ALL R a b.
+   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)"
+  by (import relation TC_DEF)
+
+constdefs
+  RTC :: "('a => 'a => bool) => 'a => 'a => bool" 
+  "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"
+
+lemma RTC_DEF: "ALL R a b.
+   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)"
+  by (import relation RTC_DEF)
+
+consts
+  RC :: "('a => 'a => bool) => 'a => 'a => bool" 
+
+defs
+  RC_primdef: "RC == %R x y. x = y | R x y"
+
+lemma RC_def: "ALL R x y. RC R x y = (x = y | R x y)"
+  by (import relation RC_def)
+
+consts
+  transitive :: "('a => 'a => bool) => bool" 
+
+defs
+  transitive_primdef: "transitive == %R. ALL x y z. R x y & R y z --> R x z"
+
+lemma transitive_def: "ALL R. transitive R = (ALL x y z. R x y & R y z --> R x z)"
+  by (import relation transitive_def)
+
+constdefs
+  pred_reflexive :: "('a => 'a => bool) => bool" 
+  "pred_reflexive == %R. ALL x. R x x"
+
+lemma reflexive_def: "ALL R. pred_reflexive R = (ALL x. R x x)"
+  by (import relation reflexive_def)
+
+lemma TC_TRANSITIVE: "ALL x. transitive (TC x)"
+  by (import relation TC_TRANSITIVE)
+
+lemma RTC_INDUCT: "ALL x xa.
+   (ALL x. xa x x) & (ALL xb y z. x xb y & xa y z --> xa xb z) -->
+   (ALL xb xc. RTC x xb xc --> xa xb xc)"
+  by (import relation RTC_INDUCT)
+
+lemma TC_RULES: "ALL x.
+   (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)"
+  by (import relation TC_RULES)
+
+lemma RTC_RULES: "ALL x.
+   (ALL xa. RTC x xa xa) &
+   (ALL xa xb xc. x xa xb & RTC x xb xc --> RTC x xa xc)"
+  by (import relation RTC_RULES)
+
+lemma RTC_STRONG_INDUCT: "ALL R P.
+   (ALL x. P x x) & (ALL x y z. R x y & RTC R y z & P y z --> P x z) -->
+   (ALL x y. RTC R x y --> P x y)"
+  by (import relation RTC_STRONG_INDUCT)
+
+lemma RTC_RTC: "ALL R x y. RTC R x y --> (ALL z. RTC R y z --> RTC R x z)"
+  by (import relation RTC_RTC)
+
+lemma RTC_TRANSITIVE: "ALL x. transitive (RTC x)"
+  by (import relation RTC_TRANSITIVE)
+
+lemma RTC_REFLEXIVE: "ALL R. pred_reflexive (RTC R)"
+  by (import relation RTC_REFLEXIVE)
+
+lemma RC_REFLEXIVE: "ALL R. pred_reflexive (RC R)"
+  by (import relation RC_REFLEXIVE)
+
+lemma TC_SUBSET: "ALL x xa xb. x xa xb --> TC x xa xb"
+  by (import relation TC_SUBSET)
+
+lemma RTC_SUBSET: "ALL R x y. R x y --> RTC R x y"
+  by (import relation RTC_SUBSET)
+
+lemma RC_SUBSET: "ALL R x y. R x y --> RC R x y"
+  by (import relation RC_SUBSET)
+
+lemma RC_RTC: "ALL R x y. RC R x y --> RTC R x y"
+  by (import relation RC_RTC)
+
+lemma TC_INDUCT: "ALL x xa.
+   (ALL xb y. x xb y --> xa xb y) &
+   (ALL x y z. xa x y & xa y z --> xa x z) -->
+   (ALL xb xc. TC x xb xc --> xa xb xc)"
+  by (import relation TC_INDUCT)
+
+lemma TC_INDUCT_LEFT1: "ALL x xa.
+   (ALL xb y. x xb y --> xa xb y) &
+   (ALL xb y z. x xb y & xa y z --> xa xb z) -->
+   (ALL xb xc. TC x xb xc --> xa xb xc)"
+  by (import relation TC_INDUCT_LEFT1)
+
+lemma TC_STRONG_INDUCT: "ALL R P.
+   (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) -->
+   (ALL u v. TC R u v --> P u v)"
+  by (import relation TC_STRONG_INDUCT)
+
+lemma TC_STRONG_INDUCT_LEFT1: "ALL R P.
+   (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) -->
+   (ALL u v. TC R u v --> P u v)"
+  by (import relation TC_STRONG_INDUCT_LEFT1)
+
+lemma TC_RTC: "ALL R x y. TC R x y --> RTC R x y"
+  by (import relation TC_RTC)
+
+lemma RTC_TC_RC: "ALL R x y. RTC R x y --> RC R x y | TC R x y"
+  by (import relation RTC_TC_RC)
+
+lemma TC_RC_EQNS: "ALL R. RC (TC R) = RTC R & TC (RC R) = RTC R"
+  by (import relation TC_RC_EQNS)
+
+lemma RC_IDEM: "ALL R. RC (RC R) = RC R"
+  by (import relation RC_IDEM)
+
+lemma TC_IDEM: "ALL R. TC (TC R) = TC R"
+  by (import relation TC_IDEM)
+
+lemma RTC_IDEM: "ALL R. RTC (RTC R) = RTC R"
+  by (import relation RTC_IDEM)
+
+lemma RTC_CASES1: "ALL x xa xb. RTC x xa xb = (xa = xb | (EX u. x xa u & RTC x u xb))"
+  by (import relation RTC_CASES1)
+
+lemma RTC_CASES2: "ALL x xa xb. RTC x xa xb = (xa = xb | (EX u. RTC x xa u & x u xb))"
+  by (import relation RTC_CASES2)
+
+lemma RTC_CASES_RTC_TWICE: "ALL x xa xb. RTC x xa xb = (EX u. RTC x xa u & RTC x u xb)"
+  by (import relation RTC_CASES_RTC_TWICE)
+
+lemma TC_CASES1: "ALL R x z. TC R x z --> R x z | (EX y. R x y & TC R y z)"
+  by (import relation TC_CASES1)
+
+lemma TC_CASES2: "ALL R x z. TC R x z --> R x z | (EX y. TC R x y & R y z)"
+  by (import relation TC_CASES2)
+
+lemma TC_MONOTONE: "ALL R Q. (ALL x y. R x y --> Q x y) --> (ALL x y. TC R x y --> TC Q x y)"
+  by (import relation TC_MONOTONE)
+
+lemma RTC_MONOTONE: "ALL R Q. (ALL x y. R x y --> Q x y) --> (ALL x y. RTC R x y --> RTC Q x y)"
+  by (import relation RTC_MONOTONE)
+
+constdefs
+  WF :: "('a => 'a => bool) => bool" 
+  "WF == %R. ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b))"
+
+lemma WF_DEF: "ALL R. WF R = (ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b)))"
+  by (import relation WF_DEF)
+
+lemma WF_INDUCTION_THM: "ALL R. WF R --> (ALL P. (ALL x. (ALL y. R y x --> P y) --> P x) --> All P)"
+  by (import relation WF_INDUCTION_THM)
+
+lemma WF_NOT_REFL: "ALL x xa xb. WF x --> x xa xb --> xa ~= xb"
+  by (import relation WF_NOT_REFL)
+
+constdefs
+  EMPTY_REL :: "'a => 'a => bool" 
+  "EMPTY_REL == %x y. False"
+
+lemma EMPTY_REL_DEF: "ALL x y. EMPTY_REL x y = False"
+  by (import relation EMPTY_REL_DEF)
+
+lemma WF_EMPTY_REL: "WF EMPTY_REL"
+  by (import relation WF_EMPTY_REL)
+
+lemma WF_SUBSET: "ALL x xa. WF x & (ALL xb y. xa xb y --> x xb y) --> WF xa"
+  by (import relation WF_SUBSET)
+
+lemma WF_TC: "ALL R. WF R --> WF (TC R)"
+  by (import relation WF_TC)
+
+consts
+  inv_image :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" 
+
+defs
+  inv_image_primdef: "relation.inv_image ==
+%(R::'b => 'b => bool) (f::'a => 'b) (x::'a) y::'a. R (f x) (f y)"
+
+lemma inv_image_def: "ALL (R::'b => 'b => bool) f::'a => 'b.
+   relation.inv_image R f = (%(x::'a) y::'a. R (f x) (f y))"
+  by (import relation inv_image_def)
+
+lemma WF_inv_image: "ALL (R::'b => 'b => bool) f::'a => 'b. WF R --> WF (relation.inv_image R f)"
+  by (import relation WF_inv_image)
+
+constdefs
+  RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" 
+  "RESTRICT == %f R x y. if R y x then f y else ARB"
+
+lemma RESTRICT_DEF: "ALL f R x. RESTRICT f R x = (%y. if R y x then f y else ARB)"
+  by (import relation RESTRICT_DEF)
+
+lemma RESTRICT_LEMMA: "ALL x xa xb xc. xa xb xc --> RESTRICT x xa xc xb = x xb"
+  by (import relation RESTRICT_LEMMA)
+
+consts
+  approx :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => ('a => 'b) => bool" 
+
+defs
+  approx_primdef: "approx == %R M x f. f = RESTRICT (%y. M (RESTRICT f R y) y) R x"
+
+lemma approx_def: "ALL R M x f. approx R M x f = (f = RESTRICT (%y. M (RESTRICT f R y) y) R x)"
+  by (import relation approx_def)
+
+consts
+  the_fun :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'a => 'b" 
+
+defs
+  the_fun_primdef: "the_fun == %R M x. Eps (approx R M x)"
+
+lemma the_fun_def: "ALL R M x. the_fun R M x = Eps (approx R M x)"
+  by (import relation the_fun_def)
+
+constdefs
+  WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" 
+  "WFREC ==
+%R M x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x"
+
+lemma WFREC_DEF: "ALL R M.
+   WFREC R M =
+   (%x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x)"
+  by (import relation WFREC_DEF)
+
+lemma WFREC_THM: "ALL R M. WF R --> (ALL x. WFREC R M x = M (RESTRICT (WFREC R M) R x) x)"
+  by (import relation WFREC_THM)
+
+lemma WFREC_COROLLARY: "ALL M R f. f = WFREC R M --> WF R --> (ALL x. f x = M (RESTRICT f R x) x)"
+  by (import relation WFREC_COROLLARY)
+
+lemma WF_RECURSION_THM: "ALL R. WF R --> (ALL M. EX! f. ALL x. f x = M (RESTRICT f R x) x)"
+  by (import relation WF_RECURSION_THM)
+
+;end_setup
+
+;setup_theory pair
+
+lemma CURRY_ONE_ONE_THM: "(curry f = curry g) = (f = g)"
+  by (import pair CURRY_ONE_ONE_THM)
+
+lemma UNCURRY_ONE_ONE_THM: "(split f = split g) = (f = g)"
+  by (import pair UNCURRY_ONE_ONE_THM)
+
+lemma pair_Axiom: "ALL f. EX x. ALL xa y. x (xa, y) = f xa y"
+  by (import pair pair_Axiom)
+
+lemma UNCURRY_CONG: "ALL M M' f.
+   M = M' & (ALL x y. M' = (x, y) --> f x y = f' x y) -->
+   split f M = split f' M'"
+  by (import pair UNCURRY_CONG)
+
+lemma ELIM_PEXISTS: "(EX p. P (fst p) (snd p)) = (EX p1. Ex (P p1))"
+  by (import pair ELIM_PEXISTS)
+
+lemma ELIM_PFORALL: "(ALL p. P (fst p) (snd p)) = (ALL p1. All (P p1))"
+  by (import pair ELIM_PFORALL)
+
+lemma PFORALL_THM: "ALL x. (ALL xa. All (x xa)) = All (split x)"
+  by (import pair PFORALL_THM)
+
+lemma PEXISTS_THM: "ALL x. (EX xa. Ex (x xa)) = Ex (split x)"
+  by (import pair PEXISTS_THM)
+
+lemma LET2_RAND: "ALL (x::'c => 'd) (xa::'a * 'b) xb::'a => 'b => 'c.
+   x (Let xa (split xb)) = (let (xa::'a, y::'b) = xa in x (xb xa y))"
+  by (import pair LET2_RAND)
+
+lemma LET2_RATOR: "ALL (x::'a1 * 'a2) (xa::'a1 => 'a2 => 'b => 'c) xb::'b.
+   Let x (split xa) xb = (let (x::'a1, y::'a2) = x in xa x y xb)"
+  by (import pair LET2_RATOR)
+
+lemma pair_case_cong: "ALL x xa xb.
+   x = xa & (ALL x y. xa = (x, y) --> xb x y = f' x y) -->
+   split xb x = split f' xa"
+  by (import pair pair_case_cong)
+
+constdefs
+  LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
+  "LEX == %R1 R2 (s, t) (u, v). R1 s u | s = u & R2 t v"
+
+lemma LEX_DEF: "ALL R1 R2. LEX R1 R2 = (%(s, t) (u, v). R1 s u | s = u & R2 t v)"
+  by (import pair LEX_DEF)
+
+lemma WF_LEX: "ALL x xa. WF x & WF xa --> WF (LEX x xa)"
+  by (import pair WF_LEX)
+
+constdefs
+  RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
+  "RPROD == %R1 R2 (s, t) (u, v). R1 s u & R2 t v"
+
+lemma RPROD_DEF: "ALL R1 R2. RPROD R1 R2 = (%(s, t) (u, v). R1 s u & R2 t v)"
+  by (import pair RPROD_DEF)
+
+lemma WF_RPROD: "ALL R Q. WF R & WF Q --> WF (RPROD R Q)"
+  by (import pair WF_RPROD)
+
+;end_setup
+
+;setup_theory num
+
+;end_setup
+
+;setup_theory prim_rec
+
+lemma LESS_0_0: "0 < Suc 0"
+  by (import prim_rec LESS_0_0)
+
+lemma LESS_LEMMA1: "ALL x xa. x < Suc xa --> x = xa | x < xa"
+  by (import prim_rec LESS_LEMMA1)
+
+lemma LESS_LEMMA2: "ALL m n. m = n | m < n --> m < Suc n"
+  by (import prim_rec LESS_LEMMA2)
+
+lemma LESS_THM: "ALL m n. (m < Suc n) = (m = n | m < n)"
+  by (import prim_rec LESS_THM)
+
+lemma LESS_SUC_IMP: "ALL x xa. x < Suc xa --> x ~= xa --> x < xa"
+  by (import prim_rec LESS_SUC_IMP)
+
+lemma EQ_LESS: "ALL n. Suc m = 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)
+
+constdefs
+  SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" 
+  "SIMP_REC_REL == %fun x f n. fun 0 = x & (ALL m<n. fun (Suc m) = f (fun m))"
+
+lemma SIMP_REC_REL: "ALL fun x f n.
+   SIMP_REC_REL fun x f n = (fun 0 = x & (ALL m<n. fun (Suc m) = f (fun m)))"
+  by (import prim_rec SIMP_REC_REL)
+
+lemma SIMP_REC_EXISTS: "ALL x f n. EX fun. SIMP_REC_REL fun x f n"
+  by (import prim_rec SIMP_REC_EXISTS)
+
+lemma SIMP_REC_REL_UNIQUE: "ALL x xa xb xc xd xe.
+   SIMP_REC_REL xb x xa xd & SIMP_REC_REL xc x xa xe -->
+   (ALL n. n < xd & n < xe --> xb n = xc n)"
+  by (import prim_rec SIMP_REC_REL_UNIQUE)
+
+lemma SIMP_REC_REL_UNIQUE_RESULT: "ALL x f n. EX! y. EX g. SIMP_REC_REL g x f (Suc n) & y = g n"
+  by (import prim_rec SIMP_REC_REL_UNIQUE_RESULT)
+
+consts
+  SIMP_REC :: "'a => ('a => 'a) => nat => 'a" 
+
+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"
+  by (import prim_rec SIMP_REC)
+
+lemma LESS_SUC_SUC: "ALL m. m < Suc m & m < Suc (Suc m)"
+  by (import prim_rec LESS_SUC_SUC)
+
+lemma SIMP_REC_THM: "ALL x f.
+   SIMP_REC x f 0 = x & (ALL m. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))"
+  by (import prim_rec SIMP_REC_THM)
+
+constdefs
+  PRE :: "nat => nat" 
+  "PRE == %m. if m = 0 then 0 else SOME n. m = Suc n"
+
+lemma PRE_DEF: "ALL m. PRE m = (if m = 0 then 0 else SOME n. m = Suc n)"
+  by (import prim_rec PRE_DEF)
+
+lemma PRE: "PRE 0 = 0 & (ALL m. PRE (Suc m) = m)"
+  by (import prim_rec PRE)
+
+constdefs
+  PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" 
+  "PRIM_REC_FUN == %x f. SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)"
+
+lemma PRIM_REC_FUN: "ALL x f. PRIM_REC_FUN x f = SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)"
+  by (import prim_rec PRIM_REC_FUN)
+
+lemma PRIM_REC_EQN: "ALL x f.
+   (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)"
+  by (import prim_rec PRIM_REC_EQN)
+
+constdefs
+  PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" 
+  "PRIM_REC == %x f m. PRIM_REC_FUN x f m (PRE m)"
+
+lemma PRIM_REC: "ALL x f m. 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 f.
+   PRIM_REC x f 0 = x & (ALL m. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)"
+  by (import prim_rec PRIM_REC_THM)
+
+lemma DC: "ALL P R a.
+   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))))"
+  by (import prim_rec DC)
+
+lemma num_Axiom_old: "ALL e f. EX! fn1. fn1 0 = e & (ALL n. fn1 (Suc n) = f (fn1 n) n)"
+  by (import prim_rec num_Axiom_old)
+
+lemma num_Axiom: "ALL e f. EX x. x 0 = e & (ALL n. x (Suc n) = f n (x n))"
+  by (import prim_rec num_Axiom)
+
+consts
+  wellfounded :: "('a => 'a => bool) => bool" 
+
+defs
+  wellfounded_primdef: "wellfounded == %R. ~ (EX f. ALL n. R (f (Suc n)) (f n))"
+
+lemma wellfounded_def: "ALL R. wellfounded R = (~ (EX f. ALL n. R (f (Suc n)) (f n)))"
+  by (import prim_rec wellfounded_def)
+
+lemma WF_IFF_WELLFOUNDED: "ALL R. WF R = wellfounded R"
+  by (import prim_rec WF_IFF_WELLFOUNDED)
+
+lemma WF_PRED: "WF (%x y. y = Suc x)"
+  by (import prim_rec WF_PRED)
+
+lemma WF_LESS: "(WF::(nat => nat => bool) => bool) (op <::nat => nat => bool)"
+  by (import prim_rec WF_LESS)
+
+consts
+  measure :: "('a => nat) => 'a => 'a => bool" 
+
+defs
+  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. WF (prim_rec.measure x)"
+  by (import prim_rec WF_measure)
+
+lemma measure_thm: "ALL x xa xb. prim_rec.measure x xa xb = (x xa < x xb)"
+  by (import prim_rec measure_thm)
+
+;end_setup
+
+;setup_theory arithmetic
+
+constdefs
+  nat_elim__magic :: "nat => nat" 
+  "nat_elim__magic == %n. n"
+
+lemma nat_elim__magic: "ALL n. nat_elim__magic n = n"
+  by (import arithmetic nat_elim__magic)
+
+consts
+  EVEN :: "nat => bool" 
+
+specification (EVEN) EVEN: "EVEN 0 = True & (ALL n. EVEN (Suc n) = (~ EVEN n))"
+  by (import arithmetic EVEN)
+
+consts
+  ODD :: "nat => bool" 
+
+specification (ODD) ODD: "ODD 0 = False & (ALL n. ODD (Suc n) = (~ ODD n))"
+  by (import arithmetic ODD)
+
+lemma TWO: "2 = Suc 1"
+  by (import arithmetic TWO)
+
+lemma NORM_0: "(0::nat) = (0::nat)"
+  by (import arithmetic NORM_0)
+
+lemma num_case_compute: "ALL n. nat_case f g n = (if n = 0 then f else g (PRE n))"
+  by (import arithmetic num_case_compute)
+
+lemma ADD_CLAUSES: "0 + m = m & m + 0 = m & Suc m + n = 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 xa. ~ (x < xa & xa < Suc x)"
+  by (import arithmetic LESS_LESS_SUC)
+
+lemma FUN_EQ_LEMMA: "ALL f x1 x2. f x1 & ~ f x2 --> x1 ~= x2"
+  by (import arithmetic FUN_EQ_LEMMA)
+
+lemma LESS_NOT_SUC: "ALL m n. m < n & n ~= Suc m --> Suc m < n"
+  by (import arithmetic LESS_NOT_SUC)
+
+lemma LESS_0_CASES: "ALL m::nat. (0::nat) = m | (0::nat) < 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. m <= Suc m"
+  by (import arithmetic LESS_EQ_SUC_REFL)
+
+lemma LESS_ADD_NONZERO: "ALL (m::nat) n::nat. n ~= (0::nat) --> 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::nat) - m = (0::nat) & m - (0::nat) = m"
+  by (import arithmetic SUB_0)
+
+lemma SUC_SUB1: "ALL m. Suc m - 1 = m"
+  by (import arithmetic SUC_SUB1)
+
+lemma PRE_SUB1: "ALL m. PRE m = m - 1"
+  by (import arithmetic PRE_SUB1)
+
+lemma MULT_CLAUSES: "ALL x xa.
+   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 n. PRE (m - n) = PRE m - n"
+  by (import arithmetic PRE_SUB)
+
+lemma ADD_EQ_1: "ALL (m::nat) n::nat.
+   (m + n = (1::nat)) =
+   (m = (1::nat) & n = (0::nat) | m = (0::nat) & n = (1::nat))"
+  by (import arithmetic ADD_EQ_1)
+
+lemma ADD_INV_0_EQ: "ALL (m::nat) n::nat. (m + n = m) = (n = (0::nat))"
+  by (import arithmetic ADD_INV_0_EQ)
+
+lemma PRE_SUC_EQ: "ALL m n. 0 < n --> (m = PRE n) = (Suc m = n)"
+  by (import arithmetic PRE_SUC_EQ)
+
+lemma INV_PRE_EQ: "ALL m n. 0 < m & 0 < n --> (PRE m = PRE n) = (m = n)"
+  by (import arithmetic INV_PRE_EQ)
+
+lemma LESS_SUC_NOT: "ALL m n. 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::nat)))"
+  by (import arithmetic LESS_ADD_1)
+
+lemma NOT_ODD_EQ_EVEN: "ALL n m. Suc (n + n) ~= m + m"
+  by (import arithmetic NOT_ODD_EQ_EVEN)
+
+lemma MULT_SUC_EQ: "ALL p m n. (n * Suc p = m * Suc p) = (n = m)"
+  by (import arithmetic MULT_SUC_EQ)
+
+lemma MULT_EXP_MONO: "ALL p q n m. (n * Suc q ^ p = m * Suc q ^ p) = (n = m)"
+  by (import arithmetic MULT_EXP_MONO)
+
+lemma LESS_ADD_SUC: "ALL m n. 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 P::nat => bool. Ex P --> (EX n::nat. P n & (ALL m<n. ~ P m))"
+  by (import arithmetic WOP)
+
+lemma INV_PRE_LESS: "ALL m. 0 < m --> (ALL n. (PRE m < PRE n) = (m < n))"
+  by (import arithmetic INV_PRE_LESS)
+
+lemma INV_PRE_LESS_EQ: "ALL n. 0 < n --> (ALL m. (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::nat) | n = (0::nat))"
+  by (import arithmetic SUB_EQ_EQ_0)
+
+lemma SUB_LESS_OR: "ALL (m::nat) n::nat. n < m --> n <= m - (1::nat)"
+  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 xa. (~ 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 n. Suc n ^ m ~= 0"
+  by (import arithmetic NOT_EXP_0)
+
+lemma ZERO_LESS_EXP: "ALL m n. 0 < Suc n ^ m"
+  by (import arithmetic ZERO_LESS_EXP)
+
+lemma ODD_OR_EVEN: "ALL x. EX xa. x = Suc (Suc 0) * xa | x = Suc (Suc 0) * xa + 1"
+  by (import arithmetic ODD_OR_EVEN)
+
+lemma LESS_EXP_SUC_MONO: "ALL n m. 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 LESS_EQ_EXISTS: "ALL (m::nat) n::nat. (m <= n) = (EX p::nat. n = m + p)"
+  by (import arithmetic LESS_EQ_EXISTS)
+
+lemma MULT_EQ_1: "ALL (x::nat) y::nat. (x * y = (1::nat)) = (x = (1::nat) & y = (1::nat))"
+  by (import arithmetic MULT_EQ_1)
+
+consts
+  FACT :: "nat => nat" 
+
+specification (FACT) FACT: "FACT 0 = 1 & (ALL n. FACT (Suc n) = Suc n * FACT n)"
+  by (import arithmetic FACT)
+
+lemma FACT_LESS: "ALL n. 0 < FACT n"
+  by (import arithmetic FACT_LESS)
+
+lemma EVEN_ODD: "ALL n. EVEN n = (~ ODD n)"
+  by (import arithmetic EVEN_ODD)
+
+lemma ODD_EVEN: "ALL x. ODD x = (~ EVEN x)"
+  by (import arithmetic ODD_EVEN)
+
+lemma EVEN_OR_ODD: "ALL x. EVEN x | ODD x"
+  by (import arithmetic EVEN_OR_ODD)
+
+lemma EVEN_AND_ODD: "ALL x. ~ (EVEN x & ODD x)"
+  by (import arithmetic EVEN_AND_ODD)
+
+lemma EVEN_ADD: "ALL m n. EVEN (m + n) = (EVEN m = EVEN n)"
+  by (import arithmetic EVEN_ADD)
+
+lemma EVEN_MULT: "ALL m n. EVEN (m * n) = (EVEN m | EVEN n)"
+  by (import arithmetic EVEN_MULT)
+
+lemma ODD_ADD: "ALL m n. ODD (m + n) = (ODD m ~= ODD n)"
+  by (import arithmetic ODD_ADD)
+
+lemma ODD_MULT: "ALL m n. ODD (m * n) = (ODD m & ODD n)"
+  by (import arithmetic ODD_MULT)
+
+lemma EVEN_DOUBLE: "ALL n. EVEN (2 * n)"
+  by (import arithmetic EVEN_DOUBLE)
+
+lemma ODD_DOUBLE: "ALL x. ODD (Suc (2 * x))"
+  by (import arithmetic ODD_DOUBLE)
+
+lemma EVEN_ODD_EXISTS: "ALL x. (EVEN x --> (EX m. x = 2 * m)) & (ODD x --> (EX m. x = Suc (2 * m)))"
+  by (import arithmetic EVEN_ODD_EXISTS)
+
+lemma EVEN_EXISTS: "ALL n. EVEN n = (EX m. n = 2 * m)"
+  by (import arithmetic EVEN_EXISTS)
+
+lemma ODD_EXISTS: "ALL n. ODD n = (EX m. n = Suc (2 * m))"
+  by (import arithmetic ODD_EXISTS)
+
+lemma NOT_SUC_LESS_EQ_0: "ALL x. ~ Suc x <= 0"
+  by (import arithmetic NOT_SUC_LESS_EQ_0)
+
+lemma NOT_LEQ: "ALL x xa. (~ x <= xa) = (Suc xa <= x)"
+  by (import arithmetic NOT_LEQ)
+
+lemma NOT_NUM_EQ: "ALL x xa. (x ~= xa) = (Suc x <= xa | Suc xa <= x)"
+  by (import arithmetic NOT_NUM_EQ)
+
+lemma NOT_GREATER_EQ: "ALL x xa. (~ xa <= x) = (Suc x <= xa)"
+  by (import arithmetic NOT_GREATER_EQ)
+
+lemma SUC_ADD_SYM: "ALL m n. Suc (m + n) = Suc n + m"
+  by (import arithmetic SUC_ADD_SYM)
+
+lemma NOT_SUC_ADD_LESS_EQ: "ALL m n. ~ 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 n. 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::nat))"
+  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::nat) < 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::nat))"
+  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::nat) < 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::nat) & 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::nat))"
+  by (import arithmetic SUB_RIGHT_EQ)
+
+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::nat) < n --> (EX (x::nat) q::nat. k = q * n + x & x < n)"
+  by (import arithmetic DA)
+
+lemma DIV_LESS_EQ: "ALL n::nat. (0::nat) < n --> (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::nat. (0::nat) < n --> (ALL k::nat. k * n mod n = (0::nat))"
+  by (import arithmetic MOD_EQ_0)
+
+lemma ZERO_MOD: "ALL n::nat. (0::nat) < n --> (0::nat) mod n = (0::nat)"
+  by (import arithmetic ZERO_MOD)
+
+lemma ZERO_DIV: "ALL n::nat. (0::nat) < n --> (0::nat) div 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::nat.
+   (0::nat) < n --> (ALL (q::nat) r::nat. (q * n + r) mod n = r mod n)"
+  by (import arithmetic MOD_TIMES)
+
+lemma MOD_PLUS: "ALL n::nat.
+   (0::nat) < n -->
+   (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::nat. (0::nat) < n --> (ALL k::nat. k mod n mod n = k mod n)"
+  by (import arithmetic MOD_MOD)
+
+lemma ADD_DIV_ADD_DIV: "ALL x::nat.
+   (0::nat) < x -->
+   (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::nat) < n & (0::nat) < m -->
+   (ALL x::nat. x mod (n * m) mod n = x mod n)"
+  by (import arithmetic MOD_MULT_MOD)
+
+lemma DIVMOD_ID: "ALL n::nat. (0::nat) < n --> n div n = (1::nat) & n mod n = (0::nat)"
+  by (import arithmetic DIVMOD_ID)
+
+lemma DIV_DIV_DIV_MULT: "ALL (x::nat) xa::nat.
+   (0::nat) < x & (0::nat) < 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::nat) < 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::nat) < 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::nat.
+   (0::nat) < n -->
+   (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::nat) < n & (0::nat) < q --> n * (p mod q) = n * p mod (n * q)"
+  by (import arithmetic MOD_COMMON_FACTOR)
+
+lemma num_case_cong: "ALL M M' b f.
+   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'"
+  by (import arithmetic num_case_cong)
+
+lemma SUC_ELIM_THM: "ALL P. (ALL n. P (Suc n) n) = (ALL n. 0 < n --> P n (n - 1))"
+  by (import arithmetic SUC_ELIM_THM)
+
+lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
+(ALL x::nat. (b = a + x --> P (0::nat)) & (a = b + x --> P x))"
+  by (import arithmetic SUB_ELIM_THM)
+
+lemma PRE_ELIM_THM: "P (PRE n) = (ALL m. (n = 0 --> P 0) & (n = Suc m --> P m))"
+  by (import arithmetic PRE_ELIM_THM)
+
+lemma MULT_INCREASES: "ALL m n. 1 < m & 0 < n --> Suc n <= m * n"
+  by (import arithmetic MULT_INCREASES)
+
+lemma EXP_ALWAYS_BIG_ENOUGH: "ALL b::nat. (1::nat) < b --> (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::nat)) = (n = (0::nat) & (0::nat) < m)"
+  by (import arithmetic EXP_EQ_0)
+
+lemma EXP_1: "ALL x::nat. (1::nat) ^ x = (1::nat) & x ^ (1::nat) = x"
+  by (import arithmetic EXP_1)
+
+lemma EXP_EQ_1: "ALL (n::nat) m::nat. (n ^ m = (1::nat)) = (n = (1::nat) | m = (0::nat))"
+  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::nat) = (0::nat) & min (0::nat) x = (0::nat)"
+  by (import arithmetic MIN_0)
+
+lemma MAX_0: "ALL x::nat. max x (0::nat) = x & max (0::nat) 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)
+
+;end_setup
+
+;setup_theory hrat
+
+constdefs
+  trat_1 :: "nat * nat" 
+  "trat_1 == (0, 0)"
+
+lemma trat_1: "trat_1 = (0, 0)"
+  by (import hrat trat_1)
+
+constdefs
+  trat_inv :: "nat * nat => nat * nat" 
+  "trat_inv == %(x, y). (y, x)"
+
+lemma trat_inv: "ALL x y. trat_inv (x, y) = (y, x)"
+  by (import hrat trat_inv)
+
+constdefs
+  trat_add :: "nat * nat => nat * nat => nat * nat" 
+  "trat_add ==
+%(x, y) (x', y').
+   (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
+
+lemma trat_add: "ALL x y x' y'.
+   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)
+
+constdefs
+  trat_mul :: "nat * nat => nat * nat => nat * nat" 
+  "trat_mul == %(x, y) (x', y'). (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
+
+lemma trat_mul: "ALL x y x' y'.
+   trat_mul (x, y) (x', y') = (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
+  by (import hrat trat_mul)
+
+consts
+  trat_sucint :: "nat => nat * nat" 
+
+specification (trat_sucint) trat_sucint: "trat_sucint 0 = trat_1 &
+(ALL n. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)"
+  by (import hrat trat_sucint)
+
+constdefs
+  trat_eq :: "nat * nat => nat * nat => bool" 
+  "trat_eq == %(x, y) (x', y'). Suc x * Suc y' = Suc x' * Suc y"
+
+lemma trat_eq: "ALL x y x' y'. 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. trat_eq p p"
+  by (import hrat TRAT_EQ_REFL)
+
+lemma TRAT_EQ_SYM: "ALL p q. trat_eq p q = trat_eq q p"
+  by (import hrat TRAT_EQ_SYM)
+
+lemma TRAT_EQ_TRANS: "ALL p q r. trat_eq p q & trat_eq q r --> trat_eq p r"
+  by (import hrat TRAT_EQ_TRANS)
+
+lemma TRAT_EQ_AP: "ALL p q. p = q --> trat_eq p q"
+  by (import hrat TRAT_EQ_AP)
+
+lemma TRAT_ADD_SYM_EQ: "ALL h i. trat_add h i = trat_add i h"
+  by (import hrat TRAT_ADD_SYM_EQ)
+
+lemma TRAT_MUL_SYM_EQ: "ALL h i. trat_mul h i = trat_mul i h"
+  by (import hrat TRAT_MUL_SYM_EQ)
+
+lemma TRAT_INV_WELLDEFINED: "ALL p q. trat_eq p q --> trat_eq (trat_inv p) (trat_inv q)"
+  by (import hrat TRAT_INV_WELLDEFINED)
+
+lemma TRAT_ADD_WELLDEFINED: "ALL p q r. 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 p2 q1 q2.
+   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 q r. 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 p2 q1 q2.
+   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 i. trat_eq (trat_add h i) (trat_add i h)"
+  by (import hrat TRAT_ADD_SYM)
+
+lemma TRAT_ADD_ASSOC: "ALL h i j. 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 i. trat_eq (trat_mul h i) (trat_mul i h)"
+  by (import hrat TRAT_MUL_SYM)
+
+lemma TRAT_MUL_ASSOC: "ALL h i j. 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 i j.
+   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. trat_eq (trat_mul trat_1 h) h"
+  by (import hrat TRAT_MUL_LID)
+
+lemma TRAT_MUL_LINV: "ALL h. trat_eq (trat_mul (trat_inv h) h) trat_1"
+  by (import hrat TRAT_MUL_LINV)
+
+lemma TRAT_NOZERO: "ALL h i. ~ trat_eq (trat_add h i) h"
+  by (import hrat TRAT_NOZERO)
+
+lemma TRAT_ADD_TOTAL: "ALL h i.
+   trat_eq h i |
+   (EX d. trat_eq h (trat_add i d)) | (EX d. trat_eq i (trat_add h d))"
+  by (import hrat TRAT_ADD_TOTAL)
+
+lemma TRAT_SUCINT_0: "ALL n. trat_eq (trat_sucint n) (n, 0)"
+  by (import hrat TRAT_SUCINT_0)
+
+lemma TRAT_ARCH: "ALL h. EX n d. trat_eq (trat_sucint n) (trat_add h d)"
+  by (import hrat TRAT_ARCH)
+
+lemma TRAT_SUCINT: "trat_eq (trat_sucint 0) trat_1 &
+(ALL n. trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))"
+  by (import hrat TRAT_SUCINT)
+
+lemma TRAT_EQ_EQUIV: "ALL p q. trat_eq p q = (trat_eq p = trat_eq q)"
+  by (import hrat TRAT_EQ_EQUIV)
+
+typedef (open) hrat = "{x. EX xa. x = trat_eq xa}" 
+  by (rule typedef_helper,import hrat hrat_TY_DEF)
+
+lemmas hrat_TY_DEF = typedef_hol2hol4 [OF type_definition_hrat]
+
+consts
+  mk_hrat :: "(nat * nat => bool) => hrat" 
+  dest_hrat :: "hrat => nat * nat => bool" 
+
+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))"
+  by (import hrat hrat_tybij)
+
+constdefs
+  hrat_1 :: "hrat" 
+  "hrat_1 == mk_hrat (trat_eq trat_1)"
+
+lemma hrat_1: "hrat_1 = mk_hrat (trat_eq trat_1)"
+  by (import hrat hrat_1)
+
+constdefs
+  hrat_inv :: "hrat => hrat" 
+  "hrat_inv == %T1. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
+
+lemma hrat_inv: "ALL T1. hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
+  by (import hrat hrat_inv)
+
+constdefs
+  hrat_add :: "hrat => hrat => hrat" 
+  "hrat_add ==
+%T1 T2.
+   mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+
+lemma hrat_add: "ALL T1 T2.
+   hrat_add T1 T2 =
+   mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+  by (import hrat hrat_add)
+
+constdefs
+  hrat_mul :: "hrat => hrat => hrat" 
+  "hrat_mul ==
+%T1 T2.
+   mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+
+lemma hrat_mul: "ALL T1 T2.
+   hrat_mul T1 T2 =
+   mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
+  by (import hrat hrat_mul)
+
+constdefs
+  hrat_sucint :: "nat => hrat" 
+  "hrat_sucint == %T1. mk_hrat (trat_eq (trat_sucint T1))"
+
+lemma hrat_sucint: "ALL T1. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))"
+  by (import hrat hrat_sucint)
+
+lemma HRAT_ADD_SYM: "ALL h i. hrat_add h i = hrat_add i h"
+  by (import hrat HRAT_ADD_SYM)
+
+lemma HRAT_ADD_ASSOC: "ALL h i j. 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 i. hrat_mul h i = hrat_mul i h"
+  by (import hrat HRAT_MUL_SYM)
+
+lemma HRAT_MUL_ASSOC: "ALL h i j. 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 i j.
+   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_mul hrat_1 h = h"
+  by (import hrat HRAT_MUL_LID)
+
+lemma HRAT_MUL_LINV: "ALL h. hrat_mul (hrat_inv h) h = hrat_1"
+  by (import hrat HRAT_MUL_LINV)
+
+lemma HRAT_NOZERO: "ALL h i. hrat_add h i ~= h"
+  by (import hrat HRAT_NOZERO)
+
+lemma HRAT_ADD_TOTAL: "ALL h i. h = i | (EX x. h = hrat_add i x) | (EX x. i = hrat_add h x)"
+  by (import hrat HRAT_ADD_TOTAL)
+
+lemma HRAT_ARCH: "ALL h. EX x xa. hrat_sucint x = hrat_add h xa"
+  by (import hrat HRAT_ARCH)
+
+lemma HRAT_SUCINT: "hrat_sucint 0 = hrat_1 &
+(ALL x. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)"
+  by (import hrat HRAT_SUCINT)
+
+;end_setup
+
+;setup_theory hreal
+
+constdefs
+  hrat_lt :: "hrat => hrat => bool" 
+  "hrat_lt == %x y. EX d. y = hrat_add x d"
+
+lemma hrat_lt: "ALL x y. hrat_lt x y = (EX d. y = hrat_add x d)"
+  by (import hreal hrat_lt)
+
+lemma HRAT_LT_REFL: "ALL x. ~ hrat_lt x x"
+  by (import hreal HRAT_LT_REFL)
+
+lemma HRAT_LT_TRANS: "ALL x y z. hrat_lt x y & hrat_lt y z --> hrat_lt x z"
+  by (import hreal HRAT_LT_TRANS)
+
+lemma HRAT_LT_ANTISYM: "ALL x y. ~ (hrat_lt x y & hrat_lt y x)"
+  by (import hreal HRAT_LT_ANTISYM)
+
+lemma HRAT_LT_TOTAL: "ALL x y. x = y | hrat_lt x y | hrat_lt y x"
+  by (import hreal HRAT_LT_TOTAL)
+
+lemma HRAT_MUL_RID: "ALL x. hrat_mul x hrat_1 = x"
+  by (import hreal HRAT_MUL_RID)
+
+lemma HRAT_MUL_RINV: "ALL x. hrat_mul x (hrat_inv x) = hrat_1"
+  by (import hreal HRAT_MUL_RINV)
+
+lemma HRAT_RDISTRIB: "ALL x y z.
+   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 y. hrat_lt x (hrat_add x y)"
+  by (import hreal HRAT_LT_ADDL)
+
+lemma HRAT_LT_ADDR: "ALL x xa. hrat_lt xa (hrat_add x xa)"
+  by (import hreal HRAT_LT_ADDR)
+
+lemma HRAT_LT_GT: "ALL x y. hrat_lt x y --> ~ hrat_lt y x"
+  by (import hreal HRAT_LT_GT)
+
+lemma HRAT_LT_NE: "ALL x y. hrat_lt x y --> x ~= y"
+  by (import hreal HRAT_LT_NE)
+
+lemma HRAT_EQ_LADD: "ALL x y z. (hrat_add x y = hrat_add x z) = (y = z)"
+  by (import hreal HRAT_EQ_LADD)
+
+lemma HRAT_EQ_LMUL: "ALL x y z. (hrat_mul x y = hrat_mul x z) = (y = z)"
+  by (import hreal HRAT_EQ_LMUL)
+
+lemma HRAT_LT_ADD2: "ALL u v x y.
+   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 y z. 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 y z. 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 v x y.
+   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 y z. 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 y z. 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 y. hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1"
+  by (import hreal HRAT_LT_LMUL1)
+
+lemma HRAT_LT_RMUL1: "ALL x y. hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1"
+  by (import hreal HRAT_LT_RMUL1)
+
+lemma HRAT_GT_LMUL1: "ALL x y. hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x"
+  by (import hreal HRAT_GT_LMUL1)
+
+lemma HRAT_LT_L1: "ALL x y. 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 y. 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 y. 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 y. 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. Ex (hrat_lt x)"
+  by (import hreal HRAT_UP)
+
+lemma HRAT_DOWN: "ALL x. EX xa. hrat_lt xa x"
+  by (import hreal HRAT_DOWN)
+
+lemma HRAT_DOWN2: "ALL x y. EX xa. hrat_lt xa x & hrat_lt xa y"
+  by (import hreal HRAT_DOWN2)
+
+lemma HRAT_MEAN: "ALL x y. hrat_lt x y --> (EX xa. hrat_lt x xa & hrat_lt xa y)"
+  by (import hreal HRAT_MEAN)
+
+constdefs
+  isacut :: "(hrat => bool) => bool" 
+  "isacut ==
+%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: "ALL C.
+   isacut 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)))"
+  by (import hreal isacut)
+
+constdefs
+  cut_of_hrat :: "hrat => hrat => bool" 
+  "cut_of_hrat == %x y. hrat_lt y x"
+
+lemma cut_of_hrat: "ALL x. cut_of_hrat x = (%y. hrat_lt y x)"
+  by (import hreal cut_of_hrat)
+
+lemma ISACUT_HRAT: "ALL h. isacut (cut_of_hrat h)"
+  by (import hreal ISACUT_HRAT)
+
+typedef (open) hreal = "Collect isacut" 
+  by (rule typedef_helper,import hreal hreal_TY_DEF)
+
+lemmas hreal_TY_DEF = typedef_hol2hol4 [OF type_definition_hreal]
+
+consts
+  hreal :: "(hrat => bool) => hreal" 
+  cut :: "hreal => hrat => bool" 
+
+specification (cut hreal) hreal_tybij: "(ALL a. hreal (hreal.cut a) = a) &
+(ALL r. isacut r = (hreal.cut (hreal r) = r))"
+  by (import hreal hreal_tybij)
+
+lemma EQUAL_CUTS: "ALL X Y. hreal.cut X = hreal.cut Y --> X = Y"
+  by (import hreal EQUAL_CUTS)
+
+lemma CUT_ISACUT: "ALL x. isacut (hreal.cut x)"
+  by (import hreal CUT_ISACUT)
+
+lemma CUT_NONEMPTY: "ALL x. Ex (hreal.cut x)"
+  by (import hreal CUT_NONEMPTY)
+
+lemma CUT_BOUNDED: "ALL x. EX xa. ~ hreal.cut x xa"
+  by (import hreal CUT_BOUNDED)
+
+lemma CUT_DOWN: "ALL x xa xb. hreal.cut x xa & hrat_lt xb xa --> hreal.cut x xb"
+  by (import hreal CUT_DOWN)
+
+lemma CUT_UP: "ALL x xa. hreal.cut x xa --> (EX y. hreal.cut x y & hrat_lt xa y)"
+  by (import hreal CUT_UP)
+
+lemma CUT_UBOUND: "ALL x xa xb. ~ hreal.cut x xa & hrat_lt xa xb --> ~ hreal.cut x xb"
+  by (import hreal CUT_UBOUND)
+
+lemma CUT_STRADDLE: "ALL X x y. hreal.cut X x & ~ hreal.cut X y --> hrat_lt x y"
+  by (import hreal CUT_STRADDLE)
+
+lemma CUT_NEARTOP_ADD: "ALL X e. EX x. hreal.cut X x & ~ hreal.cut X (hrat_add x e)"
+  by (import hreal CUT_NEARTOP_ADD)
+
+lemma CUT_NEARTOP_MUL: "ALL X u.
+   hrat_lt hrat_1 u --> (EX x. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))"
+  by (import hreal CUT_NEARTOP_MUL)
+
+constdefs
+  hreal_1 :: "hreal" 
+  "hreal_1 == hreal (cut_of_hrat hrat_1)"
+
+lemma hreal_1: "hreal_1 = hreal (cut_of_hrat hrat_1)"
+  by (import hreal hreal_1)
+
+constdefs
+  hreal_add :: "hreal => hreal => hreal" 
+  "hreal_add ==
+%X Y. hreal (%w. EX x y. w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
+
+lemma hreal_add: "ALL X Y.
+   hreal_add X Y =
+   hreal (%w. EX x y. w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
+  by (import hreal hreal_add)
+
+constdefs
+  hreal_mul :: "hreal => hreal => hreal" 
+  "hreal_mul ==
+%X Y. hreal (%w. EX x y. w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
+
+lemma hreal_mul: "ALL X Y.
+   hreal_mul X Y =
+   hreal (%w. EX x y. w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
+  by (import hreal hreal_mul)
+
+constdefs
+  hreal_inv :: "hreal => hreal" 
+  "hreal_inv ==
+%X. hreal
+     (%w. EX d. hrat_lt d hrat_1 &
+                (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
+
+lemma hreal_inv: "ALL X.
+   hreal_inv X =
+   hreal
+    (%w. EX d. hrat_lt d hrat_1 &
+               (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
+  by (import hreal hreal_inv)
+
+constdefs
+  hreal_sup :: "(hreal => bool) => hreal" 
+  "hreal_sup == %P. hreal (%w. EX X. P X & hreal.cut X w)"
+
+lemma hreal_sup: "ALL P. hreal_sup P = hreal (%w. EX X. P X & hreal.cut X w)"
+  by (import hreal hreal_sup)
+
+constdefs
+  hreal_lt :: "hreal => hreal => bool" 
+  "hreal_lt == %X Y. X ~= Y & (ALL x. hreal.cut X x --> hreal.cut Y x)"
+
+lemma hreal_lt: "ALL X Y. hreal_lt X Y = (X ~= Y & (ALL x. hreal.cut X x --> hreal.cut Y x))"
+  by (import hreal hreal_lt)
+
+lemma HREAL_INV_ISACUT: "ALL X.
+   isacut
+    (%w. EX d. hrat_lt d hrat_1 &
+               (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
+  by (import hreal HREAL_INV_ISACUT)
+
+lemma HREAL_ADD_ISACUT: "ALL X Y.
+   isacut (%w. EX x y. 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 Y.
+   isacut (%w. EX x y. 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 Y. hreal_add X Y = hreal_add Y X"
+  by (import hreal HREAL_ADD_SYM)
+
+lemma HREAL_MUL_SYM: "ALL X Y. hreal_mul X Y = hreal_mul Y X"
+  by (import hreal HREAL_MUL_SYM)
+
+lemma HREAL_ADD_ASSOC: "ALL X Y Z. 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 Y Z. 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 Y Z.
+   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_mul hreal_1 X = X"
+  by (import hreal HREAL_MUL_LID)
+
+lemma HREAL_MUL_LINV: "ALL X. hreal_mul (hreal_inv X) X = hreal_1"
+  by (import hreal HREAL_MUL_LINV)
+
+lemma HREAL_NOZERO: "ALL X Y. hreal_add X Y ~= X"
+  by (import hreal HREAL_NOZERO)
+
+constdefs
+  hreal_sub :: "hreal => hreal => hreal" 
+  "hreal_sub ==
+%Y X. hreal (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
+
+lemma hreal_sub: "ALL Y X.
+   hreal_sub Y X =
+   hreal (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
+  by (import hreal hreal_sub)
+
+lemma HREAL_LT_LEMMA: "ALL X Y. hreal_lt X Y --> (EX x. ~ hreal.cut X x & hreal.cut Y x)"
+  by (import hreal HREAL_LT_LEMMA)
+
+lemma HREAL_SUB_ISACUT: "ALL X Y.
+   hreal_lt X Y -->
+   isacut (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
+  by (import hreal HREAL_SUB_ISACUT)
+
+lemma HREAL_SUB_ADD: "ALL X Y. hreal_lt X Y --> hreal_add (hreal_sub Y X) X = Y"
+  by (import hreal HREAL_SUB_ADD)
+
+lemma HREAL_LT_TOTAL: "ALL X Y. X = Y | hreal_lt X Y | hreal_lt Y X"
+  by (import hreal HREAL_LT_TOTAL)
+
+lemma HREAL_LT: "ALL X Y. hreal_lt X Y = (EX D. Y = hreal_add X D)"
+  by (import hreal HREAL_LT)
+
+lemma HREAL_ADD_TOTAL: "ALL X Y. X = Y | (EX D. Y = hreal_add X D) | (EX D. X = hreal_add Y D)"
+  by (import hreal HREAL_ADD_TOTAL)
+
+lemma HREAL_SUP_ISACUT: "ALL P.
+   Ex P & (EX Y. ALL X. P X --> hreal_lt X Y) -->
+   isacut (%w. EX X. P X & hreal.cut X w)"
+  by (import hreal HREAL_SUP_ISACUT)
+
+lemma HREAL_SUP: "ALL P.
+   Ex P & (EX Y. ALL X. P X --> hreal_lt X Y) -->
+   (ALL Y. (EX X. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P))"
+  by (import hreal HREAL_SUP)
+
+;end_setup
+
+;setup_theory numeral
+
+lemma numeral_suc: "Suc ALT_ZERO = NUMERAL_BIT1 ALT_ZERO &
+(ALL x. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) &
+(ALL x. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))"
+  by (import numeral numeral_suc)
+
+constdefs
+  iZ :: "nat => nat" 
+  "iZ == %x. x"
+
+lemma iZ: "ALL x. iZ x = x"
+  by (import numeral iZ)
+
+constdefs
+  iiSUC :: "nat => nat" 
+  "iiSUC == %n. Suc (Suc n)"
+
+lemma iiSUC: "ALL n. iiSUC n = Suc (Suc n)"
+  by (import numeral iiSUC)
+
+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)"
+  by (import numeral numeral_distrib)
+
+lemma numeral_iisuc: "iiSUC ALT_ZERO = NUMERAL_BIT2 ALT_ZERO &
+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 xa.
+   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 xa.
+   (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 xa.
+   (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 xa.
+   (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)
+
+lemma numeral_pre: "PRE ALT_ZERO = ALT_ZERO &
+PRE (NUMERAL_BIT1 ALT_ZERO) = ALT_ZERO &
+(ALL x.
+    PRE (NUMERAL_BIT1 (NUMERAL_BIT1 x)) =
+    NUMERAL_BIT2 (PRE (NUMERAL_BIT1 x))) &
+(ALL x.
+    PRE (NUMERAL_BIT1 (NUMERAL_BIT2 x)) = NUMERAL_BIT2 (NUMERAL_BIT1 x)) &
+(ALL x. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)"
+  by (import numeral numeral_pre)
+
+lemma bit_initiality: "ALL zf b1f b2f.
+   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))"
+  by (import numeral bit_initiality)
+
+consts
+  iBIT_cases :: "nat => 'a => (nat => 'a) => (nat => 'a) => 'a" 
+
+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) (bf1::nat => 'a) bf2::nat => 'a.
+    iBIT_cases (NUMERAL_BIT1 n) zf bf1 bf2 = bf1 n) &
+(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)
+
+constdefs
+  iDUB :: "nat => nat" 
+  "iDUB == %x. x + x"
+
+lemma iDUB: "ALL x. iDUB x = x + x"
+  by (import numeral iDUB)
+
+consts
+  iSUB :: "bool => nat => nat => 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. 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. 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))))"
+  by (import numeral iSUB_DEF)
+
+lemma bit_induction: "ALL P.
+   P ALT_ZERO &
+   (ALL n. P n --> P (NUMERAL_BIT1 n)) &
+   (ALL n. P n --> P (NUMERAL_BIT2 n)) -->
+   All P"
+  by (import numeral bit_induction)
+
+lemma iSUB_THM: "ALL xa xb xc.
+   iSUB xa ALT_ZERO x = 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 xa.
+   NUMERAL (x - xa) = (if xa < x then NUMERAL (iSUB True x xa) else 0)"
+  by (import numeral numeral_sub)
+
+lemma iDUB_removal: "ALL x.
+   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 xa.
+   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)
+
+constdefs
+  iSQR :: "nat => nat" 
+  "iSQR == %x. x * x"
+
+lemma iSQR: "ALL x. iSQR x = x * x"
+  by (import numeral iSQR)
+
+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))"
+  by (import numeral numeral_exp)
+
+lemma numeral_evenodd: "ALL x.
+   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. FACT n = (if n = 0 then 1 else n * FACT (PRE n))"
+  by (import numeral numeral_fact)
+
+lemma numeral_funpow: "ALL n. (f ^ n) x = (if n = 0 then x else (f ^ (n - 1)) (f x))"
+  by (import numeral numeral_funpow)
+
+;end_setup
+
+;setup_theory ind_type
+
+lemma INJ_INVERSE2: "ALL P::'A => 'B => 'C.
+   (ALL (x1::'A) (y1::'B) (x2::'A) y2::'B.
+       (P 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)"
+  by (import ind_type INJ_INVERSE2)
+
+constdefs
+  NUMPAIR :: "nat => nat => nat" 
+  "NUMPAIR == %x y. 2 ^ x * (2 * y + 1)"
+
+lemma NUMPAIR: "ALL x y. NUMPAIR x y = 2 ^ x * (2 * y + 1)"
+  by (import ind_type NUMPAIR)
+
+lemma NUMPAIR_INJ_LEMMA: "ALL x xa xb xc. NUMPAIR x xa = NUMPAIR xb xc --> x = xb"
+  by (import ind_type NUMPAIR_INJ_LEMMA)
+
+lemma NUMPAIR_INJ: "ALL x1 y1 x2 y2. (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
+  by (import ind_type NUMPAIR_INJ)
+
+consts
+  NUMSND :: "nat => nat" 
+  NUMFST :: "nat => nat" 
+
+specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL x y. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y"
+  by (import ind_type NUMPAIR_DEST)
+
+constdefs
+  NUMSUM :: "bool => nat => nat" 
+  "NUMSUM == %b x. if b then Suc (2 * x) else 2 * x"
+
+lemma NUMSUM: "ALL b x. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)"
+  by (import ind_type NUMSUM)
+
+lemma NUMSUM_INJ: "ALL b1 x1 b2 x2. (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
+  by (import ind_type NUMSUM_INJ)
+
+consts
+  NUMRIGHT :: "nat => nat" 
+  NUMLEFT :: "nat => bool" 
+
+specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL x y. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y"
+  by (import ind_type NUMSUM_DEST)
+
+constdefs
+  INJN :: "nat => nat => 'a => bool" 
+  "INJN == %m n a. n = m"
+
+lemma INJN: "ALL m. INJN m = (%n a. n = m)"
+  by (import ind_type INJN)
+
+lemma INJN_INJ: "ALL n1 n2. (INJN n1 = INJN n2) = (n1 = n2)"
+  by (import ind_type INJN_INJ)
+
+constdefs
+  INJA :: "'a => nat => 'a => bool" 
+  "INJA == %a n b. b = a"
+
+lemma INJA: "ALL a. INJA a = (%n b. b = a)"
+  by (import ind_type INJA)
+
+lemma INJA_INJ: "ALL a1 a2. (INJA a1 = INJA a2) = (a1 = a2)"
+  by (import ind_type INJA_INJ)
+
+constdefs
+  INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" 
+  "INJF == %f n. f (NUMFST n) (NUMSND n)"
+
+lemma INJF: "ALL f. INJF f = (%n. f (NUMFST n) (NUMSND n))"
+  by (import ind_type INJF)
+
+lemma INJF_INJ: "ALL f1 f2. (INJF f1 = INJF f2) = (f1 = f2)"
+  by (import ind_type INJF_INJ)
+
+constdefs
+  INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" 
+  "INJP ==
+%f1 f2 n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a"
+
+lemma INJP: "ALL f1 f2.
+   INJP f1 f2 =
+   (%n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)"
+  by (import ind_type INJP)
+
+lemma INJP_INJ: "ALL f1 f1' f2 f2'. (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
+  by (import ind_type INJP_INJ)
+
+constdefs
+  ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" 
+  "ZCONSTR == %c i r. INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
+
+lemma ZCONSTR: "ALL c i r. ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
+  by (import ind_type ZCONSTR)
+
+constdefs
+  ZBOT :: "nat => 'a => bool" 
+  "ZBOT == INJP (INJN 0) (SOME z. True)"
+
+lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z. True)"
+  by (import ind_type ZBOT)
+
+lemma ZCONSTR_ZBOT: "ALL x xa xb. ZCONSTR x xa xb ~= ZBOT"
+  by (import ind_type ZCONSTR_ZBOT)
+
+constdefs
+  ZRECSPACE :: "(nat => 'a => bool) => bool" 
+  "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"
+
+lemma ZRECSPACE: "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)"
+  by (import ind_type ZRECSPACE)
+
+lemma ZRECSPACE_rules: "(op &::bool => bool => bool)
+ ((ZRECSPACE::(nat => 'a => bool) => bool) (ZBOT::nat => 'a => bool))
+ ((All::(nat => bool) => bool)
+   (%c::nat.
+       (All::('a => bool) => bool)
+        (%i::'a.
+            (All::((nat => nat => 'a => bool) => bool) => bool)
+             (%r::nat => nat => 'a => bool.
+                 (op -->::bool => bool => bool)
+                  ((All::(nat => bool) => bool)
+                    (%n::nat.
+                        (ZRECSPACE::(nat => 'a => bool) => bool) (r n)))
+                  ((ZRECSPACE::(nat => 'a => bool) => bool)
+                    ((ZCONSTR::nat
+                               => 'a => (nat => nat => 'a => bool)
+  => nat => 'a => bool)
+                      c i r))))))"
+  by (import ind_type ZRECSPACE_rules)
+
+lemma ZRECSPACE_ind: "ALL x.
+   x ZBOT & (ALL c i r. (ALL n. x (r n)) --> x (ZCONSTR c i r)) -->
+   (ALL a0. ZRECSPACE a0 --> x a0)"
+  by (import ind_type ZRECSPACE_ind)
+
+lemma ZRECSPACE_cases: "ALL a0.
+   ZRECSPACE a0 =
+   (a0 = ZBOT | (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE (r n))))"
+  by (import ind_type ZRECSPACE_cases)
+
+typedef (open) ('a) recspace = "(Collect::((nat => 'a => bool) => bool) => (nat => 'a => bool) set)
+ (ZRECSPACE::(nat => 'a => bool) => bool)" 
+  by (rule typedef_helper,import ind_type recspace_TY_DEF)
+
+lemmas recspace_TY_DEF = typedef_hol2hol4 [OF type_definition_recspace]
+
+consts
+  mk_rec :: "(nat => 'a => bool) => 'a recspace" 
+  dest_rec :: "'a recspace => nat => 'a => bool" 
+
+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))"
+  by (import ind_type recspace_repfns)
+
+constdefs
+  BOTTOM :: "'a recspace" 
+  "BOTTOM == mk_rec ZBOT"
+
+lemma BOTTOM: "BOTTOM = mk_rec ZBOT"
+  by (import ind_type BOTTOM)
+
+constdefs
+  CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" 
+  "CONSTR == %c i r. mk_rec (ZCONSTR c i (%n. dest_rec (r n)))"
+
+lemma CONSTR: "ALL c i r. CONSTR c i r = mk_rec (ZCONSTR c i (%n. dest_rec (r n)))"
+  by (import ind_type CONSTR)
+
+lemma MK_REC_INJ: "ALL x y. 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 y. (dest_rec x = dest_rec y) = (x = y)"
+  by (import ind_type DEST_REC_INJ)
+
+lemma CONSTR_BOT: "ALL c i r. CONSTR c i r ~= BOTTOM"
+  by (import ind_type CONSTR_BOT)
+
+lemma CONSTR_INJ: "ALL c1 i1 r1 c2 i2 r2.
+   (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.
+   P BOTTOM & (ALL c i r. (ALL n. P (r n)) --> P (CONSTR c i r)) --> All P"
+  by (import ind_type CONSTR_IND)
+
+lemma CONSTR_REC: "ALL Fn. EX f. ALL c i r. f (CONSTR c i r) = Fn c i r (%n. f (r n))"
+  by (import ind_type CONSTR_REC)
+
+consts
+  FCONS :: "'a => (nat => 'a) => nat => 'a" 
+
+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)"
+  by (import ind_type FCONS)
+
+constdefs
+  FNIL :: "nat => 'a" 
+  "FNIL == %n. SOME x. True"
+
+lemma FNIL: "ALL n. FNIL n = (SOME x. True)"
+  by (import ind_type FNIL)
+
+constdefs
+  ISO :: "('a => 'b) => ('b => 'a) => bool" 
+  "ISO == %f g. (ALL x. f (g x) = x) & (ALL y. g (f y) = y)"
+
+lemma ISO: "ALL f g. ISO f g = ((ALL x. f (g x) = x) & (ALL y. g (f y) = y))"
+  by (import ind_type ISO)
+
+lemma ISO_REFL: "ISO (%x. x) (%x. x)"
+  by (import ind_type ISO_REFL)
+
+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)))"
+  by (import ind_type ISO_FUN)
+
+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))"
+  by (import ind_type ISO_USAGE)
+
+;end_setup
+
+;setup_theory divides
+
+lemma ONE_DIVIDES_ALL: "All (op dvd (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 NOT_LT_DIV: "ALL (a::nat) b::nat. (0::nat) < b & b < a --> ~ a dvd b"
+  by (import divides NOT_LT_DIV)
+
+lemma DIVIDES_FACT: "ALL b. 0 < b --> b dvd FACT b"
+  by (import divides DIVIDES_FACT)
+
+lemma DIVIDES_MULT_LEFT: "ALL (x::nat) xa::nat. (x * xa dvd xa) = (xa = (0::nat) | x = (1::nat))"
+  by (import divides DIVIDES_MULT_LEFT)
+
+;end_setup
+
+;setup_theory prime
+
+consts
+  prime :: "nat => bool" 
+
+defs
+  prime_primdef: "prime.prime == %a. a ~= 1 & (ALL b. b dvd a --> b = a | b = 1)"
+
+lemma prime_def: "ALL a. prime.prime a = (a ~= 1 & (ALL b. b dvd a --> b = a | b = 1))"
+  by (import prime prime_def)
+
+lemma NOT_PRIME_0: "~ prime.prime 0"
+  by (import prime NOT_PRIME_0)
+
+lemma NOT_PRIME_1: "~ prime.prime 1"
+  by (import prime NOT_PRIME_1)
+
+;end_setup
+
+;setup_theory list
+
+consts
+  EL :: "nat => 'a list => 'a" 
+
+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))"
+  by (import list EL)
+
+lemma NULL: "(op &::bool => bool => bool) ((null::'a list => bool) ([]::'a list))
+ ((All::('a => bool) => bool)
+   (%x::'a.
+       (All::('a list => bool) => bool)
+        (%xa::'a list.
+            (Not::bool => bool)
+             ((null::'a list => bool)
+               ((op #::'a => 'a list => 'a list) x xa)))))"
+  by (import list NULL)
+
+lemma list_case_compute: "ALL l. list_case b f l = (if null l then b else f (hd l) (tl l))"
+  by (import list list_case_compute)
+
+lemma LIST_NOT_EQ: "ALL l1 l2. l1 ~= l2 --> (ALL x xa. x # l1 ~= xa # l2)"
+  by (import list LIST_NOT_EQ)
+
+lemma NOT_EQ_LIST: "ALL h1 h2. h1 ~= h2 --> (ALL x xa. h1 # x ~= h2 # xa)"
+  by (import list NOT_EQ_LIST)
+
+lemma EQ_LIST: "ALL h1 h2. h1 = h2 --> (ALL l1 l2. l1 = l2 --> h1 # l1 = h2 # l2)"
+  by (import list EQ_LIST)
+
+lemma CONS: "ALL l. ~ null l --> hd l # tl l = l"
+  by (import list CONS)
+
+lemma MAP_EQ_NIL: "ALL l f. (map f l = []) = (l = []) & ([] = map f l) = (l = [])"
+  by (import list MAP_EQ_NIL)
+
+lemma EVERY_EL: "ALL l P. list_all P l = (ALL n<length l. P (EL n l))"
+  by (import list EVERY_EL)
+
+lemma EVERY_CONJ: "ALL l. list_all (%x. P x & Q x) l = (list_all P l & list_all Q l)"
+  by (import list EVERY_CONJ)
+
+lemma EVERY_MEM: "ALL P l. list_all P l = (ALL e. e mem l --> P e)"
+  by (import list EVERY_MEM)
+
+lemma EXISTS_MEM: "ALL P l. list_exists P l = (EX e. e mem l & P e)"
+  by (import list EXISTS_MEM)
+
+lemma MEM_APPEND: "ALL e l1 l2. e mem l1 @ l2 = (e mem l1 | e mem l2)"
+  by (import list MEM_APPEND)
+
+lemma EXISTS_APPEND: "ALL P l1 l2. list_exists P (l1 @ l2) = (list_exists P l1 | list_exists P l2)"
+  by (import list EXISTS_APPEND)
+
+lemma NOT_EVERY: "ALL P l. (~ list_all P l) = list_exists (Not o P) l"
+  by (import list NOT_EVERY)
+
+lemma NOT_EXISTS: "ALL P l. (~ list_exists P l) = list_all (Not o P) l"
+  by (import list NOT_EXISTS)
+
+lemma MEM_MAP: "ALL l f x. x mem map f l = (EX y. x = f y & y mem l)"
+  by (import list MEM_MAP)
+
+lemma LENGTH_CONS: "ALL l n. (length l = Suc n) = (EX h l'. length l' = n & l = h # l')"
+  by (import list LENGTH_CONS)
+
+lemma LENGTH_EQ_CONS: "ALL P n.
+   (ALL l. length l = Suc n --> P l) =
+   (ALL l. length l = n --> (ALL x. P (x # l)))"
+  by (import list LENGTH_EQ_CONS)
+
+lemma LENGTH_EQ_NIL: "ALL P. (ALL l. length l = 0 --> P l) = P []"
+  by (import list LENGTH_EQ_NIL)
+
+lemma CONS_ACYCLIC: "ALL l x. l ~= x # l & x # l ~= l"
+  by (import list CONS_ACYCLIC)
+
+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 = []))"
+  by (import list APPEND_eq_NIL)
+
+lemma APPEND_11: "(ALL (l1::'a list) (l2::'a list) l3::'a list.
+    (l1 @ l2 = l1 @ l3) = (l2 = l3)) &
+(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. EL n l = (if n = 0 then hd l else EL (PRE n) (tl l))"
+  by (import list EL_compute)
+
+lemma WF_LIST_PRED: "WF (%L1 L2. EX h. L2 = h # L1)"
+  by (import list WF_LIST_PRED)
+
+lemma list_size_cong: "ALL M N f f'.
+   M = N & (ALL x. 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 l' b b' f f'.
+   l = l' & b = b' & (ALL x a. 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 l' b b' f f'.
+   l = l' & b = b' & (ALL x a. 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 l2 f f'.
+   l1 = l2 & (ALL x. x mem l2 --> f x = f' x) --> map f l1 = map f' l2"
+  by (import list MAP_CONG)
+
+lemma EXISTS_CONG: "ALL l1 l2 P P'.
+   l1 = l2 & (ALL x. x mem l2 --> P x = P' x) -->
+   list_exists P l1 = list_exists P' l2"
+  by (import list EXISTS_CONG)
+
+lemma EVERY_CONG: "ALL l1 l2 P P'.
+   l1 = l2 & (ALL x. 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 Q. (ALL x. P x --> Q x) --> (ALL l. list_all P l --> list_all Q l)"
+  by (import list EVERY_MONOTONIC)
+
+lemma LENGTH_ZIP: "ALL l1 l2.
+   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.
+   length (fst (unzip pl)) = length pl & length (snd (unzip pl)) = length pl"
+  by (import list LENGTH_UNZIP)
+
+lemma ZIP_UNZIP: "ALL l. ZIP (unzip l) = l"
+  by (import list ZIP_UNZIP)
+
+lemma UNZIP_ZIP: "ALL l1 l2. length l1 = length l2 --> unzip (zip l1 l2) = (l1, l2)"
+  by (import list UNZIP_ZIP)
+
+lemma ZIP_MAP: "ALL l1 l2 f1 f2.
+   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)"
+  by (import list ZIP_MAP)
+
+lemma MEM_ZIP: "ALL l1 l2 p.
+   length l1 = length l2 -->
+   p mem zip l1 l2 = (EX n<length l1. p = (EL n l1, EL n l2))"
+  by (import list MEM_ZIP)
+
+lemma EL_ZIP: "ALL l1 l2 n.
+   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 l1 l2.
+   length l1 = length l2 -->
+   (ALL f. map2 f l1 l2 = map (split f) (zip l1 l2))"
+  by (import list MAP2_ZIP)
+
+lemma MEM_EL: "ALL l x. x mem l = (EX n<length l. x = EL n l)"
+  by (import list MEM_EL)
+
+lemma LAST_CONS: "(ALL x::'a. last [x] = x) &
+(ALL (x::'a) (xa::'a) xb::'a list. last (x # xa # xb) = last (xa # xb))"
+  by (import list LAST_CONS)
+
+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)
+
+;end_setup
+
+;setup_theory pred_set
+
+lemma EXTENSION: "ALL s t. (s = t) = (ALL x. IN x s = IN x t)"
+  by (import pred_set EXTENSION)
+
+lemma NOT_EQUAL_SETS: "ALL x xa. (x ~= xa) = (EX xb. 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)
+
+consts
+  GSPEC :: "('b => 'a * bool) => 'a => bool" 
+
+specification (GSPEC) GSPECIFICATION: "ALL (f::'b => 'a * bool) v::'a. IN v (GSPEC f) = (EX x::'b. (v, True) = f x)"
+  by (import pred_set GSPECIFICATION)
+
+lemma SET_MINIMUM: "ALL (s::'a => bool) M::'a => nat.
+   (EX x::'a. IN x s) =
+   (EX x::'a. IN x s & (ALL y::'a. IN y s --> M x <= M y))"
+  by (import pred_set SET_MINIMUM)
+
+constdefs
+  EMPTY :: "'a => bool" 
+  "EMPTY == %x. False"
+
+lemma EMPTY_DEF: "EMPTY = (%x. False)"
+  by (import pred_set EMPTY_DEF)
+
+lemma NOT_IN_EMPTY: "ALL x. ~ IN x EMPTY"
+  by (import pred_set NOT_IN_EMPTY)
+
+lemma MEMBER_NOT_EMPTY: "ALL x. (EX xa. IN xa x) = (x ~= EMPTY)"
+  by (import pred_set MEMBER_NOT_EMPTY)
+
+constdefs
+  UNIV :: "'a => bool" 
+  "pred_set.UNIV == %x. True"
+
+lemma UNIV_DEF: "pred_set.UNIV = (%x. True)"
+  by (import pred_set UNIV_DEF)
+
+lemma IN_UNIV: "ALL x. IN x pred_set.UNIV"
+  by (import pred_set IN_UNIV)
+
+lemma UNIV_NOT_EMPTY: "pred_set.UNIV ~= EMPTY"
+  by (import pred_set UNIV_NOT_EMPTY)
+
+lemma EMPTY_NOT_UNIV: "EMPTY ~= pred_set.UNIV"
+  by (import pred_set EMPTY_NOT_UNIV)
+
+lemma EQ_UNIV: "(ALL x. IN x s) = (s = pred_set.UNIV)"
+  by (import pred_set EQ_UNIV)
+
+constdefs
+  SUBSET :: "('a => bool) => ('a => bool) => bool" 
+  "SUBSET == %s t. ALL x. IN x s --> IN x t"
+
+lemma SUBSET_DEF: "ALL s t. SUBSET s t = (ALL x. IN x s --> IN x t)"
+  by (import pred_set SUBSET_DEF)
+
+lemma SUBSET_TRANS: "ALL x xa xb. SUBSET x xa & SUBSET xa xb --> SUBSET x xb"
+  by (import pred_set SUBSET_TRANS)
+
+lemma SUBSET_REFL: "ALL x. SUBSET x x"
+  by (import pred_set SUBSET_REFL)
+
+lemma SUBSET_ANTISYM: "ALL x xa. 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. SUBSET x EMPTY = (x = EMPTY)"
+  by (import pred_set SUBSET_EMPTY)
+
+lemma SUBSET_UNIV: "ALL x. SUBSET x pred_set.UNIV"
+  by (import pred_set SUBSET_UNIV)
+
+lemma UNIV_SUBSET: "ALL x. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)"
+  by (import pred_set UNIV_SUBSET)
+
+constdefs
+  PSUBSET :: "('a => bool) => ('a => bool) => bool" 
+  "PSUBSET == %s t. SUBSET s t & s ~= t"
+
+lemma PSUBSET_DEF: "ALL s t. PSUBSET s t = (SUBSET s t & s ~= t)"
+  by (import pred_set PSUBSET_DEF)
+
+lemma PSUBSET_TRANS: "ALL x xa xb. PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb"
+  by (import pred_set PSUBSET_TRANS)
+
+lemma PSUBSET_IRREFL: "ALL x. ~ PSUBSET x x"
+  by (import pred_set PSUBSET_IRREFL)
+
+lemma NOT_PSUBSET_EMPTY: "ALL x. ~ PSUBSET x EMPTY"
+  by (import pred_set NOT_PSUBSET_EMPTY)
+
+lemma NOT_UNIV_PSUBSET: "ALL x. ~ PSUBSET pred_set.UNIV x"
+  by (import pred_set NOT_UNIV_PSUBSET)
+
+lemma PSUBSET_UNIV: "ALL x. PSUBSET x pred_set.UNIV = (EX xa. ~ IN xa x)"
+  by (import pred_set PSUBSET_UNIV)
+
+constdefs
+  UNION :: "('a => bool) => ('a => bool) => 'a => bool" 
+  "pred_set.UNION == %s t. GSPEC (%x. (x, IN x s | IN x t))"
+
+lemma UNION_DEF: "ALL s t. pred_set.UNION s t = GSPEC (%x. (x, IN x s | IN x t))"
+  by (import pred_set UNION_DEF)
+
+lemma IN_UNION: "ALL x xa xb. IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)"
+  by (import pred_set IN_UNION)
+
+lemma UNION_ASSOC: "ALL x xa xb.
+   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. pred_set.UNION x x = x"
+  by (import pred_set UNION_IDEMPOT)
+
+lemma UNION_COMM: "ALL x xa. pred_set.UNION x xa = pred_set.UNION xa x"
+  by (import pred_set UNION_COMM)
+
+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))"
+  by (import pred_set SUBSET_UNION)
+
+lemma UNION_SUBSET: "ALL s t u. 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 xa. SUBSET x xa = (pred_set.UNION x xa = xa)"
+  by (import pred_set SUBSET_UNION_ABSORPTION)
+
+lemma UNION_EMPTY: "(ALL x::'a => bool. pred_set.UNION EMPTY x = x) &
+(ALL x::'a => bool. pred_set.UNION x EMPTY = x)"
+  by (import pred_set UNION_EMPTY)
+
+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)"
+  by (import pred_set UNION_UNIV)
+
+lemma EMPTY_UNION: "ALL x xa. (pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)"
+  by (import pred_set EMPTY_UNION)
+
+constdefs
+  INTER :: "('a => bool) => ('a => bool) => 'a => bool" 
+  "pred_set.INTER == %s t. GSPEC (%x. (x, IN x s & IN x t))"
+
+lemma INTER_DEF: "ALL s t. pred_set.INTER s t = GSPEC (%x. (x, IN x s & IN x t))"
+  by (import pred_set INTER_DEF)
+
+lemma IN_INTER: "ALL x xa xb. IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)"
+  by (import pred_set IN_INTER)
+
+lemma INTER_ASSOC: "ALL x xa xb.
+   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. pred_set.INTER x x = x"
+  by (import pred_set INTER_IDEMPOT)
+
+lemma INTER_COMM: "ALL x xa. pred_set.INTER x xa = pred_set.INTER xa x"
+  by (import pred_set INTER_COMM)
+
+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)"
+  by (import pred_set INTER_SUBSET)
+
+lemma SUBSET_INTER: "ALL s t u. 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 xa. SUBSET x xa = (pred_set.INTER x xa = x)"
+  by (import pred_set SUBSET_INTER_ABSORPTION)
+
+lemma INTER_EMPTY: "(ALL x::'a => bool. pred_set.INTER EMPTY x = EMPTY) &
+(ALL x::'a => bool. pred_set.INTER x EMPTY = EMPTY)"
+  by (import pred_set INTER_EMPTY)
+
+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)"
+  by (import pred_set INTER_UNIV)
+
+lemma UNION_OVER_INTER: "ALL x xa xb.
+   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 xa xb.
+   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)
+
+constdefs
+  DISJOINT :: "('a => bool) => ('a => bool) => bool" 
+  "DISJOINT == %s t. pred_set.INTER s t = EMPTY"
+
+lemma DISJOINT_DEF: "ALL s t. DISJOINT s t = (pred_set.INTER s t = EMPTY)"
+  by (import pred_set DISJOINT_DEF)
+
+lemma IN_DISJOINT: "ALL x xa. DISJOINT x xa = (~ (EX xb. IN xb x & IN xb xa))"
+  by (import pred_set IN_DISJOINT)
+
+lemma DISJOINT_SYM: "ALL x xa. DISJOINT x xa = DISJOINT xa x"
+  by (import pred_set DISJOINT_SYM)
+
+lemma DISJOINT_EMPTY: "ALL x. DISJOINT EMPTY x & DISJOINT x EMPTY"
+  by (import pred_set DISJOINT_EMPTY)
+
+lemma DISJOINT_EMPTY_REFL: "ALL x. (x = EMPTY) = DISJOINT x x"
+  by (import pred_set DISJOINT_EMPTY_REFL)
+
+lemma DISJOINT_UNION: "ALL x xa xb.
+   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 t u.
+   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)
+
+constdefs
+  DIFF :: "('a => bool) => ('a => bool) => 'a => bool" 
+  "DIFF == %s t. GSPEC (%x. (x, IN x s & ~ IN x t))"
+
+lemma DIFF_DEF: "ALL s t. DIFF s t = GSPEC (%x. (x, IN x s & ~ IN x t))"
+  by (import pred_set DIFF_DEF)
+
+lemma IN_DIFF: "ALL s t x. IN x (DIFF s t) = (IN x s & ~ IN x t)"
+  by (import pred_set IN_DIFF)
+
+lemma DIFF_EMPTY: "ALL s. DIFF s EMPTY = s"
+  by (import pred_set DIFF_EMPTY)
+
+lemma EMPTY_DIFF: "ALL s. DIFF EMPTY s = EMPTY"
+  by (import pred_set EMPTY_DIFF)
+
+lemma DIFF_UNIV: "ALL s. DIFF s pred_set.UNIV = EMPTY"
+  by (import pred_set DIFF_UNIV)
+
+lemma DIFF_DIFF: "ALL x xa. DIFF (DIFF x xa) xa = DIFF x xa"
+  by (import pred_set DIFF_DIFF)
+
+lemma DIFF_EQ_EMPTY: "ALL x. DIFF x x = EMPTY"
+  by (import pred_set DIFF_EQ_EMPTY)
+
+constdefs
+  INSERT :: "'a => ('a => bool) => 'a => bool" 
+  "INSERT == %x s. GSPEC (%y. (y, y = x | IN y s))"
+
+lemma INSERT_DEF: "ALL x s. INSERT x s = GSPEC (%y. (y, y = x | IN y s))"
+  by (import pred_set INSERT_DEF)
+
+lemma IN_INSERT: "ALL x xa xb. IN x (INSERT xa xb) = (x = xa | IN x xb)"
+  by (import pred_set IN_INSERT)
+
+lemma COMPONENT: "ALL x xa. IN x (INSERT x xa)"
+  by (import pred_set COMPONENT)
+
+lemma SET_CASES: "ALL x. x = EMPTY | (EX xa xb. x = INSERT xa xb & ~ IN xa xb)"
+  by (import pred_set SET_CASES)
+
+lemma DECOMPOSITION: "ALL s x. IN x s = (EX t. s = INSERT x t & ~ IN x t)"
+  by (import pred_set DECOMPOSITION)
+
+lemma ABSORPTION: "ALL x xa. IN x xa = (INSERT x xa = xa)"
+  by (import pred_set ABSORPTION)
+
+lemma INSERT_INSERT: "ALL x xa. INSERT x (INSERT x xa) = INSERT x xa"
+  by (import pred_set INSERT_INSERT)
+
+lemma INSERT_COMM: "ALL x xa xb. INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)"
+  by (import pred_set INSERT_COMM)
+
+lemma INSERT_UNIV: "ALL x. INSERT x pred_set.UNIV = pred_set.UNIV"
+  by (import pred_set INSERT_UNIV)
+
+lemma NOT_INSERT_EMPTY: "ALL x xa. INSERT x xa ~= EMPTY"
+  by (import pred_set NOT_INSERT_EMPTY)
+
+lemma NOT_EMPTY_INSERT: "ALL x xa. EMPTY ~= INSERT x xa"
+  by (import pred_set NOT_EMPTY_INSERT)
+
+lemma INSERT_UNION: "ALL x s t.
+   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 s t. 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 s t.
+   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 xa xb. DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)"
+  by (import pred_set DISJOINT_INSERT)
+
+lemma INSERT_SUBSET: "ALL x xa xb. SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)"
+  by (import pred_set INSERT_SUBSET)
+
+lemma SUBSET_INSERT: "ALL x xa. ~ IN x xa --> (ALL xb. SUBSET xa (INSERT x xb) = SUBSET xa xb)"
+  by (import pred_set SUBSET_INSERT)
+
+lemma INSERT_DIFF: "ALL s t x.
+   DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))"
+  by (import pred_set INSERT_DIFF)
+
+constdefs
+  DELETE :: "('a => bool) => 'a => 'a => bool" 
+  "DELETE == %s x. DIFF s (INSERT x EMPTY)"
+
+lemma DELETE_DEF: "ALL s x. DELETE s x = DIFF s (INSERT x EMPTY)"
+  by (import pred_set DELETE_DEF)
+
+lemma IN_DELETE: "ALL x xa xb. IN xa (DELETE x xb) = (IN xa x & xa ~= xb)"
+  by (import pred_set IN_DELETE)
+
+lemma DELETE_NON_ELEMENT: "ALL x xa. (~ IN x xa) = (DELETE xa x = xa)"
+  by (import pred_set DELETE_NON_ELEMENT)
+
+lemma IN_DELETE_EQ: "ALL s x x'. (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. DELETE EMPTY x = EMPTY"
+  by (import pred_set EMPTY_DELETE)
+
+lemma DELETE_DELETE: "ALL x xa. DELETE (DELETE xa x) x = DELETE xa x"
+  by (import pred_set DELETE_DELETE)
+
+lemma DELETE_COMM: "ALL x xa xb. DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x"
+  by (import pred_set DELETE_COMM)
+
+lemma DELETE_SUBSET: "ALL x xa. SUBSET (DELETE xa x) xa"
+  by (import pred_set DELETE_SUBSET)
+
+lemma SUBSET_DELETE: "ALL x xa xb. SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)"
+  by (import pred_set SUBSET_DELETE)
+
+lemma SUBSET_INSERT_DELETE: "ALL x s t. SUBSET s (INSERT x t) = SUBSET (DELETE s x) t"
+  by (import pred_set SUBSET_INSERT_DELETE)
+
+lemma DIFF_INSERT: "ALL x xa xb. DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa"
+  by (import pred_set DIFF_INSERT)
+
+lemma PSUBSET_INSERT_SUBSET: "ALL x xa. PSUBSET x xa = (EX xb. ~ IN xb x & SUBSET (INSERT xb x) xa)"
+  by (import pred_set PSUBSET_INSERT_SUBSET)
+
+lemma PSUBSET_MEMBER: "ALL s t. PSUBSET s t = (SUBSET s t & (EX y. IN y t & ~ IN y s))"
+  by (import pred_set PSUBSET_MEMBER)
+
+lemma DELETE_INSERT: "ALL x xa xb.
+   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 xa. IN x xa --> INSERT x (DELETE xa x) = xa"
+  by (import pred_set INSERT_DELETE)
+
+lemma DELETE_INTER: "ALL x xa xb.
+   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 xa xb. DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x"
+  by (import pred_set DISJOINT_DELETE_SYM)
+
+consts
+  CHOICE :: "('a => bool) => 'a" 
+
+specification (CHOICE) CHOICE_DEF: "ALL x. x ~= EMPTY --> IN (CHOICE x) x"
+  by (import pred_set CHOICE_DEF)
+
+constdefs
+  REST :: "('a => bool) => 'a => bool" 
+  "REST == %s. DELETE s (CHOICE s)"
+
+lemma REST_DEF: "ALL s. REST s = DELETE s (CHOICE s)"
+  by (import pred_set REST_DEF)
+
+lemma CHOICE_NOT_IN_REST: "ALL x. ~ IN (CHOICE x) (REST x)"
+  by (import pred_set CHOICE_NOT_IN_REST)
+
+lemma CHOICE_INSERT_REST: "ALL s. s ~= EMPTY --> INSERT (CHOICE s) (REST s) = s"
+  by (import pred_set CHOICE_INSERT_REST)
+
+lemma REST_SUBSET: "ALL x. SUBSET (REST x) x"
+  by (import pred_set REST_SUBSET)
+
+lemma REST_PSUBSET: "ALL x. x ~= EMPTY --> PSUBSET (REST x) x"
+  by (import pred_set REST_PSUBSET)
+
+constdefs
+  SING :: "('a => bool) => bool" 
+  "SING == %s. EX x. s = INSERT x EMPTY"
+
+lemma SING_DEF: "ALL s. SING s = (EX x. s = INSERT x EMPTY)"
+  by (import pred_set SING_DEF)
+
+lemma SING: "ALL x. SING (INSERT x EMPTY)"
+  by (import pred_set SING)
+
+lemma IN_SING: "ALL x xa. IN x (INSERT xa EMPTY) = (x = xa)"
+  by (import pred_set IN_SING)
+
+lemma NOT_SING_EMPTY: "ALL x. INSERT x EMPTY ~= EMPTY"
+  by (import pred_set NOT_SING_EMPTY)
+
+lemma NOT_EMPTY_SING: "ALL x. EMPTY ~= INSERT x EMPTY"
+  by (import pred_set NOT_EMPTY_SING)
+
+lemma EQUAL_SING: "ALL x xa. (INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)"
+  by (import pred_set EQUAL_SING)
+
+lemma DISJOINT_SING_EMPTY: "ALL x. DISJOINT (INSERT x EMPTY) EMPTY"
+  by (import pred_set DISJOINT_SING_EMPTY)
+
+lemma INSERT_SING_UNION: "ALL x xa. INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x"
+  by (import pred_set INSERT_SING_UNION)
+
+lemma SING_DELETE: "ALL x. DELETE (INSERT x EMPTY) x = EMPTY"
+  by (import pred_set SING_DELETE)
+
+lemma DELETE_EQ_SING: "ALL x xa. IN xa x --> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)"
+  by (import pred_set DELETE_EQ_SING)
+
+lemma CHOICE_SING: "ALL x. CHOICE (INSERT x EMPTY) = x"
+  by (import pred_set CHOICE_SING)
+
+lemma REST_SING: "ALL x. REST (INSERT x EMPTY) = EMPTY"
+  by (import pred_set REST_SING)
+
+lemma SING_IFF_EMPTY_REST: "ALL x. SING x = (x ~= EMPTY & REST x = EMPTY)"
+  by (import pred_set SING_IFF_EMPTY_REST)
+
+constdefs
+  IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" 
+  "IMAGE == %f s. GSPEC (%x. (f x, IN x s))"
+
+lemma IMAGE_DEF: "ALL f s. IMAGE f s = GSPEC (%x. (f x, IN x s))"
+  by (import pred_set IMAGE_DEF)
+
+lemma IN_IMAGE: "ALL (x::'b) (xa::'a => bool) xb::'a => 'b.
+   IN x (IMAGE xb xa) = (EX xc::'a. x = xb xc & IN xc xa)"
+  by (import pred_set IN_IMAGE)
+
+lemma IMAGE_IN: "ALL x xa. IN x xa --> (ALL xb. IN (xb x) (IMAGE xb xa))"
+  by (import pred_set IMAGE_IN)
+
+lemma IMAGE_EMPTY: "ALL x. IMAGE x EMPTY = EMPTY"
+  by (import pred_set IMAGE_EMPTY)
+
+lemma IMAGE_ID: "ALL x. IMAGE (%x. x) x = x"
+  by (import pred_set IMAGE_ID)
+
+lemma IMAGE_COMPOSE: "ALL (x::'b => 'c) (xa::'a => 'b) xb::'a => bool.
+   IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)"
+  by (import pred_set IMAGE_COMPOSE)
+
+lemma IMAGE_INSERT: "ALL x xa xb. IMAGE x (INSERT xa xb) = INSERT (x xa) (IMAGE x xb)"
+  by (import pred_set IMAGE_INSERT)
+
+lemma IMAGE_EQ_EMPTY: "ALL s x. (IMAGE x s = EMPTY) = (s = EMPTY)"
+  by (import pred_set IMAGE_EQ_EMPTY)
+
+lemma IMAGE_DELETE: "ALL f x s. ~ IN x s --> IMAGE f (DELETE s x) = IMAGE f s"
+  by (import pred_set IMAGE_DELETE)
+
+lemma IMAGE_UNION: "ALL x xa xb.
+   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 xa. SUBSET x xa --> (ALL xb. SUBSET (IMAGE xb x) (IMAGE xb xa))"
+  by (import pred_set IMAGE_SUBSET)
+
+lemma IMAGE_INTER: "ALL f s t.
+   SUBSET (IMAGE f (pred_set.INTER s t))
+    (pred_set.INTER (IMAGE f s) (IMAGE f t))"
+  by (import pred_set IMAGE_INTER)
+
+constdefs
+  INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+  "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)"
+
+lemma INJ_DEF: "ALL f s t.
+   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))"
+  by (import pred_set INJ_DEF)
+
+lemma INJ_ID: "ALL x. INJ (%x. x) x x"
+  by (import pred_set INJ_ID)
+
+lemma INJ_COMPOSE: "ALL x xa xb xc xd. 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. All (INJ x EMPTY) & (ALL xa. INJ x xa EMPTY = (xa = EMPTY))"
+  by (import pred_set INJ_EMPTY)
+
+constdefs
+  SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+  "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))"
+
+lemma SURJ_DEF: "ALL f s t.
+   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)))"
+  by (import pred_set SURJ_DEF)
+
+lemma SURJ_ID: "ALL x. SURJ (%x. x) x x"
+  by (import pred_set SURJ_ID)
+
+lemma SURJ_COMPOSE: "ALL x xa xb xc xd. 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.
+   (ALL xa. SURJ x EMPTY xa = (xa = EMPTY)) &
+   (ALL xa. SURJ x xa EMPTY = (xa = EMPTY))"
+  by (import pred_set SURJ_EMPTY)
+
+lemma IMAGE_SURJ: "ALL x xa xb. SURJ x xa xb = (IMAGE x xa = xb)"
+  by (import pred_set IMAGE_SURJ)
+
+constdefs
+  BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
+  "BIJ == %f s t. INJ f s t & SURJ f s t"
+
+lemma BIJ_DEF: "ALL f s t. BIJ f s t = (INJ f s t & SURJ f s t)"
+  by (import pred_set BIJ_DEF)
+
+lemma BIJ_ID: "ALL x. BIJ (%x. x) x x"
+  by (import pred_set BIJ_ID)
+
+lemma BIJ_EMPTY: "ALL x.
+   (ALL xa. BIJ x EMPTY xa = (xa = EMPTY)) &
+   (ALL xa. BIJ x xa EMPTY = (xa = EMPTY))"
+  by (import pred_set BIJ_EMPTY)
+
+lemma BIJ_COMPOSE: "ALL x xa xb xc xd. BIJ x xb xc & BIJ xa xc xd --> BIJ (xa o x) xb xd"
+  by (import pred_set BIJ_COMPOSE)
+
+consts
+  LINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
+
+specification (LINV) LINV_DEF: "ALL f s t. INJ f s t --> (ALL x. IN x s --> LINV f s (f x) = x)"
+  by (import pred_set LINV_DEF)
+
+consts
+  RINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
+
+specification (RINV) RINV_DEF: "ALL f s t. SURJ f s t --> (ALL x. IN x t --> f (RINV f s x) = x)"
+  by (import pred_set RINV_DEF)
+
+constdefs
+  FINITE :: "('a => bool) => bool" 
+  "FINITE ==
+%s. ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s"
+
+lemma FINITE_DEF: "ALL s.
+   FINITE s =
+   (ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s)"
+  by (import pred_set FINITE_DEF)
+
+lemma FINITE_EMPTY: "FINITE EMPTY"
+  by (import pred_set FINITE_EMPTY)
+
+lemma FINITE_INDUCT: "ALL P.
+   P EMPTY &
+   (ALL s. FINITE s & P s --> (ALL e. ~ IN e s --> P (INSERT e s))) -->
+   (ALL s. FINITE s --> P s)"
+  by (import pred_set FINITE_INDUCT)
+
+lemma FINITE_INSERT: "ALL x s. FINITE (INSERT x s) = FINITE s"
+  by (import pred_set FINITE_INSERT)
+
+lemma FINITE_DELETE: "ALL x s. FINITE (DELETE s x) = FINITE s"
+  by (import pred_set FINITE_DELETE)
+
+lemma FINITE_UNION: "ALL s t. FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)"
+  by (import pred_set FINITE_UNION)
+
+lemma INTER_FINITE: "ALL s. FINITE s --> (ALL t. FINITE (pred_set.INTER s t))"
+  by (import pred_set INTER_FINITE)
+
+lemma SUBSET_FINITE: "ALL s. FINITE s --> (ALL t. SUBSET t s --> FINITE t)"
+  by (import pred_set SUBSET_FINITE)
+
+lemma PSUBSET_FINITE: "ALL x. FINITE x --> (ALL xa. PSUBSET xa x --> FINITE xa)"
+  by (import pred_set PSUBSET_FINITE)
+
+lemma FINITE_DIFF: "ALL s. FINITE s --> (ALL t. FINITE (DIFF s t))"
+  by (import pred_set FINITE_DIFF)
+
+lemma FINITE_SING: "ALL x. FINITE (INSERT x EMPTY)"
+  by (import pred_set FINITE_SING)
+
+lemma SING_FINITE: "ALL x. SING x --> FINITE x"
+  by (import pred_set SING_FINITE)
+
+lemma IMAGE_FINITE: "ALL s. FINITE s --> (ALL f. FINITE (IMAGE f s))"
+  by (import pred_set IMAGE_FINITE)
+
+consts
+  CARD :: "('a => bool) => nat" 
+
+specification (CARD) CARD_DEF: "(op &::bool => bool => bool)
+ ((op =::nat => nat => bool)
+   ((CARD::('a => bool) => nat) (EMPTY::'a => bool)) (0::nat))
+ ((All::(('a => bool) => bool) => bool)
+   (%s::'a => bool.
+       (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s)
+        ((All::('a => bool) => bool)
+          (%x::'a.
+              (op =::nat => nat => bool)
+               ((CARD::('a => bool) => nat)
+                 ((INSERT::'a => ('a => bool) => 'a => bool) x s))
+               ((If::bool => nat => nat => nat)
+                 ((IN::'a => ('a => bool) => bool) x s)
+                 ((CARD::('a => bool) => nat) s)
+                 ((Suc::nat => nat) ((CARD::('a => bool) => nat) s)))))))"
+  by (import pred_set CARD_DEF)
+
+lemma CARD_EMPTY: "CARD EMPTY = 0"
+  by (import pred_set CARD_EMPTY)
+
+lemma CARD_INSERT: "ALL s.
+   FINITE s -->
+   (ALL x. 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. FINITE s --> (CARD s = 0) = (s = EMPTY)"
+  by (import pred_set CARD_EQ_0)
+
+lemma CARD_DELETE: "ALL s.
+   FINITE s -->
+   (ALL x. 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. FINITE s --> (ALL t. CARD (pred_set.INTER s t) <= CARD s)"
+  by (import pred_set CARD_INTER_LESS_EQ)
+
+lemma CARD_UNION: "ALL s.
+   FINITE s -->
+   (ALL t.
+       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. FINITE s --> (ALL t. SUBSET t s --> CARD t <= CARD s)"
+  by (import pred_set CARD_SUBSET)
+
+lemma CARD_PSUBSET: "ALL s. FINITE s --> (ALL t. PSUBSET t s --> CARD t < CARD s)"
+  by (import pred_set CARD_PSUBSET)
+
+lemma CARD_SING: "ALL x. CARD (INSERT x EMPTY) = 1"
+  by (import pred_set CARD_SING)
+
+lemma SING_IFF_CARD1: "ALL x. SING x = (CARD x = 1 & FINITE x)"
+  by (import pred_set SING_IFF_CARD1)
+
+lemma CARD_DIFF: "ALL t.
+   FINITE t -->
+   (ALL s.
+       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.
+   FINITE t -->
+   (ALL s. FINITE s --> CARD t < CARD s --> 0 < CARD (DIFF s t))"
+  by (import pred_set LESS_CARD_DIFF)
+
+lemma FINITE_COMPLETE_INDUCTION: "ALL P.
+   (ALL x. (ALL y. PSUBSET y x --> P y) --> FINITE x --> P x) -->
+   (ALL x. FINITE x --> P x)"
+  by (import pred_set FINITE_COMPLETE_INDUCTION)
+
+constdefs
+  INFINITE :: "('a => bool) => bool" 
+  "INFINITE == %s. ~ FINITE s"
+
+lemma INFINITE_DEF: "ALL s. INFINITE s = (~ FINITE s)"
+  by (import pred_set INFINITE_DEF)
+
+lemma NOT_IN_FINITE: "(op =::bool => bool => bool)
+ ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool))
+ ((All::(('a => bool) => bool) => bool)
+   (%s::'a => bool.
+       (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s)
+        ((Ex::('a => bool) => bool)
+          (%x::'a.
+              (Not::bool => bool) ((IN::'a => ('a => bool) => bool) x s)))))"
+  by (import pred_set NOT_IN_FINITE)
+
+lemma INFINITE_INHAB: "ALL x. INFINITE x --> (EX xa. IN xa x)"
+  by (import pred_set INFINITE_INHAB)
+
+lemma IMAGE_11_INFINITE: "ALL f.
+   (ALL x y. f x = f y --> x = y) -->
+   (ALL s. INFINITE s --> INFINITE (IMAGE f s))"
+  by (import pred_set IMAGE_11_INFINITE)
+
+lemma INFINITE_SUBSET: "ALL x. INFINITE x --> (ALL xa. SUBSET x xa --> INFINITE xa)"
+  by (import pred_set INFINITE_SUBSET)
+
+lemma IN_INFINITE_NOT_FINITE: "ALL x xa. INFINITE x & FINITE xa --> (EX xb. IN xb x & ~ IN xb xa)"
+  by (import pred_set IN_INFINITE_NOT_FINITE)
+
+lemma INFINITE_UNIV: "(op =::bool => bool => bool)
+ ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool))
+ ((Ex::(('a => 'a) => bool) => bool)
+   (%f::'a => 'a.
+       (op &::bool => bool => bool)
+        ((All::('a => bool) => bool)
+          (%x::'a.
+              (All::('a => bool) => bool)
+               (%y::'a.
+                   (op -->::bool => bool => bool)
+                    ((op =::'a => 'a => bool) (f x) (f y))
+                    ((op =::'a => 'a => bool) x y))))
+        ((Ex::('a => bool) => bool)
+          (%y::'a.
+              (All::('a => bool) => bool)
+               (%x::'a.
+                   (Not::bool => bool)
+                    ((op =::'a => 'a => bool) (f x) y))))))"
+  by (import pred_set INFINITE_UNIV)
+
+lemma FINITE_PSUBSET_INFINITE: "ALL x. INFINITE x = (ALL xa. FINITE xa --> SUBSET xa x --> PSUBSET xa x)"
+  by (import pred_set FINITE_PSUBSET_INFINITE)
+
+lemma FINITE_PSUBSET_UNIV: "(op =::bool => bool => bool)
+ ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool))
+ ((All::(('a => bool) => bool) => bool)
+   (%s::'a => bool.
+       (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s)
+        ((PSUBSET::('a => bool) => ('a => bool) => bool) s
+          (pred_set.UNIV::'a => bool))))"
+  by (import pred_set FINITE_PSUBSET_UNIV)
+
+lemma INFINITE_DIFF_FINITE: "ALL s t. INFINITE s & FINITE t --> DIFF s t ~= EMPTY"
+  by (import pred_set INFINITE_DIFF_FINITE)
+
+lemma FINITE_ISO_NUM: "ALL s.
+   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)))"
+  by (import pred_set FINITE_ISO_NUM)
+
+lemma FINITE_WEAK_ENUMERATE: "ALL x::'a => bool.
+   FINITE x =
+   (EX (f::nat => 'a) b::nat. ALL e::'a. IN e x = (EX n<b. e = f n))"
+  by (import pred_set FINITE_WEAK_ENUMERATE)
+
+constdefs
+  BIGUNION :: "(('a => bool) => bool) => 'a => bool" 
+  "BIGUNION == %P. GSPEC (%x. (x, EX p. IN p P & IN x p))"
+
+lemma BIGUNION: "ALL P. BIGUNION P = GSPEC (%x. (x, EX p. IN p P & IN x p))"
+  by (import pred_set BIGUNION)
+
+lemma IN_BIGUNION: "ALL x xa. IN x (BIGUNION xa) = (EX s. IN x s & IN s xa)"
+  by (import pred_set IN_BIGUNION)
+
+lemma BIGUNION_EMPTY: "BIGUNION EMPTY = EMPTY"
+  by (import pred_set BIGUNION_EMPTY)
+
+lemma BIGUNION_SING: "ALL x. BIGUNION (INSERT x EMPTY) = x"
+  by (import pred_set BIGUNION_SING)
+
+lemma BIGUNION_UNION: "ALL x xa.
+   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 => bool) => bool) t::'a => bool.
+    DISJOINT (BIGUNION s) t =
+    (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 => bool. IN xb x --> DISJOINT xa xb))"
+  by (import pred_set DISJOINT_BIGUNION)
+
+lemma BIGUNION_INSERT: "ALL x xa. BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)"
+  by (import pred_set BIGUNION_INSERT)
+
+lemma BIGUNION_SUBSET: "ALL X P. SUBSET (BIGUNION P) X = (ALL Y. IN Y P --> SUBSET Y X)"
+  by (import pred_set BIGUNION_SUBSET)
+
+lemma FINITE_BIGUNION: "ALL x. FINITE x & (ALL s. IN s x --> FINITE s) --> FINITE (BIGUNION x)"
+  by (import pred_set FINITE_BIGUNION)
+
+constdefs
+  BIGINTER :: "(('a => bool) => bool) => 'a => bool" 
+  "BIGINTER == %B. GSPEC (%x. (x, ALL P. IN P B --> IN x P))"
+
+lemma BIGINTER: "ALL B. BIGINTER B = GSPEC (%x. (x, ALL P. IN P B --> IN x P))"
+  by (import pred_set BIGINTER)
+
+lemma IN_BIGINTER: "IN x (BIGINTER B) = (ALL P. IN P B --> IN x P)"
+  by (import pred_set IN_BIGINTER)
+
+lemma BIGINTER_INSERT: "ALL P B. BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)"
+  by (import pred_set BIGINTER_INSERT)
+
+lemma BIGINTER_EMPTY: "BIGINTER EMPTY = pred_set.UNIV"
+  by (import pred_set BIGINTER_EMPTY)
+
+lemma BIGINTER_INTER: "ALL x xa. BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa"
+  by (import pred_set BIGINTER_INTER)
+
+lemma BIGINTER_SING: "ALL x. BIGINTER (INSERT x EMPTY) = x"
+  by (import pred_set BIGINTER_SING)
+
+lemma SUBSET_BIGINTER: "ALL X P. SUBSET X (BIGINTER P) = (ALL x. IN x P --> SUBSET X x)"
+  by (import pred_set SUBSET_BIGINTER)
+
+lemma DISJOINT_BIGINTER: "ALL x xa xb.
+   IN xa xb & DISJOINT xa x -->
+   DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x"
+  by (import pred_set DISJOINT_BIGINTER)
+
+constdefs
+  CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" 
+  "CROSS == %P Q. GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))"
+
+lemma CROSS_DEF: "ALL P Q. CROSS P Q = GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))"
+  by (import pred_set CROSS_DEF)
+
+lemma IN_CROSS: "ALL x xa xb. IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)"
+  by (import pred_set IN_CROSS)
+
+lemma CROSS_EMPTY: "ALL x. CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY"
+  by (import pred_set CROSS_EMPTY)
+
+lemma CROSS_INSERT_LEFT: "ALL x xa xb.
+   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 xa xb.
+   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 xa. FINITE x & FINITE xa --> FINITE (CROSS x xa)"
+  by (import pred_set FINITE_CROSS)
+
+lemma CROSS_SINGS: "ALL x xa. CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY"
+  by (import pred_set CROSS_SINGS)
+
+lemma CARD_SING_CROSS: "ALL x s. FINITE s --> CARD (CROSS (INSERT x EMPTY) s) = CARD s"
+  by (import pred_set CARD_SING_CROSS)
+
+lemma CARD_CROSS: "ALL x xa. FINITE x & FINITE xa --> CARD (CROSS x xa) = CARD x * CARD xa"
+  by (import pred_set CARD_CROSS)
+
+lemma CROSS_SUBSET: "ALL x xa xb xc.
+   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 Q. FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)"
+  by (import pred_set FINITE_CROSS_EQ)
+
+constdefs
+  COMPL :: "('a => bool) => 'a => bool" 
+  "COMPL == DIFF pred_set.UNIV"
+
+lemma COMPL_DEF: "ALL P. COMPL P = DIFF pred_set.UNIV P"
+  by (import pred_set COMPL_DEF)
+
+lemma IN_COMPL: "ALL x xa. IN x (COMPL xa) = (~ IN x xa)"
+  by (import pred_set IN_COMPL)
+
+lemma COMPL_COMPL: "ALL x. COMPL (COMPL x) = x"
+  by (import pred_set COMPL_COMPL)
+
+lemma COMPL_CLAUSES: "ALL x.
+   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 xa.
+   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 xa. pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))"
+  by (import pred_set INTER_UNION_COMPL)
+
+lemma COMPL_EMPTY: "COMPL EMPTY = pred_set.UNIV"
+  by (import pred_set COMPL_EMPTY)
+
+consts
+  count :: "nat => nat => bool" 
+
+defs
+  count_primdef: "count == %n. GSPEC (%m. (m, m < n))"
+
+lemma count_def: "ALL n. count n = GSPEC (%m. (m, m < n))"
+  by (import pred_set count_def)
+
+lemma IN_COUNT: "ALL m n. IN m (count n) = (m < n)"
+  by (import pred_set IN_COUNT)
+
+lemma COUNT_ZERO: "count 0 = EMPTY"
+  by (import pred_set COUNT_ZERO)
+
+lemma COUNT_SUC: "ALL n. count (Suc n) = INSERT n (count n)"
+  by (import pred_set COUNT_SUC)
+
+lemma FINITE_COUNT: "ALL n. FINITE (count n)"
+  by (import pred_set FINITE_COUNT)
+
+lemma CARD_COUNT: "ALL n. CARD (count n) = n"
+  by (import pred_set CARD_COUNT)
+
+constdefs
+  ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" 
+  "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)"
+
+lemma ITSET_tupled_primitive_def: "ALL f.
+   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)"
+  by (import pred_set ITSET_tupled_primitive_def)
+
+constdefs
+  ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" 
+  "ITSET == %f x x1. ITSET_tupled f (x, x1)"
+
+lemma ITSET_curried_def: "ALL f x x1. ITSET f x x1 = ITSET_tupled f (x, x1)"
+  by (import pred_set ITSET_curried_def)
+
+lemma ITSET_IND: "ALL P.
+   (ALL s b.
+       (FINITE s & s ~= EMPTY --> P (REST s) (f (CHOICE s) b)) -->
+       P s b) -->
+   (ALL v. All (P v))"
+  by (import pred_set ITSET_IND)
+
+lemma ITSET_THM: "ALL s f b.
+   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 xa. ITSET x EMPTY xa = xa"
+  by (import pred_set ITSET_EMPTY)
+
+;end_setup
+
+;setup_theory operator
+
+constdefs
+  ASSOC :: "('a => 'a => 'a) => bool" 
+  "ASSOC == %f. ALL x y z. f x (f y z) = f (f x y) z"
+
+lemma ASSOC_DEF: "ALL f. ASSOC f = (ALL x y z. f x (f y z) = f (f x y) z)"
+  by (import operator ASSOC_DEF)
+
+constdefs
+  COMM :: "('a => 'a => 'b) => bool" 
+  "COMM == %f. ALL x y. f x y = f y x"
+
+lemma COMM_DEF: "ALL f. COMM f = (ALL x y. f x y = f y x)"
+  by (import operator COMM_DEF)
+
+constdefs
+  FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" 
+  "FCOMM == %f g. ALL x y z. g x (f y z) = f (g x y) z"
+
+lemma FCOMM_DEF: "ALL f g. FCOMM f g = (ALL x y z. g x (f y z) = f (g x y) z)"
+  by (import operator FCOMM_DEF)
+
+constdefs
+  RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" 
+  "RIGHT_ID == %f e. ALL x. f x e = x"
+
+lemma RIGHT_ID_DEF: "ALL f e. RIGHT_ID f e = (ALL x. f x e = x)"
+  by (import operator RIGHT_ID_DEF)
+
+constdefs
+  LEFT_ID :: "('a => 'b => 'b) => 'a => bool" 
+  "LEFT_ID == %f e. ALL x. f e x = x"
+
+lemma LEFT_ID_DEF: "ALL f e. LEFT_ID f e = (ALL x. f e x = x)"
+  by (import operator LEFT_ID_DEF)
+
+constdefs
+  MONOID :: "('a => 'a => 'a) => 'a => bool" 
+  "MONOID == %f e. ASSOC f & RIGHT_ID f e & LEFT_ID f e"
+
+lemma MONOID_DEF: "ALL f e. MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)"
+  by (import operator MONOID_DEF)
+
+lemma ASSOC_CONJ: "ASSOC op &"
+  by (import operator ASSOC_CONJ)
+
+lemma ASSOC_DISJ: "ASSOC op |"
+  by (import operator ASSOC_DISJ)
+
+lemma FCOMM_ASSOC: "ALL x. FCOMM x x = ASSOC x"
+  by (import operator FCOMM_ASSOC)
+
+lemma MONOID_CONJ_T: "MONOID op & True"
+  by (import operator MONOID_CONJ_T)
+
+lemma MONOID_DISJ_F: "MONOID op | False"
+  by (import operator MONOID_DISJ_F)
+
+;end_setup
+
+;setup_theory rich_list
+
+consts
+  SNOC :: "'a => 'a list => 'a list" 
+
+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)"
+  by (import rich_list SNOC)
+
+consts
+  SCANL :: "('b => 'a => 'b) => 'b => 'a list => 'b list" 
+
+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)"
+  by (import rich_list SCANL)
+
+consts
+  SCANR :: "('a => 'b => 'b) => 'b => 'a list => 'b 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 l. x mem l = list_exists (op = x) l"
+  by (import rich_list IS_EL_DEF)
+
+constdefs
+  AND_EL :: "bool list => bool" 
+  "AND_EL == list_all I"
+
+lemma AND_EL_DEF: "AND_EL = list_all I"
+  by (import rich_list AND_EL_DEF)
+
+constdefs
+  OR_EL :: "bool list => bool" 
+  "OR_EL == list_exists I"
+
+lemma OR_EL_DEF: "OR_EL = list_exists I"
+  by (import rich_list OR_EL_DEF)
+
+consts
+  FIRSTN :: "nat => 'a list => 'a list" 
+
+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)"
+  by (import rich_list FIRSTN)
+
+consts
+  BUTFIRSTN :: "nat => 'a list => 'a list" 
+
+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)"
+  by (import rich_list BUTFIRSTN)
+
+consts
+  SEG :: "nat => nat => 'a list => 'a 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 l. last (SNOC x l) = x"
+  by (import rich_list LAST)
+
+lemma BUTLAST: "ALL x l. butlast (SNOC x l) = l"
+  by (import rich_list BUTLAST)
+
+consts
+  LASTN :: "nat => 'a list => 'a 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)
+
+consts
+  BUTLASTN :: "nat => 'a list => 'a 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 list. EL (0::nat) x = hd x) &
+(ALL (x::nat) xa::'a list. EL (Suc x) xa = EL x (tl xa))"
+  by (import rich_list EL)
+
+consts
+  ELL :: "nat => 'a list => 'a" 
+
+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))"
+  by (import rich_list ELL)
+
+consts
+  IS_PREFIX :: "'a list => 'a list => bool" 
+
+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 l. SNOC x l = l @ [x]"
+  by (import rich_list SNOC_APPEND)
+
+lemma REVERSE: "rev [] = [] & (ALL (x::'a) xa::'a list. rev (x # xa) = SNOC x (rev xa))"
+  by (import rich_list REVERSE)
+
+lemma REVERSE_SNOC: "ALL x l. rev (SNOC x l) = x # rev l"
+  by (import rich_list REVERSE_SNOC)
+
+lemma SNOC_Axiom: "ALL (e::'b) f::'a => 'a list => 'b => 'b.
+   EX x::'a list => 'b.
+      x [] = e & (ALL (xa::'a) l::'a list. x (SNOC xa l) = f xa l (x l))"
+  by (import rich_list SNOC_Axiom)
+
+consts
+  IS_SUFFIX :: "'a list => 'a list => bool" 
+
+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)
+
+consts
+  IS_SUBLIST :: "'a list => 'a list => bool" 
+
+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)
+
+consts
+  SPLITP :: "('a => bool) => 'a list => 'a list * 'a 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)
+
+constdefs
+  PREFIX :: "('a => bool) => 'a list => 'a list" 
+  "PREFIX == %P l. fst (SPLITP (Not o P) l)"
+
+lemma PREFIX_DEF: "ALL P l. PREFIX P l = fst (SPLITP (Not o P) l)"
+  by (import rich_list PREFIX_DEF)
+
+constdefs
+  SUFFIX :: "('a => bool) => 'a list => 'a list" 
+  "SUFFIX == %P. foldl (%l' x. if P x then SNOC x l' else []) []"
+
+lemma SUFFIX_DEF: "ALL P l. SUFFIX P l = foldl (%l' x. if P x then SNOC x l' else []) [] l"
+  by (import rich_list SUFFIX_DEF)
+
+constdefs
+  UNZIP_FST :: "('a * 'b) list => 'a list" 
+  "UNZIP_FST == %l. fst (unzip l)"
+
+lemma UNZIP_FST_DEF: "ALL l. UNZIP_FST l = fst (unzip l)"
+  by (import rich_list UNZIP_FST_DEF)
+
+constdefs
+  UNZIP_SND :: "('a * 'b) list => 'b list" 
+  "UNZIP_SND == %l. snd (unzip l)"
+
+lemma UNZIP_SND_DEF: "ALL l. UNZIP_SND l = snd (unzip l)"
+  by (import rich_list UNZIP_SND_DEF)
+
+consts
+  GENLIST :: "(nat => 'a) => nat => 'a list" 
+
+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))"
+  by (import rich_list GENLIST)
+
+consts
+  REPLICATE :: "nat => 'a => 'a list" 
+
+specification (REPLICATE) REPLICATE: "(ALL x::'a. REPLICATE (0::nat) x = []) &
+(ALL (n::nat) x::'a. REPLICATE (Suc n) x = x # REPLICATE n x)"
+  by (import rich_list REPLICATE)
+
+lemma LENGTH_MAP2: "ALL l1 l2.
+   length l1 = length l2 -->
+   (ALL f.
+       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. null l = (l = [])"
+  by (import rich_list NULL_EQ_NIL)
+
+lemma LENGTH_EQ: "ALL x y. x = y --> length x = length y"
+  by (import rich_list LENGTH_EQ)
+
+lemma LENGTH_NOT_NULL: "ALL l. (0 < length l) = (~ null l)"
+  by (import rich_list LENGTH_NOT_NULL)
+
+lemma SNOC_INDUCT: "ALL P. P [] & (ALL l. P l --> (ALL x. P (SNOC x l))) --> All P"
+  by (import rich_list SNOC_INDUCT)
+
+lemma SNOC_CASES: "ALL x'. x' = [] | (EX x l. x' = SNOC x l)"
+  by (import rich_list SNOC_CASES)
+
+lemma LENGTH_SNOC: "ALL x l. length (SNOC x l) = Suc (length l)"
+  by (import rich_list LENGTH_SNOC)
+
+lemma NOT_NIL_SNOC: "ALL x xa. [] ~= SNOC x xa"
+  by (import rich_list NOT_NIL_SNOC)
+
+lemma NOT_SNOC_NIL: "ALL x xa. SNOC x xa ~= []"
+  by (import rich_list NOT_SNOC_NIL)
+
+lemma SNOC_11: "ALL x l x' l'. (SNOC x l = SNOC x' l') = (x = x' & l = l')"
+  by (import rich_list SNOC_11)
+
+lemma SNOC_EQ_LENGTH_EQ: "ALL x1 l1 x2 l2. SNOC x1 l1 = SNOC x2 l2 --> length l1 = length l2"
+  by (import rich_list SNOC_EQ_LENGTH_EQ)
+
+lemma SNOC_REVERSE_CONS: "ALL x xa. SNOC x xa = rev (x # rev xa)"
+  by (import rich_list SNOC_REVERSE_CONS)
+
+lemma MAP_SNOC: "ALL x xa xb. map x (SNOC xa xb) = SNOC (x xa) (map x xb)"
+  by (import rich_list MAP_SNOC)
+
+lemma FOLDR_SNOC: "ALL f e x l. foldr f (SNOC x l) e = foldr f l (f x e)"
+  by (import rich_list FOLDR_SNOC)
+
+lemma FOLDL_SNOC: "ALL (f::'b => 'a => 'b) (e::'b) (x::'a) l::'a list.
+   foldl f e (SNOC x l) = f (foldl f e l) x"
+  by (import rich_list FOLDL_SNOC)
+
+lemma FOLDR_FOLDL: "ALL f e. MONOID f e --> (ALL l. foldr f l e = foldl f e l)"
+  by (import rich_list FOLDR_FOLDL)
+
+lemma LENGTH_FOLDR: "ALL l. length l = foldr (%x. Suc) l 0"
+  by (import rich_list LENGTH_FOLDR)
+
+lemma LENGTH_FOLDL: "ALL l. length l = foldl (%l' x. Suc l') 0 l"
+  by (import rich_list LENGTH_FOLDL)
+
+lemma MAP_FOLDR: "ALL f l. map f l = foldr (%x. op # (f x)) l []"
+  by (import rich_list MAP_FOLDR)
+
+lemma MAP_FOLDL: "ALL f l. map f l = foldl (%l' x. SNOC (f x) l') [] l"
+  by (import rich_list MAP_FOLDL)
+
+lemma MAP_o: "ALL (f::'b => 'c) g::'a => 'b. map (f o g) = map f o map g"
+  by (import rich_list MAP_o)
+
+lemma FILTER_FOLDR: "ALL P l. filter P l = foldr (%x l'. if P x then x # l' else l') l []"
+  by (import rich_list FILTER_FOLDR)
+
+lemma FILTER_SNOC: "ALL P x l.
+   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 l. filter P l = foldl (%l' x. if P x then SNOC x l' else l') [] l"
+  by (import rich_list FILTER_FOLDL)
+
+lemma FILTER_COMM: "ALL f1 f2 l. filter f1 (filter f2 l) = filter f2 (filter f1 l)"
+  by (import rich_list FILTER_COMM)
+
+lemma FILTER_IDEM: "ALL f l. filter f (filter f l) = filter f l"
+  by (import rich_list FILTER_IDEM)
+
+lemma FILTER_MAP: "ALL (f1::'b => bool) (f2::'a => 'b) l::'a list.
+   filter f1 (map f2 l) = map f2 (filter (f1 o f2) l)"
+  by (import rich_list FILTER_MAP)
+
+lemma LENGTH_SEG: "ALL n k l. n + k <= length l --> length (SEG n k l) = n"
+  by (import rich_list LENGTH_SEG)
+
+lemma APPEND_NIL: "(ALL l::'a list. l @ [] = l) & (ALL x::'a list. [] @ x = x)"
+  by (import rich_list APPEND_NIL)
+
+lemma APPEND_SNOC: "ALL l1 x l2. l1 @ SNOC x l2 = SNOC x (l1 @ l2)"
+  by (import rich_list APPEND_SNOC)
+
+lemma APPEND_FOLDR: "ALL l1 l2. l1 @ l2 = foldr op # l1 l2"
+  by (import rich_list APPEND_FOLDR)
+
+lemma APPEND_FOLDL: "ALL l1 l2. l1 @ l2 = foldl (%l' x. SNOC x l') l1 l2"
+  by (import rich_list APPEND_FOLDL)
+
+lemma CONS_APPEND: "ALL x l. x # l = [x] @ l"
+  by (import rich_list CONS_APPEND)
+
+lemma ASSOC_APPEND: "ASSOC op @"
+  by (import rich_list ASSOC_APPEND)
+
+lemma MONOID_APPEND_NIL: "MONOID op @ []"
+  by (import rich_list MONOID_APPEND_NIL)
+
+lemma APPEND_LENGTH_EQ: "ALL l1 l1'.
+   length l1 = length l1' -->
+   (ALL l2 l2'.
+       length l2 = length l2' -->
+       (l1 @ l2 = l1' @ l2') = (l1 = l1' & l2 = l2'))"
+  by (import rich_list APPEND_LENGTH_EQ)
+
+lemma FLAT_SNOC: "ALL x l. concat (SNOC x l) = concat l @ x"
+  by (import rich_list FLAT_SNOC)
+
+lemma FLAT_FOLDR: "ALL l. concat l = foldr op @ l []"
+  by (import rich_list FLAT_FOLDR)
+
+lemma FLAT_FOLDL: "ALL l. concat l = foldl op @ [] l"
+  by (import rich_list FLAT_FOLDL)
+
+lemma LENGTH_FLAT: "ALL l. length (concat l) = sum (map size l)"
+  by (import rich_list LENGTH_FLAT)
+
+lemma REVERSE_FOLDR: "ALL l. rev l = foldr SNOC l []"
+  by (import rich_list REVERSE_FOLDR)
+
+lemma REVERSE_FOLDL: "ALL l. rev l = foldl (%l' x. x # l') [] l"
+  by (import rich_list REVERSE_FOLDL)
+
+lemma ALL_EL_SNOC: "ALL P x l. 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 => bool) (f::'a => 'b) l::'a 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 x l. list_exists P (SNOC x l) = (P x | list_exists P l)"
+  by (import rich_list SOME_EL_SNOC)
+
+lemma IS_EL_SNOC: "ALL y x l. y mem SNOC x l = (y = x | y mem l)"
+  by (import rich_list IS_EL_SNOC)
+
+lemma SUM_SNOC: "ALL x l. sum (SNOC x l) = sum l + x"
+  by (import rich_list SUM_SNOC)
+
+lemma SUM_FOLDL: "ALL l. sum l = foldl op + 0 l"
+  by (import rich_list SUM_FOLDL)
+
+lemma IS_PREFIX_APPEND: "ALL l1 l2. IS_PREFIX l1 l2 = (EX l. l1 = l2 @ l)"
+  by (import rich_list IS_PREFIX_APPEND)
+
+lemma IS_SUFFIX_APPEND: "ALL l1 l2. IS_SUFFIX l1 l2 = (EX l. l1 = l @ l2)"
+  by (import rich_list IS_SUFFIX_APPEND)
+
+lemma IS_SUBLIST_APPEND: "ALL l1 l2. IS_SUBLIST l1 l2 = (EX l l'. l1 = l @ l2 @ l')"
+  by (import rich_list IS_SUBLIST_APPEND)
+
+lemma IS_PREFIX_IS_SUBLIST: "ALL l1 l2. IS_PREFIX l1 l2 --> IS_SUBLIST l1 l2"
+  by (import rich_list IS_PREFIX_IS_SUBLIST)
+
+lemma IS_SUFFIX_IS_SUBLIST: "ALL l1 l2. IS_SUFFIX l1 l2 --> IS_SUBLIST l1 l2"
+  by (import rich_list IS_SUFFIX_IS_SUBLIST)
+
+lemma IS_PREFIX_REVERSE: "ALL l1 l2. IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2"
+  by (import rich_list IS_PREFIX_REVERSE)
+
+lemma IS_SUFFIX_REVERSE: "ALL l2 l1. IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2"
+  by (import rich_list IS_SUFFIX_REVERSE)
+
+lemma IS_SUBLIST_REVERSE: "ALL l1 l2. IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2"
+  by (import rich_list IS_SUBLIST_REVERSE)
+
+lemma PREFIX_FOLDR: "ALL P x. PREFIX P x = foldr (%x l'. if P x then x # l' else []) x []"
+  by (import rich_list PREFIX_FOLDR)
+
+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 l. IS_PREFIX l (PREFIX P l)"
+  by (import rich_list IS_PREFIX_PREFIX)
+
+lemma LENGTH_SCANL: "ALL (f::'b => 'a => 'b) (e::'b) l::'a list.
+   length (SCANL f e l) = Suc (length l)"
+  by (import rich_list LENGTH_SCANL)
+
+lemma LENGTH_SCANR: "ALL f e l. length (SCANR f e l) = Suc (length l)"
+  by (import rich_list LENGTH_SCANR)
+
+lemma COMM_MONOID_FOLDL: "ALL x.
+   COMM x -->
+   (ALL xa. MONOID x xa --> (ALL e l. foldl x e l = x e (foldl x xa l)))"
+  by (import rich_list COMM_MONOID_FOLDL)
+
+lemma COMM_MONOID_FOLDR: "ALL x.
+   COMM x -->
+   (ALL xa. MONOID x xa --> (ALL e l. foldr x l e = x e (foldr x l xa)))"
+  by (import rich_list COMM_MONOID_FOLDR)
+
+lemma FCOMM_FOLDR_APPEND: "ALL x xa.
+   FCOMM x xa -->
+   (ALL xb.
+       LEFT_ID x xb -->
+       (ALL l1 l2.
+           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 xa.
+   FCOMM x xa -->
+   (ALL xb.
+       RIGHT_ID xa xb -->
+       (ALL l1 l2.
+           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 xa xb. foldl x xa [xb] = x xa xb"
+  by (import rich_list FOLDL_SINGLE)
+
+lemma FOLDR_SINGLE: "ALL x xa xb. foldr x [xb] xa = x xb xa"
+  by (import rich_list FOLDR_SINGLE)
+
+lemma FOLDR_CONS_NIL: "ALL l. foldr op # l [] = l"
+  by (import rich_list FOLDR_CONS_NIL)
+
+lemma FOLDL_SNOC_NIL: "ALL l. foldl (%xs x. SNOC x xs) [] l = l"
+  by (import rich_list FOLDL_SNOC_NIL)
+
+lemma FOLDR_REVERSE: "ALL x xa xb. foldr x (rev xb) xa = foldl (%xa y. x y xa) xa xb"
+  by (import rich_list FOLDR_REVERSE)
+
+lemma FOLDL_REVERSE: "ALL x xa xb. foldl x xa (rev xb) = foldr (%xa y. x y xa) xb xa"
+  by (import rich_list FOLDL_REVERSE)
+
+lemma FOLDR_MAP: "ALL (f::'a => 'a => 'a) (e::'a) (g::'b => 'a) l::'b list.
+   foldr f (map g l) e = foldr (%x::'b. f (g x)) l e"
+  by (import rich_list FOLDR_MAP)
+
+lemma FOLDL_MAP: "ALL (f::'a => 'a => 'a) (e::'a) (g::'b => 'a) l::'b list.
+   foldl f e (map g l) = foldl (%(x::'a) y::'b. f x (g y)) e l"
+  by (import rich_list FOLDL_MAP)
+
+lemma ALL_EL_FOLDR: "ALL P l. list_all P l = foldr (%x. op & (P x)) l True"
+  by (import rich_list ALL_EL_FOLDR)
+
+lemma ALL_EL_FOLDL: "ALL P l. list_all P l = foldl (%l' x. l' & P x) True l"
+  by (import rich_list ALL_EL_FOLDL)
+
+lemma SOME_EL_FOLDR: "ALL P l. list_exists P l = foldr (%x. op | (P x)) l False"
+  by (import rich_list SOME_EL_FOLDR)
+
+lemma SOME_EL_FOLDL: "ALL P l. list_exists P l = foldl (%l' x. l' | P x) False l"
+  by (import rich_list SOME_EL_FOLDL)
+
+lemma ALL_EL_FOLDR_MAP: "ALL x xa. 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 xa. 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 xa. list_exists x xa = foldr op | (map x xa) False"
+  by (import rich_list SOME_EL_FOLDR_MAP)
+
+lemma SOME_EL_FOLDL_MAP: "ALL x xa. list_exists x xa = foldl op | False (map x xa)"
+  by (import rich_list SOME_EL_FOLDL_MAP)
+
+lemma FOLDR_FILTER: "ALL (f::'a => 'a => 'a) (e::'a) (P::'a => bool) l::'a list.
+   foldr f (filter P l) e =
+   foldr (%(x::'a) y::'a. if P x then f x y else y) l e"
+  by (import rich_list FOLDR_FILTER)
+
+lemma FOLDL_FILTER: "ALL (f::'a => 'a => 'a) (e::'a) (P::'a => bool) l::'a list.
+   foldl f e (filter P l) =
+   foldl (%(x::'a) y::'a. if P y then f x y else x) e l"
+  by (import rich_list FOLDL_FILTER)
+
+lemma ASSOC_FOLDR_FLAT: "ALL f.
+   ASSOC f -->
+   (ALL e.
+       LEFT_ID f e -->
+       (ALL l. 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.
+   ASSOC f -->
+   (ALL e.
+       RIGHT_ID f e -->
+       (ALL l. 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 => bool) (f::'a => 'b) l::'a list.
+   list_exists P (map f l) = list_exists (P o f) l"
+  by (import rich_list SOME_EL_MAP)
+
+lemma SOME_EL_DISJ: "ALL P Q l.
+   list_exists (%x. P x | Q x) l = (list_exists P l | list_exists Q l)"
+  by (import rich_list SOME_EL_DISJ)
+
+lemma IS_EL_FOLDR: "ALL x xa. x mem xa = foldr (%xa. op | (x = xa)) xa False"
+  by (import rich_list IS_EL_FOLDR)
+
+lemma IS_EL_FOLDL: "ALL x xa. x mem xa = foldl (%l' xa. l' | x = xa) False xa"
+  by (import rich_list IS_EL_FOLDL)
+
+lemma NULL_FOLDR: "ALL l. null l = foldr (%x l'. False) l True"
+  by (import rich_list NULL_FOLDR)
+
+lemma NULL_FOLDL: "ALL l. null l = foldl (%x l'. False) True l"
+  by (import rich_list NULL_FOLDL)
+
+lemma FILTER_REVERSE: "ALL P l. filter P (rev l) = rev (filter P l)"
+  by (import rich_list FILTER_REVERSE)
+
+lemma SEG_LENGTH_ID: "ALL l. SEG (length l) 0 l = l"
+  by (import rich_list SEG_LENGTH_ID)
+
+lemma SEG_SUC_CONS: "ALL m n l x. SEG m (Suc n) (x # l) = SEG m n l"
+  by (import rich_list SEG_SUC_CONS)
+
+lemma SEG_0_SNOC: "ALL m l x. 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 l. n <= length l --> BUTLASTN n l = SEG (length l - n) 0 l"
+  by (import rich_list BUTLASTN_SEG)
+
+lemma LASTN_CONS: "ALL n l. n <= length l --> (ALL x. LASTN n (x # l) = LASTN n l)"
+  by (import rich_list LASTN_CONS)
+
+lemma LENGTH_LASTN: "ALL n l. n <= length l --> length (LASTN n l) = n"
+  by (import rich_list LENGTH_LASTN)
+
+lemma LASTN_LENGTH_ID: "ALL l. LASTN (length l) l = l"
+  by (import rich_list LASTN_LENGTH_ID)
+
+lemma LASTN_LASTN: "ALL l n m. 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. FIRSTN (length l) l = l"
+  by (import rich_list FIRSTN_LENGTH_ID)
+
+lemma FIRSTN_SNOC: "ALL n l. n <= length l --> (ALL x. FIRSTN n (SNOC x l) = FIRSTN n l)"
+  by (import rich_list FIRSTN_SNOC)
+
+lemma BUTLASTN_LENGTH_NIL: "ALL l. BUTLASTN (length l) l = []"
+  by (import rich_list BUTLASTN_LENGTH_NIL)
+
+lemma BUTLASTN_SUC_BUTLAST: "ALL n l. n < length l --> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)"
+  by (import rich_list BUTLASTN_SUC_BUTLAST)
+
+lemma BUTLASTN_BUTLAST: "ALL n l. n < length l --> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)"
+  by (import rich_list BUTLASTN_BUTLAST)
+
+lemma LENGTH_BUTLASTN: "ALL n l. n <= length l --> length (BUTLASTN n l) = length l - n"
+  by (import rich_list LENGTH_BUTLASTN)
+
+lemma BUTLASTN_BUTLASTN: "ALL m n l.
+   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 l. n <= length l --> BUTLASTN n l @ LASTN n l = l"
+  by (import rich_list APPEND_BUTLASTN_LASTN)
+
+lemma APPEND_FIRSTN_LASTN: "ALL m n l. m + n = length l --> FIRSTN n l @ LASTN m l = l"
+  by (import rich_list APPEND_FIRSTN_LASTN)
+
+lemma BUTLASTN_APPEND2: "ALL n l1 l2. n <= length l2 --> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2"
+  by (import rich_list BUTLASTN_APPEND2)
+
+lemma BUTLASTN_LENGTH_APPEND: "ALL l2 l1. BUTLASTN (length l2) (l1 @ l2) = l1"
+  by (import rich_list BUTLASTN_LENGTH_APPEND)
+
+lemma LASTN_LENGTH_APPEND: "ALL l2 l1. LASTN (length l2) (l1 @ l2) = l2"
+  by (import rich_list LASTN_LENGTH_APPEND)
+
+lemma BUTLASTN_CONS: "ALL n l. n <= length l --> (ALL x. BUTLASTN n (x # l) = x # BUTLASTN n l)"
+  by (import rich_list BUTLASTN_CONS)
+
+lemma BUTLASTN_LENGTH_CONS: "ALL l x. BUTLASTN (length l) (x # l) = [x]"
+  by (import rich_list BUTLASTN_LENGTH_CONS)
+
+lemma LAST_LASTN_LAST: "ALL n l. n <= length l --> 0 < n --> last (LASTN n l) = last l"
+  by (import rich_list LAST_LASTN_LAST)
+
+lemma BUTLASTN_LASTN_NIL: "ALL n l. n <= length l --> BUTLASTN n (LASTN n l) = []"
+  by (import rich_list BUTLASTN_LASTN_NIL)
+
+lemma LASTN_BUTLASTN: "ALL n m l.
+   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 n l.
+   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. l ~= [] --> LASTN 1 l = [last l]"
+  by (import rich_list LASTN_1)
+
+lemma BUTLASTN_1: "ALL l. l ~= [] --> BUTLASTN 1 l = butlast l"
+  by (import rich_list BUTLASTN_1)
+
+lemma BUTLASTN_APPEND1: "ALL l2 n.
+   length l2 <= n -->
+   (ALL l1. BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1)"
+  by (import rich_list BUTLASTN_APPEND1)
+
+lemma LASTN_APPEND2: "ALL n l2. n <= length l2 --> (ALL l1. LASTN n (l1 @ l2) = LASTN n l2)"
+  by (import rich_list LASTN_APPEND2)
+
+lemma LASTN_APPEND1: "ALL l2 n.
+   length l2 <= n -->
+   (ALL l1. LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2)"
+  by (import rich_list LASTN_APPEND1)
+
+lemma LASTN_MAP: "ALL n l. n <= length l --> (ALL f. LASTN n (map f l) = map f (LASTN n l))"
+  by (import rich_list LASTN_MAP)
+
+lemma BUTLASTN_MAP: "ALL n l.
+   n <= length l --> (ALL f. BUTLASTN n (map f l) = map f (BUTLASTN n l))"
+  by (import rich_list BUTLASTN_MAP)
+
+lemma ALL_EL_LASTN: "ALL P l. list_all P l --> (ALL m<=length l. list_all P (LASTN m l))"
+  by (import rich_list ALL_EL_LASTN)
+
+lemma ALL_EL_BUTLASTN: "ALL P l. list_all P l --> (ALL m<=length l. list_all P (BUTLASTN m l))"
+  by (import rich_list ALL_EL_BUTLASTN)
+
+lemma LENGTH_FIRSTN: "ALL n l. n <= length l --> length (FIRSTN n l) = n"
+  by (import rich_list LENGTH_FIRSTN)
+
+lemma FIRSTN_FIRSTN: "ALL m l. m <= length l --> (ALL n<=m. FIRSTN n (FIRSTN m l) = FIRSTN n l)"
+  by (import rich_list FIRSTN_FIRSTN)
+
+lemma LENGTH_BUTFIRSTN: "ALL n l. n <= length l --> length (BUTFIRSTN n l) = length l - n"
+  by (import rich_list LENGTH_BUTFIRSTN)
+
+lemma BUTFIRSTN_LENGTH_NIL: "ALL l. BUTFIRSTN (length l) l = []"
+  by (import rich_list BUTFIRSTN_LENGTH_NIL)
+
+lemma BUTFIRSTN_APPEND1: "ALL n l1.
+   n <= length l1 --> (ALL l2. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2)"
+  by (import rich_list BUTFIRSTN_APPEND1)
+
+lemma BUTFIRSTN_APPEND2: "ALL l1 n.
+   length l1 <= n -->
+   (ALL l2. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2)"
+  by (import rich_list BUTFIRSTN_APPEND2)
+
+lemma BUTFIRSTN_BUTFIRSTN: "ALL n m l.
+   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 l. n <= length l --> FIRSTN n l @ BUTFIRSTN n l = l"
+  by (import rich_list APPEND_FIRSTN_BUTFIRSTN)
+
+lemma LASTN_SEG: "ALL n l. n <= length l --> LASTN n l = SEG n (length l - n) l"
+  by (import rich_list LASTN_SEG)
+
+lemma FIRSTN_SEG: "ALL n l. n <= length l --> FIRSTN n l = SEG n 0 l"
+  by (import rich_list FIRSTN_SEG)
+
+lemma BUTFIRSTN_SEG: "ALL n l. n <= length l --> BUTFIRSTN n l = SEG (length l - n) n l"
+  by (import rich_list BUTFIRSTN_SEG)
+
+lemma BUTFIRSTN_SNOC: "ALL n l.
+   n <= length l -->
+   (ALL x. BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l))"
+  by (import rich_list BUTFIRSTN_SNOC)
+
+lemma APPEND_BUTLASTN_BUTFIRSTN: "ALL m n l. m + n = length l --> BUTLASTN m l @ BUTFIRSTN n l = l"
+  by (import rich_list APPEND_BUTLASTN_BUTFIRSTN)
+
+lemma SEG_SEG: "ALL n1 m1 n2 m2 l.
+   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 m l1. n + m <= length l1 --> (ALL l2. SEG n m (l1 @ l2) = SEG n m l1)"
+  by (import rich_list SEG_APPEND1)
+
+lemma SEG_APPEND2: "ALL l1 m n l2.
+   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 m l. 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 l1 n l2.
+   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 xa. SEG 1 (length x) (SNOC xa x) = [xa]"
+  by (import rich_list SEG_LENGTH_SNOC)
+
+lemma SEG_SNOC: "ALL n m l. n + m <= length l --> (ALL x. SEG n m (SNOC x l) = SEG n m l)"
+  by (import rich_list SEG_SNOC)
+
+lemma ELL_SEG: "ALL n l. 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 l. SNOC x l = foldr op # l [x]"
+  by (import rich_list SNOC_FOLDR)
+
+lemma IS_EL_FOLDR_MAP: "ALL x xa. 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 xa. x mem xa = foldl op | False (map (op = x) xa)"
+  by (import rich_list IS_EL_FOLDL_MAP)
+
+lemma FILTER_FILTER: "ALL P Q l. filter P (filter Q l) = [x:l. P x & Q x]"
+  by (import rich_list FILTER_FILTER)
+
+lemma FCOMM_FOLDR_FLAT: "ALL g f.
+   FCOMM g f -->
+   (ALL e.
+       LEFT_ID g e -->
+       (ALL l. 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 g.
+   FCOMM f g -->
+   (ALL e.
+       RIGHT_ID g e -->
+       (ALL l. 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 => 'a => 'a.
+   (ALL (a::'a) (b::'a) c::'a. f a (f b c) = f b (f a c)) -->
+   (ALL (e::'a) (g::'b => 'a) l::'b 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 => 'a => 'a.
+   (ALL (a::'a) (b::'a) c::'a. f a (f b c) = f b (f a c)) -->
+   (ALL (e::'a) (P::'a => bool) l::'a 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. COMM f --> ASSOC f --> (ALL e l. foldr f (rev l) e = foldr f l e)"
+  by (import rich_list COMM_ASSOC_FOLDR_REVERSE)
+
+lemma COMM_ASSOC_FOLDL_REVERSE: "ALL f. COMM f --> ASSOC f --> (ALL e l. foldl f e (rev l) = foldl f e l)"
+  by (import rich_list COMM_ASSOC_FOLDL_REVERSE)
+
+lemma ELL_LAST: "ALL l. ~ null l --> ELL 0 l = last l"
+  by (import rich_list ELL_LAST)
+
+lemma ELL_0_SNOC: "ALL l x. ELL 0 (SNOC x l) = x"
+  by (import rich_list ELL_0_SNOC)
+
+lemma ELL_SNOC: "ALL n. 0 < n --> (ALL x l. ELL n (SNOC x l) = ELL (PRE n) l)"
+  by (import rich_list ELL_SNOC)
+
+lemma ELL_SUC_SNOC: "ALL n x xa. ELL (Suc n) (SNOC x xa) = ELL n xa"
+  by (import rich_list ELL_SUC_SNOC)
+
+lemma ELL_CONS: "ALL n l. n < length l --> (ALL x. ELL n (x # l) = ELL n l)"
+  by (import rich_list ELL_CONS)
+
+lemma ELL_LENGTH_CONS: "ALL l x. ELL (length l) (x # l) = x"
+  by (import rich_list ELL_LENGTH_CONS)
+
+lemma ELL_LENGTH_SNOC: "ALL l x. 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 l2. n < length l2 --> (ALL l1. ELL n (l1 @ l2) = ELL n l2)"
+  by (import rich_list ELL_APPEND2)
+
+lemma ELL_APPEND1: "ALL l2 n.
+   length l2 <= n --> (ALL l1. ELL n (l1 @ l2) = ELL (n - length l2) l1)"
+  by (import rich_list ELL_APPEND1)
+
+lemma ELL_PRE_LENGTH: "ALL l. l ~= [] --> ELL (PRE (length l)) l = hd l"
+  by (import rich_list ELL_PRE_LENGTH)
+
+lemma EL_LENGTH_SNOC: "ALL l x. EL (length l) (SNOC x l) = x"
+  by (import rich_list EL_LENGTH_SNOC)
+
+lemma EL_PRE_LENGTH: "ALL l. l ~= [] --> EL (PRE (length l)) l = last l"
+  by (import rich_list EL_PRE_LENGTH)
+
+lemma EL_SNOC: "ALL n l. n < length l --> (ALL x. EL n (SNOC x l) = EL n l)"
+  by (import rich_list EL_SNOC)
+
+lemma EL_ELL: "ALL n l. n < length l --> EL n l = ELL (PRE (length l - n)) l"
+  by (import rich_list EL_ELL)
+
+lemma EL_LENGTH_APPEND: "ALL l2 l1. ~ null l2 --> EL (length l1) (l1 @ l2) = hd l2"
+  by (import rich_list EL_LENGTH_APPEND)
+
+lemma ELL_EL: "ALL n l. n < length l --> ELL n l = EL (PRE (length l - n)) l"
+  by (import rich_list ELL_EL)
+
+lemma ELL_MAP: "ALL n l f. n < length l --> ELL n (map f l) = f (ELL n l)"
+  by (import rich_list ELL_MAP)
+
+lemma LENGTH_BUTLAST: "ALL l. l ~= [] --> length (butlast l) = PRE (length l)"
+  by (import rich_list LENGTH_BUTLAST)
+
+lemma BUTFIRSTN_LENGTH_APPEND: "ALL l1 l2. BUTFIRSTN (length l1) (l1 @ l2) = l2"
+  by (import rich_list BUTFIRSTN_LENGTH_APPEND)
+
+lemma FIRSTN_APPEND1: "ALL n l1. n <= length l1 --> (ALL l2. FIRSTN n (l1 @ l2) = FIRSTN n l1)"
+  by (import rich_list FIRSTN_APPEND1)
+
+lemma FIRSTN_APPEND2: "ALL l1 n.
+   length l1 <= n -->
+   (ALL l2. FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2)"
+  by (import rich_list FIRSTN_APPEND2)
+
+lemma FIRSTN_LENGTH_APPEND: "ALL l1 l2. FIRSTN (length l1) (l1 @ l2) = l1"
+  by (import rich_list FIRSTN_LENGTH_APPEND)
+
+lemma REVERSE_FLAT: "ALL l. rev (concat l) = concat (rev (map rev l))"
+  by (import rich_list REVERSE_FLAT)
+
+lemma MAP_FILTER: "ALL f P l.
+   (ALL x. 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. concat (rev l) = rev (concat (map rev l))"
+  by (import rich_list FLAT_REVERSE)
+
+lemma FLAT_FLAT: "ALL l. concat (concat l) = concat (map concat l)"
+  by (import rich_list FLAT_FLAT)
+
+lemma ALL_EL_REVERSE: "ALL P l. list_all P (rev l) = list_all P l"
+  by (import rich_list ALL_EL_REVERSE)
+
+lemma SOME_EL_REVERSE: "ALL P l. list_exists P (rev l) = list_exists P l"
+  by (import rich_list SOME_EL_REVERSE)
+
+lemma ALL_EL_SEG: "ALL P l.
+   list_all P l --> (ALL m k. m + k <= length l --> list_all P (SEG m k l))"
+  by (import rich_list ALL_EL_SEG)
+
+lemma ALL_EL_FIRSTN: "ALL P l. list_all P l --> (ALL m<=length l. list_all P (FIRSTN m l))"
+  by (import rich_list ALL_EL_FIRSTN)
+
+lemma ALL_EL_BUTFIRSTN: "ALL P l. list_all P l --> (ALL m<=length l. list_all P (BUTFIRSTN m l))"
+  by (import rich_list ALL_EL_BUTFIRSTN)
+
+lemma SOME_EL_SEG: "ALL m k l.
+   m + k <= length l -->
+   (ALL P. list_exists P (SEG m k l) --> list_exists P l)"
+  by (import rich_list SOME_EL_SEG)
+
+lemma SOME_EL_FIRSTN: "ALL m l.
+   m <= length l --> (ALL P. list_exists P (FIRSTN m l) --> list_exists P l)"
+  by (import rich_list SOME_EL_FIRSTN)
+
+lemma SOME_EL_BUTFIRSTN: "ALL m l.
+   m <= length l -->
+   (ALL P. list_exists P (BUTFIRSTN m l) --> list_exists P l)"
+  by (import rich_list SOME_EL_BUTFIRSTN)
+
+lemma SOME_EL_LASTN: "ALL m l.
+   m <= length l --> (ALL P. list_exists P (LASTN m l) --> list_exists P l)"
+  by (import rich_list SOME_EL_LASTN)
+
+lemma SOME_EL_BUTLASTN: "ALL m l.
+   m <= length l -->
+   (ALL P. list_exists P (BUTLASTN m l) --> list_exists P l)"
+  by (import rich_list SOME_EL_BUTLASTN)
+
+lemma IS_EL_REVERSE: "ALL x l. x mem rev l = x mem l"
+  by (import rich_list IS_EL_REVERSE)
+
+lemma IS_EL_FILTER: "ALL P x. P x --> (ALL l. x mem filter P l = x mem l)"
+  by (import rich_list IS_EL_FILTER)
+
+lemma IS_EL_SEG: "ALL n m l. n + m <= length l --> (ALL x. x mem SEG n m l --> x mem l)"
+  by (import rich_list IS_EL_SEG)
+
+lemma IS_EL_SOME_EL: "ALL x l. x mem l = list_exists (op = x) l"
+  by (import rich_list IS_EL_SOME_EL)
+
+lemma IS_EL_FIRSTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem FIRSTN x xa --> xb mem xa)"
+  by (import rich_list IS_EL_FIRSTN)
+
+lemma IS_EL_BUTFIRSTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem BUTFIRSTN x xa --> xb mem xa)"
+  by (import rich_list IS_EL_BUTFIRSTN)
+
+lemma IS_EL_BUTLASTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem BUTLASTN x xa --> xb mem xa)"
+  by (import rich_list IS_EL_BUTLASTN)
+
+lemma IS_EL_LASTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem LASTN x xa --> xb mem xa)"
+  by (import rich_list IS_EL_LASTN)
+
+lemma ZIP_SNOC: "ALL l1 l2.
+   length l1 = length l2 -->
+   (ALL x1 x2. zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2))"
+  by (import rich_list ZIP_SNOC)
+
+lemma UNZIP_SNOC: "ALL x l.
+   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. length (UNZIP_FST x) = length x"
+  by (import rich_list LENGTH_UNZIP_FST)
+
+lemma LENGTH_UNZIP_SND: "ALL x. length (UNZIP_SND x) = length x"
+  by (import rich_list LENGTH_UNZIP_SND)
+
+lemma SUM_APPEND: "ALL l1 l2. sum (l1 @ l2) = sum l1 + sum l2"
+  by (import rich_list SUM_APPEND)
+
+lemma SUM_REVERSE: "ALL l. sum (rev l) = sum l"
+  by (import rich_list SUM_REVERSE)
+
+lemma SUM_FLAT: "ALL l. sum (concat l) = sum (map sum l)"
+  by (import rich_list SUM_FLAT)
+
+lemma EL_APPEND1: "ALL n l1 l2. n < length l1 --> EL n (l1 @ l2) = EL n l1"
+  by (import rich_list EL_APPEND1)
+
+lemma EL_APPEND2: "ALL l1 n.
+   length l1 <= n --> (ALL l2. EL n (l1 @ l2) = EL (n - length l1) l2)"
+  by (import rich_list EL_APPEND2)
+
+lemma EL_MAP: "ALL n l. n < length l --> (ALL f. EL n (map f l) = f (EL n l))"
+  by (import rich_list EL_MAP)
+
+lemma EL_CONS: "ALL n. 0 < n --> (ALL x l. EL n (x # l) = EL (PRE n) l)"
+  by (import rich_list EL_CONS)
+
+lemma EL_SEG: "ALL n l. n < length l --> EL n l = hd (SEG 1 n l)"
+  by (import rich_list EL_SEG)
+
+lemma EL_IS_EL: "ALL n l. n < length l --> EL n l mem l"
+  by (import rich_list EL_IS_EL)
+
+lemma TL_SNOC: "ALL x l. tl (SNOC x l) = (if null l then [] else SNOC x (tl l))"
+  by (import rich_list TL_SNOC)
+
+lemma EL_REVERSE: "ALL n l. 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 l. n < length l --> EL n (rev l) = ELL n l"
+  by (import rich_list EL_REVERSE_ELL)
+
+lemma ELL_LENGTH_APPEND: "ALL l1 l2. ~ null l1 --> ELL (length l2) (l1 @ l2) = last l1"
+  by (import rich_list ELL_LENGTH_APPEND)
+
+lemma ELL_IS_EL: "ALL n l. n < length l --> ELL n l mem l"
+  by (import rich_list ELL_IS_EL)
+
+lemma ELL_REVERSE: "ALL n l. 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 l. n < length l --> ELL n (rev l) = EL n l"
+  by (import rich_list ELL_REVERSE_EL)
+
+lemma FIRSTN_BUTLASTN: "ALL n l. n <= length l --> FIRSTN n l = BUTLASTN (length l - n) l"
+  by (import rich_list FIRSTN_BUTLASTN)
+
+lemma BUTLASTN_FIRSTN: "ALL n l. n <= length l --> BUTLASTN n l = FIRSTN (length l - n) l"
+  by (import rich_list BUTLASTN_FIRSTN)
+
+lemma LASTN_BUTFIRSTN: "ALL n l. n <= length l --> LASTN n l = BUTFIRSTN (length l - n) l"
+  by (import rich_list LASTN_BUTFIRSTN)
+
+lemma BUTFIRSTN_LASTN: "ALL n l. n <= length l --> BUTFIRSTN n l = LASTN (length l - n) l"
+  by (import rich_list BUTFIRSTN_LASTN)
+
+lemma SEG_LASTN_BUTLASTN: "ALL n m l.
+   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 l. n <= length l --> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)"
+  by (import rich_list BUTFIRSTN_REVERSE)
+
+lemma BUTLASTN_REVERSE: "ALL n l. n <= length l --> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)"
+  by (import rich_list BUTLASTN_REVERSE)
+
+lemma LASTN_REVERSE: "ALL n l. n <= length l --> LASTN n (rev l) = rev (FIRSTN n l)"
+  by (import rich_list LASTN_REVERSE)
+
+lemma FIRSTN_REVERSE: "ALL n l. n <= length l --> FIRSTN n (rev l) = rev (LASTN n l)"
+  by (import rich_list FIRSTN_REVERSE)
+
+lemma SEG_REVERSE: "ALL n m l.
+   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 n. length (GENLIST f n) = n"
+  by (import rich_list LENGTH_GENLIST)
+
+lemma LENGTH_REPLICATE: "ALL n x. length (REPLICATE n x) = n"
+  by (import rich_list LENGTH_REPLICATE)
+
+lemma IS_EL_REPLICATE: "ALL n. 0 < n --> (ALL x. x mem REPLICATE n x)"
+  by (import rich_list IS_EL_REPLICATE)
+
+lemma ALL_EL_REPLICATE: "ALL x n. list_all (op = x) (REPLICATE n x)"
+  by (import rich_list ALL_EL_REPLICATE)
+
+lemma AND_EL_FOLDL: "ALL l. AND_EL l = foldl op & True l"
+  by (import rich_list AND_EL_FOLDL)
+
+lemma AND_EL_FOLDR: "ALL l. AND_EL l = foldr op & l True"
+  by (import rich_list AND_EL_FOLDR)
+
+lemma OR_EL_FOLDL: "ALL l. OR_EL l = foldl op | False l"
+  by (import rich_list OR_EL_FOLDL)
+
+lemma OR_EL_FOLDR: "ALL l. OR_EL l = foldr op | l False"
+  by (import rich_list OR_EL_FOLDR)
+
+;end_setup
+
+;setup_theory state_transformer
+
+constdefs
+  UNIT :: "'b => 'a => 'b * 'a" 
+  "(op ==::('b => 'a => 'b * 'a) => ('b => 'a => 'b * 'a) => prop)
+ (UNIT::'b => 'a => 'b * 'a) (Pair::'b => 'a => 'b * 'a)"
+
+lemma UNIT_DEF: "ALL x::'b. UNIT x = Pair x"
+  by (import state_transformer UNIT_DEF)
+
+constdefs
+  BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" 
+  "BIND == %g f. split f o g"
+
+lemma BIND_DEF: "ALL g f. BIND g f = split f o g"
+  by (import state_transformer BIND_DEF)
+
+constdefs
+  MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" 
+  "MMAP == %(f::'c => 'b) m::'a => 'c * 'a. BIND m (UNIT o f)"
+
+lemma MMAP_DEF: "ALL (f::'c => 'b) m::'a => 'c * 'a. MMAP f m = BIND m (UNIT o f)"
+  by (import state_transformer MMAP_DEF)
+
+constdefs
+  JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" 
+  "JOIN == %z. BIND z I"
+
+lemma JOIN_DEF: "ALL z. JOIN z = BIND z I"
+  by (import state_transformer JOIN_DEF)
+
+lemma BIND_LEFT_UNIT: "ALL k x. BIND (UNIT x) k = k x"
+  by (import state_transformer BIND_LEFT_UNIT)
+
+lemma UNIT_UNCURRY: "ALL x. split UNIT x = x"
+  by (import state_transformer UNIT_UNCURRY)
+
+lemma BIND_RIGHT_UNIT: "ALL k. BIND k UNIT = k"
+  by (import state_transformer BIND_RIGHT_UNIT)
+
+lemma BIND_ASSOC: "ALL x xa xb. BIND x (%a. BIND (xa a) xb) = BIND (BIND x xa) xb"
+  by (import state_transformer BIND_ASSOC)
+
+lemma MMAP_ID: "MMAP I = I"
+  by (import state_transformer MMAP_ID)
+
+lemma MMAP_COMP: "ALL (f::'c => 'd) g::'b => 'c. MMAP (f o g) = MMAP f o MMAP g"
+  by (import state_transformer MMAP_COMP)
+
+lemma MMAP_UNIT: "ALL f::'b => 'c. MMAP f o UNIT = UNIT o f"
+  by (import state_transformer MMAP_UNIT)
+
+lemma MMAP_JOIN: "ALL f::'b => 'c. MMAP f o JOIN = JOIN o MMAP (MMAP f)"
+  by (import state_transformer MMAP_JOIN)
+
+lemma JOIN_UNIT: "JOIN o UNIT = I"
+  by (import state_transformer JOIN_UNIT)
+
+lemma JOIN_MMAP_UNIT: "JOIN o MMAP UNIT = I"
+  by (import state_transformer JOIN_MMAP_UNIT)
+
+lemma JOIN_MAP_JOIN: "JOIN o MMAP JOIN = JOIN o JOIN"
+  by (import state_transformer JOIN_MAP_JOIN)
+
+lemma JOIN_MAP: "ALL x xa. BIND x xa = JOIN (MMAP xa x)"
+  by (import state_transformer JOIN_MAP)
+
+lemma FST_o_UNIT: "ALL x. fst o UNIT x = K x"
+  by (import state_transformer FST_o_UNIT)
+
+lemma SND_o_UNIT: "ALL x. snd o UNIT x = I"
+  by (import state_transformer SND_o_UNIT)
+
+lemma FST_o_MMAP: "ALL x xa. fst o MMAP x xa = x o (fst o xa)"
+  by (import state_transformer FST_o_MMAP)
+
+;end_setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4Prob.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,3981 @@
+theory HOL4Prob = HOL4Real:
+
+;setup_theory prob_extra
+
+lemma BOOL_BOOL_CASES_THM: "ALL f. 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. 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 q. Suc q * (p div Suc q) <= p"
+  by (import prob_extra DIV_THEN_MULT)
+
+lemma DIV_TWO_UNIQUE: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(nat => bool) => bool)
+      (%q::nat.
+          (All::(nat => bool) => bool)
+           (%r::nat.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op =::nat => nat => bool) n
+                    ((op +::nat => nat => nat)
+                      ((op *::nat => nat => nat)
+                        ((number_of::bin => nat)
+                          ((op BIT::bin => bool => bin)
+                            ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                              (True::bool))
+                            (False::bool)))
+                        q)
+                      r))
+                  ((op |::bool => bool => bool)
+                    ((op =::nat => nat => bool) r (0::nat))
+                    ((op =::nat => nat => bool) r (1::nat))))
+                ((op &::bool => bool => bool)
+                  ((op =::nat => nat => bool) q
+                    ((op div::nat => nat => nat) n
+                      ((number_of::bin => nat)
+                        ((op BIT::bin => bool => bin)
+                          ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                            (True::bool))
+                          (False::bool)))))
+                  ((op =::nat => nat => bool) r
+                    ((op mod::nat => nat => nat) n
+                      ((number_of::bin => nat)
+                        ((op BIT::bin => bool => bin)
+                          ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                            (True::bool))
+                          (False::bool)))))))))"
+  by (import prob_extra DIV_TWO_UNIQUE)
+
+lemma DIVISION_TWO: "ALL n::nat.
+   n = (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::nat) * (n div (2::nat)) + n mod (2::nat)"
+  by (import prob_extra DIV_TWO)
+
+lemma MOD_TWO: "ALL n. n mod 2 = (if EVEN n then 0 else 1)"
+  by (import prob_extra MOD_TWO)
+
+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::(nat => bool) => bool)
+ (%m::nat.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((op <::nat => nat => bool)
+             ((op div::nat => nat => nat) m
+               ((number_of::bin => nat)
+                 ((op BIT::bin => bool => bin)
+                   ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                     (True::bool))
+                   (False::bool))))
+             ((op div::nat => nat => nat) n
+               ((number_of::bin => nat)
+                 ((op BIT::bin => bool => bin)
+                   ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                     (True::bool))
+                   (False::bool)))))
+           ((op <::nat => nat => bool) m n)))"
+  by (import prob_extra DIV_TWO_MONO)
+
+lemma DIV_TWO_MONO_EVEN: "(All::(nat => bool) => bool)
+ (%m::nat.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool) ((EVEN::nat => bool) n)
+           ((op =::bool => bool => bool)
+             ((op <::nat => nat => bool)
+               ((op div::nat => nat => nat) m
+                 ((number_of::bin => nat)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool))))
+               ((op div::nat => nat => nat) n
+                 ((number_of::bin => nat)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool)))))
+             ((op <::nat => nat => bool) m n))))"
+  by (import prob_extra DIV_TWO_MONO_EVEN)
+
+lemma DIV_TWO_CANCEL: "ALL n. 2 * n div 2 = n & Suc (2 * n) div 2 = n"
+  by (import prob_extra DIV_TWO_CANCEL)
+
+lemma EXP_DIV_TWO: "ALL n::nat. (2::nat) ^ Suc n div (2::nat) = (2::nat) ^ n"
+  by (import prob_extra EXP_DIV_TWO)
+
+lemma EVEN_EXP_TWO: "ALL n. EVEN (2 ^ n) = (n ~= 0)"
+  by (import prob_extra EVEN_EXP_TWO)
+
+lemma DIV_TWO_EXP: "ALL (n::nat) k::nat.
+   (k div (2::nat) < (2::nat) ^ n) = (k < (2::nat) ^ Suc n)"
+  by (import prob_extra DIV_TWO_EXP)
+
+consts
+  inf :: "(real => bool) => real" 
+
+defs
+  inf_primdef: "inf == %P. - sup (IMAGE uminus P)"
+
+lemma inf_def: "ALL P. inf P = - sup (IMAGE uminus P)"
+  by (import prob_extra inf_def)
+
+lemma INF_DEF_ALT: "ALL P. inf P = - sup (%r. P (- r))"
+  by (import prob_extra INF_DEF_ALT)
+
+lemma REAL_SUP_EXISTS_UNIQUE: "(All::((real => bool) => bool) => bool)
+ (%P::real => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) ((Ex::(real => bool) => bool) P)
+        ((Ex::(real => bool) => bool)
+          (%z::real.
+              (All::(real => bool) => bool)
+               (%x::real.
+                   (op -->::bool => bool => bool) (P x)
+                    ((op <=::real => real => bool) x z)))))
+      ((Ex1::(real => bool) => bool)
+        (%s::real.
+            (All::(real => bool) => bool)
+             (%y::real.
+                 (op =::bool => bool => bool)
+                  ((Ex::(real => bool) => bool)
+                    (%x::real.
+                        (op &::bool => bool => bool) (P x)
+                         ((op <::real => real => bool) y x)))
+                  ((op <::real => real => bool) y s)))))"
+  by (import prob_extra REAL_SUP_EXISTS_UNIQUE)
+
+lemma REAL_SUP_MAX: "(All::((real => bool) => bool) => bool)
+ (%P::real => bool.
+     (All::(real => bool) => bool)
+      (%z::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool) (P z)
+             ((All::(real => bool) => bool)
+               (%x::real.
+                   (op -->::bool => bool => bool) (P x)
+                    ((op <=::real => real => bool) x z))))
+           ((op =::real => real => bool) ((sup::(real => bool) => real) P)
+             z)))"
+  by (import prob_extra REAL_SUP_MAX)
+
+lemma REAL_INF_MIN: "(All::((real => bool) => bool) => bool)
+ (%P::real => bool.
+     (All::(real => bool) => bool)
+      (%z::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool) (P z)
+             ((All::(real => bool) => bool)
+               (%x::real.
+                   (op -->::bool => bool => bool) (P x)
+                    ((op <=::real => real => bool) z x))))
+           ((op =::real => real => bool) ((inf::(real => bool) => real) P)
+             z)))"
+  by (import prob_extra REAL_INF_MIN)
+
+lemma HALF_POS: "(0::real) < (1::real) / (2::real)"
+  by (import prob_extra HALF_POS)
+
+lemma HALF_CANCEL: "(2::real) * ((1::real) / (2::real)) = (1::real)"
+  by (import prob_extra HALF_CANCEL)
+
+lemma POW_HALF_POS: "ALL n::nat. (0::real) < ((1::real) / (2::real)) ^ n"
+  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::bin => real)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool))))
+               n)
+             ((op ^::real => nat => real)
+               ((op /::real => real => real) (1::real)
+                 ((number_of::bin => real)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool))))
+               m))))"
+  by (import prob_extra POW_HALF_MONO)
+
+lemma POW_HALF_TWICE: "ALL n::nat.
+   ((1::real) / (2::real)) ^ n = (2::real) * ((1::real) / (2::real)) ^ Suc n"
+  by (import prob_extra POW_HALF_TWICE)
+
+lemma X_HALF_HALF: "ALL x::real. (1::real) / (2::real) * x + (1::real) / (2::real) * x = x"
+  by (import prob_extra X_HALF_HALF)
+
+lemma REAL_SUP_LE_X: "(All::((real => bool) => bool) => bool)
+ (%P::real => bool.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool) ((Ex::(real => bool) => bool) P)
+             ((All::(real => bool) => bool)
+               (%r::real.
+                   (op -->::bool => bool => bool) (P r)
+                    ((op <=::real => real => bool) r x))))
+           ((op <=::real => real => bool) ((sup::(real => bool) => real) P)
+             x)))"
+  by (import prob_extra REAL_SUP_LE_X)
+
+lemma REAL_X_LE_SUP: "(All::((real => bool) => bool) => bool)
+ (%P::real => bool.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool) ((Ex::(real => bool) => bool) P)
+             ((op &::bool => bool => bool)
+               ((Ex::(real => bool) => bool)
+                 (%z::real.
+                     (All::(real => bool) => bool)
+                      (%r::real.
+                          (op -->::bool => bool => bool) (P r)
+                           ((op <=::real => real => bool) r z))))
+               ((Ex::(real => bool) => bool)
+                 (%r::real.
+                     (op &::bool => bool => bool) (P r)
+                      ((op <=::real => real => bool) x r)))))
+           ((op <=::real => real => bool) x
+             ((sup::(real => bool) => real) P))))"
+  by (import prob_extra REAL_X_LE_SUP)
+
+lemma ABS_BETWEEN_LE: "ALL (x::real) (y::real) d::real.
+   ((0::real) <= d & x - d <= y & y <= x + d) = (abs (y - x) <= d)"
+  by (import prob_extra ABS_BETWEEN_LE)
+
+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: "(1::real) / (2::real) < (1::real)"
+  by (import prob_extra HALF_LT_1)
+
+lemma POW_HALF_EXP: "ALL n::nat. ((1::real) / (2::real)) ^ n = inverse (real ((2::nat) ^ n))"
+  by (import prob_extra POW_HALF_EXP)
+
+lemma INV_SUC_POS: "ALL n. 0 < 1 / real (Suc n)"
+  by (import prob_extra INV_SUC_POS)
+
+lemma INV_SUC_MAX: "ALL x. 1 / real (Suc x) <= 1"
+  by (import prob_extra INV_SUC_MAX)
+
+lemma INV_SUC: "ALL n. 0 < 1 / real (Suc n) & 1 / real (Suc n) <= 1"
+  by (import prob_extra INV_SUC)
+
+lemma ABS_UNIT_INTERVAL: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <=::real => real => bool) x (1::real))
+               ((op &::bool => bool => bool)
+                 ((op <=::real => real => bool) (0::real) y)
+                 ((op <=::real => real => bool) y (1::real)))))
+           ((op <=::real => real => bool)
+             ((abs::real => real) ((op -::real => real => real) x y))
+             (1::real))))"
+  by (import prob_extra ABS_UNIT_INTERVAL)
+
+lemma MEM_NIL: "ALL l. (ALL x. ~ x mem l) = (l = [])"
+  by (import prob_extra MEM_NIL)
+
+lemma MAP_MEM: "ALL f l x. x mem map f l = (EX y. y mem l & x = f y)"
+  by (import prob_extra MAP_MEM)
+
+lemma MEM_NIL_MAP_CONS: "ALL x l. ~ [] mem map (op # x) l"
+  by (import prob_extra MEM_NIL_MAP_CONS)
+
+lemma FILTER_TRUE: "ALL l. [x:l. True] = l"
+  by (import prob_extra FILTER_TRUE)
+
+lemma FILTER_FALSE: "ALL l. [x:l. False] = []"
+  by (import prob_extra FILTER_FALSE)
+
+lemma FILTER_MEM: "(All::(('a => bool) => bool) => bool)
+ (%P::'a => bool.
+     (All::('a => bool) => bool)
+      (%x::'a.
+          (All::('a list => bool) => bool)
+           (%l::'a list.
+               (op -->::bool => bool => bool)
+                ((op mem::'a => 'a list => bool) x
+                  ((filter::('a => bool) => 'a list => 'a list) P l))
+                (P x))))"
+  by (import prob_extra FILTER_MEM)
+
+lemma MEM_FILTER: "(All::(('a => bool) => bool) => bool)
+ (%P::'a => bool.
+     (All::('a list => bool) => bool)
+      (%l::'a list.
+          (All::('a => bool) => bool)
+           (%x::'a.
+               (op -->::bool => bool => bool)
+                ((op mem::'a => 'a list => bool) x
+                  ((filter::('a => bool) => 'a list => 'a list) P l))
+                ((op mem::'a => 'a list => bool) x l))))"
+  by (import prob_extra MEM_FILTER)
+
+lemma FILTER_OUT_ELT: "ALL x l. x mem l | [y:l. y ~= x] = l"
+  by (import prob_extra FILTER_OUT_ELT)
+
+lemma IS_PREFIX_NIL: "ALL x. IS_PREFIX x [] & IS_PREFIX [] x = (x = [])"
+  by (import prob_extra IS_PREFIX_NIL)
+
+lemma IS_PREFIX_REFL: "ALL x. IS_PREFIX x x"
+  by (import prob_extra IS_PREFIX_REFL)
+
+lemma IS_PREFIX_ANTISYM: "(All::('a list => bool) => bool)
+ (%x::'a list.
+     (All::('a list => bool) => bool)
+      (%y::'a list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((IS_PREFIX::'a list => 'a list => bool) y x)
+             ((IS_PREFIX::'a list => 'a list => bool) x y))
+           ((op =::'a list => 'a list => bool) x y)))"
+  by (import prob_extra IS_PREFIX_ANTISYM)
+
+lemma IS_PREFIX_TRANS: "(All::('a list => bool) => bool)
+ (%x::'a list.
+     (All::('a list => bool) => bool)
+      (%y::'a list.
+          (All::('a list => bool) => bool)
+           (%z::'a list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((IS_PREFIX::'a list => 'a list => bool) x y)
+                  ((IS_PREFIX::'a list => 'a list => bool) y z))
+                ((IS_PREFIX::'a list => 'a list => bool) x z))))"
+  by (import prob_extra IS_PREFIX_TRANS)
+
+lemma IS_PREFIX_BUTLAST: "ALL x y. IS_PREFIX (x # y) (butlast (x # y))"
+  by (import prob_extra IS_PREFIX_BUTLAST)
+
+lemma IS_PREFIX_LENGTH: "(All::('a list => bool) => bool)
+ (%x::'a list.
+     (All::('a list => bool) => bool)
+      (%y::'a list.
+          (op -->::bool => bool => bool)
+           ((IS_PREFIX::'a list => 'a list => bool) y x)
+           ((op <=::nat => nat => bool) ((size::'a list => nat) x)
+             ((size::'a list => nat) y))))"
+  by (import prob_extra IS_PREFIX_LENGTH)
+
+lemma IS_PREFIX_LENGTH_ANTI: "(All::('a list => bool) => bool)
+ (%x::'a list.
+     (All::('a list => bool) => bool)
+      (%y::'a list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((IS_PREFIX::'a list => 'a list => bool) y x)
+             ((op =::nat => nat => bool) ((size::'a list => nat) x)
+               ((size::'a list => nat) y)))
+           ((op =::'a list => 'a list => bool) x y)))"
+  by (import prob_extra IS_PREFIX_LENGTH_ANTI)
+
+lemma IS_PREFIX_SNOC: "ALL x y z. 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 => 'c => 'c) (e::'c) (g::'a => 'b) l::'a list.
+   foldr f (map g l) e = foldr (%x::'a. f (g x)) l e"
+  by (import prob_extra FOLDR_MAP)
+
+lemma LAST_MEM: "ALL h t. last (h # t) mem 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"
+  by (import prob_extra LAST_MAP_CONS)
+
+lemma EXISTS_LONGEST: "(All::('a list => bool) => bool)
+ (%x::'a list.
+     (All::('a list list => bool) => bool)
+      (%y::'a list list.
+          (Ex::('a list => bool) => bool)
+           (%z::'a list.
+               (op &::bool => bool => bool)
+                ((op mem::'a list => 'a list list => bool) z
+                  ((op #::'a list => 'a list list => 'a list list) x y))
+                ((All::('a list => bool) => bool)
+                  (%w::'a list.
+                      (op -->::bool => bool => bool)
+                       ((op mem::'a list => 'a list list => bool) w
+                         ((op #::'a list => 'a list list => 'a list list) x
+                           y))
+                       ((op <=::nat => nat => bool)
+                         ((size::'a list => nat) w)
+                         ((size::'a list => nat) z)))))))"
+  by (import prob_extra EXISTS_LONGEST)
+
+lemma UNION_DEF_ALT: "ALL s t. pred_set.UNION s t = (%x. s x | t x)"
+  by (import prob_extra UNION_DEF_ALT)
+
+lemma INTER_UNION_RDISTRIB: "ALL p q r.
+   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 xa. (x = xa) = (SUBSET x xa & SUBSET xa x)"
+  by (import prob_extra SUBSET_EQ)
+
+lemma INTER_IS_EMPTY: "ALL s t. (pred_set.INTER s t = EMPTY) = (ALL x. ~ s x | ~ t x)"
+  by (import prob_extra INTER_IS_EMPTY)
+
+lemma UNION_DISJOINT_SPLIT: "(All::(('a => bool) => bool) => bool)
+ (%s::'a => bool.
+     (All::(('a => bool) => bool) => bool)
+      (%t::'a => bool.
+          (All::(('a => bool) => bool) => bool)
+           (%u::'a => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op =::('a => bool) => ('a => bool) => bool)
+                    ((pred_set.UNION::('a => bool)
+=> ('a => bool) => 'a => bool)
+                      s t)
+                    ((pred_set.UNION::('a => bool)
+=> ('a => bool) => 'a => bool)
+                      s u))
+                  ((op &::bool => bool => bool)
+                    ((op =::('a => bool) => ('a => bool) => bool)
+                      ((pred_set.INTER::('a => bool)
+  => ('a => bool) => 'a => bool)
+                        s t)
+                      (EMPTY::'a => bool))
+                    ((op =::('a => bool) => ('a => bool) => bool)
+                      ((pred_set.INTER::('a => bool)
+  => ('a => bool) => 'a => bool)
+                        s u)
+                      (EMPTY::'a => bool))))
+                ((op =::('a => bool) => ('a => bool) => bool) t u))))"
+  by (import prob_extra UNION_DISJOINT_SPLIT)
+
+lemma GSPEC_DEF_ALT: "ALL f. GSPEC f = (%v. EX x. (v, True) = f x)"
+  by (import prob_extra GSPEC_DEF_ALT)
+
+;end_setup
+
+;setup_theory prob_canon
+
+consts
+  alg_twin :: "bool list => bool list => bool" 
+
+defs
+  alg_twin_primdef: "alg_twin == %x y. EX l. x = SNOC True l & y = SNOC False l"
+
+lemma alg_twin_def: "ALL x y. alg_twin x y = (EX l. x = SNOC True l & y = SNOC False l)"
+  by (import prob_canon alg_twin_def)
+
+constdefs
+  alg_order_tupled :: "bool list * bool list => bool" 
+  "(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)))"
+
+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)))"
+  by (import prob_canon alg_order_tupled_primitive_def)
+
+consts
+  alg_order :: "bool list => bool list => bool" 
+
+defs
+  alg_order_primdef: "alg_order == %x x1. alg_order_tupled (x, x1)"
+
+lemma alg_order_curried_def: "ALL x x1. alg_order x x1 = alg_order_tupled (x, x1)"
+  by (import prob_canon alg_order_curried_def)
+
+lemma alg_order_ind: "(All::((bool list => bool list => bool) => bool) => bool)
+ (%P::bool list => bool list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(bool => bool) => bool)
+          (%x::bool.
+              (All::(bool list => bool) => bool)
+               (%xa::bool list.
+                   P ([]::bool list)
+                    ((op #::bool => bool list => bool list) x xa))))
+        ((op &::bool => bool => bool) (P ([]::bool list) ([]::bool list))
+          ((op &::bool => bool => bool)
+            ((All::(bool => bool) => bool)
+              (%x::bool.
+                  (All::(bool list => bool) => bool)
+                   (%xa::bool list.
+                       P ((op #::bool => bool list => bool list) x xa)
+                        ([]::bool list))))
+            ((All::(bool => bool) => bool)
+              (%x::bool.
+                  (All::(bool list => bool) => bool)
+                   (%xa::bool list.
+                       (All::(bool => bool) => bool)
+                        (%xb::bool.
+                            (All::(bool list => bool) => bool)
+                             (%xc::bool list.
+                                 (op -->::bool => bool => bool) (P xa xc)
+                                  (P ((op #::bool => bool list => bool list)
+ x xa)
+                                    ((op #::bool => bool list => bool list)
+xb xc))))))))))
+      ((All::(bool list => bool) => bool)
+        (%x::bool list. (All::(bool list => bool) => bool) (P x))))"
+  by (import prob_canon alg_order_ind)
+
+lemma alg_order_def: "alg_order [] (v6 # v7) = True &
+alg_order [] [] = True &
+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)
+
+consts
+  alg_sorted :: "bool list list => bool" 
+
+defs
+  alg_sorted_primdef: "alg_sorted ==
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_sorted.
+     list_case True
+      (%v2. list_case True
+             (%v6 v7. alg_order v2 v6 & alg_sorted (v6 # v7))))"
+
+lemma alg_sorted_primitive_def: "alg_sorted =
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_sorted.
+     list_case True
+      (%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::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(bool list => bool) => bool)
+          (%x::bool list.
+              (All::(bool list => bool) => bool)
+               (%y::bool list.
+                   (All::(bool list list => bool) => bool)
+                    (%z::bool list list.
+                        (op -->::bool => bool => bool)
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              y z))
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              x ((op #::bool list
+  => bool list list => bool list list)
+                                  y z)))))))
+        ((op &::bool => bool => bool)
+          ((All::(bool list => bool) => bool)
+            (%v::bool list.
+                P ((op #::bool list => bool list list => bool list list) v
+                    ([]::bool list list))))
+          (P ([]::bool list list))))
+      ((All::(bool list list => bool) => bool) P))"
+  by (import prob_canon alg_sorted_ind)
+
+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
+  alg_prefixfree :: "bool list list => bool" 
+
+defs
+  alg_prefixfree_primdef: "alg_prefixfree ==
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_prefixfree.
+     list_case True
+      (%v2. list_case True
+             (%v6 v7. ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))"
+
+lemma alg_prefixfree_primitive_def: "alg_prefixfree =
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_prefixfree.
+     list_case True
+      (%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::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(bool list => bool) => bool)
+          (%x::bool list.
+              (All::(bool list => bool) => bool)
+               (%y::bool list.
+                   (All::(bool list list => bool) => bool)
+                    (%z::bool list list.
+                        (op -->::bool => bool => bool)
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              y z))
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              x ((op #::bool list
+  => bool list list => bool list list)
+                                  y z)))))))
+        ((op &::bool => bool => bool)
+          ((All::(bool list => bool) => bool)
+            (%v::bool list.
+                P ((op #::bool list => bool list list => bool list list) v
+                    ([]::bool list list))))
+          (P ([]::bool list list))))
+      ((All::(bool list list => bool) => bool) P))"
+  by (import prob_canon alg_prefixfree_ind)
+
+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
+  alg_twinfree :: "bool list list => bool" 
+
+defs
+  alg_twinfree_primdef: "alg_twinfree ==
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_twinfree.
+     list_case True
+      (%v2. list_case True
+             (%v6 v7. ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))"
+
+lemma alg_twinfree_primitive_def: "alg_twinfree =
+WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z)))
+ (%alg_twinfree.
+     list_case True
+      (%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::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(bool list => bool) => bool)
+          (%x::bool list.
+              (All::(bool list => bool) => bool)
+               (%y::bool list.
+                   (All::(bool list list => bool) => bool)
+                    (%z::bool list list.
+                        (op -->::bool => bool => bool)
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              y z))
+                         (P ((op #::bool list
+                                    => bool list list => bool list list)
+                              x ((op #::bool list
+  => bool list list => bool list list)
+                                  y z)))))))
+        ((op &::bool => bool => bool)
+          ((All::(bool list => bool) => bool)
+            (%v::bool list.
+                P ((op #::bool list => bool list list => bool list list) v
+                    ([]::bool list list))))
+          (P ([]::bool list list))))
+      ((All::(bool list list => bool) => bool) P))"
+  by (import prob_canon alg_twinfree_ind)
+
+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 t. 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. 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)
+
+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. 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
+     else alg_canon_prefs l (h # t)))"
+  by (import prob_canon alg_canon_find_def)
+
+consts
+  alg_canon1 :: "bool list list => bool list list" 
+
+defs
+  alg_canon1_primdef: "alg_canon1 == FOLDR alg_canon_find []"
+
+lemma alg_canon1_def: "alg_canon1 = FOLDR alg_canon_find []"
+  by (import prob_canon alg_canon1_def)
+
+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. 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)
+
+consts
+  alg_canon2 :: "bool list list => bool list list" 
+
+defs
+  alg_canon2_primdef: "alg_canon2 == FOLDR alg_canon_merge []"
+
+lemma alg_canon2_def: "alg_canon2 = FOLDR alg_canon_merge []"
+  by (import prob_canon alg_canon2_def)
+
+consts
+  alg_canon :: "bool list list => bool list list" 
+
+defs
+  alg_canon_primdef: "alg_canon == %l. alg_canon2 (alg_canon1 l)"
+
+lemma alg_canon_def: "ALL l. 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. alg_canon l = l"
+
+lemma algebra_canon_def: "ALL l. algebra_canon l = (alg_canon l = l)"
+  by (import prob_canon algebra_canon_def)
+
+lemma ALG_TWIN_NIL: "ALL l. ~ alg_twin l [] & ~ alg_twin [] l"
+  by (import prob_canon ALG_TWIN_NIL)
+
+lemma ALG_TWIN_SING: "ALL x l.
+   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 y z h t.
+   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 t t'. alg_twin (h # t) (h # t') = alg_twin t t'"
+  by (import prob_canon ALG_TWIN_REDUCE)
+
+lemma ALG_TWINS_PREFIX: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%l::bool list.
+          (op -->::bool => bool => bool)
+           ((IS_PREFIX::bool list => bool list => bool) x l)
+           ((op |::bool => bool => bool)
+             ((op =::bool list => bool list => bool) x l)
+             ((op |::bool => bool => bool)
+               ((IS_PREFIX::bool list => bool list => bool) x
+                 ((SNOC::bool => bool list => bool list) (True::bool) l))
+               ((IS_PREFIX::bool list => bool list => bool) x
+                 ((SNOC::bool => bool list => bool list) (False::bool)
+                   l))))))"
+  by (import prob_canon ALG_TWINS_PREFIX)
+
+lemma ALG_ORDER_NIL: "ALL x. alg_order [] x & alg_order x [] = (x = [])"
+  by (import prob_canon ALG_ORDER_NIL)
+
+lemma ALG_ORDER_REFL: "ALL x. alg_order x x"
+  by (import prob_canon ALG_ORDER_REFL)
+
+lemma ALG_ORDER_ANTISYM: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_order::bool list => bool list => bool) x y)
+             ((alg_order::bool list => bool list => bool) y x))
+           ((op =::bool list => bool list => bool) x y)))"
+  by (import prob_canon ALG_ORDER_ANTISYM)
+
+lemma ALG_ORDER_TRANS: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (All::(bool list => bool) => bool)
+           (%z::bool list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_order::bool list => bool list => bool) x y)
+                  ((alg_order::bool list => bool list => bool) y z))
+                ((alg_order::bool list => bool list => bool) x z))))"
+  by (import prob_canon ALG_ORDER_TRANS)
+
+lemma ALG_ORDER_TOTAL: "ALL x y. alg_order x y | alg_order y x"
+  by (import prob_canon ALG_ORDER_TOTAL)
+
+lemma ALG_ORDER_PREFIX: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (op -->::bool => bool => bool)
+           ((IS_PREFIX::bool list => bool list => bool) y x)
+           ((alg_order::bool list => bool list => bool) x y)))"
+  by (import prob_canon ALG_ORDER_PREFIX)
+
+lemma ALG_ORDER_PREFIX_ANTI: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_order::bool list => bool list => bool) x y)
+             ((IS_PREFIX::bool list => bool list => bool) x y))
+           ((op =::bool list => bool list => bool) x y)))"
+  by (import prob_canon ALG_ORDER_PREFIX_ANTI)
+
+lemma ALG_ORDER_PREFIX_MONO: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (All::(bool list => bool) => bool)
+           (%z::bool list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_order::bool list => bool list => bool) x y)
+                  ((op &::bool => bool => bool)
+                    ((alg_order::bool list => bool list => bool) y z)
+                    ((IS_PREFIX::bool list => bool list => bool) z x)))
+                ((IS_PREFIX::bool list => bool list => bool) y x))))"
+  by (import prob_canon ALG_ORDER_PREFIX_MONO)
+
+lemma ALG_ORDER_PREFIX_TRANS: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (All::(bool list => bool) => bool)
+           (%z::bool list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_order::bool list => bool list => bool) x y)
+                  ((IS_PREFIX::bool list => bool list => bool) y z))
+                ((op |::bool => bool => bool)
+                  ((alg_order::bool list => bool list => bool) x z)
+                  ((IS_PREFIX::bool list => bool list => bool) x z)))))"
+  by (import prob_canon ALG_ORDER_PREFIX_TRANS)
+
+lemma ALG_ORDER_SNOC: "ALL x l. ~ alg_order (SNOC x l) l"
+  by (import prob_canon ALG_ORDER_SNOC)
+
+lemma ALG_SORTED_MIN: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_sorted::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((All::(bool list => bool) => bool)
+             (%x::bool list.
+                 (op -->::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) x t)
+                  ((alg_order::bool list => bool list => bool) h x)))))"
+  by (import prob_canon ALG_SORTED_MIN)
+
+lemma ALG_SORTED_DEF_ALT: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op =::bool => bool => bool)
+           ((alg_sorted::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((op &::bool => bool => bool)
+             ((All::(bool list => bool) => bool)
+               (%x::bool list.
+                   (op -->::bool => bool => bool)
+                    ((op mem::bool list => bool list list => bool) x t)
+                    ((alg_order::bool list => bool list => bool) h x)))
+             ((alg_sorted::bool list list => bool) t))))"
+  by (import prob_canon ALG_SORTED_DEF_ALT)
+
+lemma ALG_SORTED_TL: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_sorted::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((alg_sorted::bool list list => bool) t)))"
+  by (import prob_canon ALG_SORTED_TL)
+
+lemma ALG_SORTED_MONO: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (All::(bool list list => bool) => bool)
+           (%z::bool list list.
+               (op -->::bool => bool => bool)
+                ((alg_sorted::bool list list => bool)
+                  ((op #::bool list => bool list list => bool list list) x
+                    ((op #::bool list => bool list list => bool list list) y
+                      z)))
+                ((alg_sorted::bool list list => bool)
+                  ((op #::bool list => bool list list => bool list list) x
+                    z)))))"
+  by (import prob_canon ALG_SORTED_MONO)
+
+lemma ALG_SORTED_TLS: "ALL l b. alg_sorted (map (op # b) l) = alg_sorted l"
+  by (import prob_canon ALG_SORTED_TLS)
+
+lemma ALG_SORTED_STEP: "ALL l1 l2.
+   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 h' t t'.
+   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::((bool list => bool) => bool) => bool)
+ (%P::bool list => bool.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_sorted::bool list list => bool) b)
+           ((alg_sorted::bool list list => bool)
+             ((filter::(bool list => bool)
+                       => bool list list => bool list list)
+               P b))))"
+  by (import prob_canon ALG_SORTED_FILTER)
+
+lemma ALG_PREFIXFREE_TL: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_prefixfree::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((alg_prefixfree::bool list list => bool) t)))"
+  by (import prob_canon ALG_PREFIXFREE_TL)
+
+lemma ALG_PREFIXFREE_MONO: "(All::(bool list => bool) => bool)
+ (%x::bool list.
+     (All::(bool list => bool) => bool)
+      (%y::bool list.
+          (All::(bool list list => bool) => bool)
+           (%z::bool list list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_sorted::bool list list => bool)
+                    ((op #::bool list => bool list list => bool list list) x
+                      ((op #::bool list => bool list list => bool list list)
+                        y z)))
+                  ((alg_prefixfree::bool list list => bool)
+                    ((op #::bool list => bool list list => bool list list) x
+                      ((op #::bool list => bool list list => bool list list)
+                        y z))))
+                ((alg_prefixfree::bool list list => bool)
+                  ((op #::bool list => bool list list => bool list list) x
+                    z)))))"
+  by (import prob_canon ALG_PREFIXFREE_MONO)
+
+lemma ALG_PREFIXFREE_ELT: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool)
+               ((op #::bool list => bool list list => bool list list) h t))
+             ((alg_prefixfree::bool list list => bool)
+               ((op #::bool list => bool list list => bool list list) h t)))
+           ((All::(bool list => bool) => bool)
+             (%x::bool list.
+                 (op -->::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) x t)
+                  ((op &::bool => bool => bool)
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) x h))
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) h
+                        x)))))))"
+  by (import prob_canon ALG_PREFIXFREE_ELT)
+
+lemma ALG_PREFIXFREE_TLS: "ALL l b. alg_prefixfree (map (op # b) l) = alg_prefixfree l"
+  by (import prob_canon ALG_PREFIXFREE_TLS)
+
+lemma ALG_PREFIXFREE_STEP: "ALL l1 l2.
+   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 h' t t'.
+   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::((bool list => bool) => bool) => bool)
+ (%P::bool list => bool.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool) b)
+             ((alg_prefixfree::bool list list => bool) b))
+           ((alg_prefixfree::bool list list => bool)
+             ((filter::(bool list => bool)
+                       => bool list list => bool list list)
+               P b))))"
+  by (import prob_canon ALG_PREFIXFREE_FILTER)
+
+lemma ALG_TWINFREE_TL: "(All::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_twinfree::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((alg_twinfree::bool list list => bool) t)))"
+  by (import prob_canon ALG_TWINFREE_TL)
+
+lemma ALG_TWINFREE_TLS: "ALL l b. alg_twinfree (map (op # b) l) = alg_twinfree l"
+  by (import prob_canon ALG_TWINFREE_TLS)
+
+lemma ALG_TWINFREE_STEP1: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_twinfree::bool list list => bool)
+             ((op @::bool list list => bool list list => bool list list)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (True::bool)) l1)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (False::bool))
+                 l2)))
+           ((op &::bool => bool => bool)
+             ((alg_twinfree::bool list list => bool) l1)
+             ((alg_twinfree::bool list list => bool) l2))))"
+  by (import prob_canon ALG_TWINFREE_STEP1)
+
+lemma ALG_TWINFREE_STEP2: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op |::bool => bool => bool)
+               ((Not::bool => bool)
+                 ((op mem::bool list => bool list list => bool)
+                   ([]::bool list) l1))
+               ((Not::bool => bool)
+                 ((op mem::bool list => bool list list => bool)
+                   ([]::bool list) l2)))
+             ((op &::bool => bool => bool)
+               ((alg_twinfree::bool list list => bool) l1)
+               ((alg_twinfree::bool list list => bool) l2)))
+           ((alg_twinfree::bool list list => bool)
+             ((op @::bool list list => bool list list => bool list list)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (True::bool)) l1)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (False::bool))
+                 l2)))))"
+  by (import prob_canon ALG_TWINFREE_STEP2)
+
+lemma ALG_TWINFREE_STEP: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((op |::bool => bool => bool)
+             ((Not::bool => bool)
+               ((op mem::bool list => bool list list => bool)
+                 ([]::bool list) l1))
+             ((Not::bool => bool)
+               ((op mem::bool list => bool list list => bool)
+                 ([]::bool list) l2)))
+           ((op =::bool => bool => bool)
+             ((alg_twinfree::bool list list => bool)
+               ((op @::bool list list => bool list list => bool list list)
+                 ((map::(bool list => bool list)
+                        => bool list list => bool list list)
+                   ((op #::bool => bool list => bool list) (True::bool)) l1)
+                 ((map::(bool list => bool list)
+                        => bool list list => bool list list)
+                   ((op #::bool => bool list => bool list) (False::bool))
+                   l2)))
+             ((op &::bool => bool => bool)
+               ((alg_twinfree::bool list list => bool) l1)
+               ((alg_twinfree::bool list list => bool) l2)))))"
+  by (import prob_canon ALG_TWINFREE_STEP)
+
+lemma ALG_LONGEST_HD: "ALL h t. length h <= alg_longest (h # t)"
+  by (import prob_canon ALG_LONGEST_HD)
+
+lemma ALG_LONGEST_TL: "ALL h t. alg_longest t <= alg_longest (h # t)"
+  by (import prob_canon ALG_LONGEST_TL)
+
+lemma ALG_LONGEST_TLS: "ALL h t b. 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 l2.
+   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 b. hd (alg_canon_prefs l b) = l"
+  by (import prob_canon ALG_CANON_PREFS_HD)
+
+lemma ALG_CANON_PREFS_DELETES: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (All::(bool list => bool) => bool)
+           (%x::bool list.
+               (op -->::bool => bool => bool)
+                ((op mem::bool list => bool list list => bool) x
+                  ((alg_canon_prefs::bool list
+                                     => bool list list => bool list list)
+                    l b))
+                ((op mem::bool list => bool list list => bool) x
+                  ((op #::bool list => bool list list => bool list list) l
+                    b)))))"
+  by (import prob_canon ALG_CANON_PREFS_DELETES)
+
+lemma ALG_CANON_PREFS_SORTED: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_sorted::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) l b))
+           ((alg_sorted::bool list list => bool)
+             ((alg_canon_prefs::bool list
+                                => bool list list => bool list list)
+               l b))))"
+  by (import prob_canon ALG_CANON_PREFS_SORTED)
+
+lemma ALG_CANON_PREFS_PREFIXFREE: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool) b)
+             ((alg_prefixfree::bool list list => bool) b))
+           ((alg_prefixfree::bool list list => bool)
+             ((alg_canon_prefs::bool list
+                                => bool list list => bool list list)
+               l b))))"
+  by (import prob_canon ALG_CANON_PREFS_PREFIXFREE)
+
+lemma ALG_CANON_PREFS_CONSTANT: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_prefixfree::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) l b))
+           ((op =::bool list list => bool list list => bool)
+             ((alg_canon_prefs::bool list
+                                => bool list list => bool list list)
+               l b)
+             ((op #::bool list => bool list list => bool list list) l b))))"
+  by (import prob_canon ALG_CANON_PREFS_CONSTANT)
+
+lemma ALG_CANON_FIND_HD: "ALL l h t.
+   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::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (All::(bool list => bool) => bool)
+           (%x::bool list.
+               (op -->::bool => bool => bool)
+                ((op mem::bool list => bool list list => bool) x
+                  ((alg_canon_find::bool list
+                                    => bool list list => bool list list)
+                    l b))
+                ((op mem::bool list => bool list list => bool) x
+                  ((op #::bool list => bool list list => bool list list) l
+                    b)))))"
+  by (import prob_canon ALG_CANON_FIND_DELETES)
+
+lemma ALG_CANON_FIND_SORTED: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_sorted::bool list list => bool) b)
+           ((alg_sorted::bool list list => bool)
+             ((alg_canon_find::bool list
+                               => bool list list => bool list list)
+               l b))))"
+  by (import prob_canon ALG_CANON_FIND_SORTED)
+
+lemma ALG_CANON_FIND_PREFIXFREE: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool) b)
+             ((alg_prefixfree::bool list list => bool) b))
+           ((alg_prefixfree::bool list list => bool)
+             ((alg_canon_find::bool list
+                               => bool list list => bool list list)
+               l b))))"
+  by (import prob_canon ALG_CANON_FIND_PREFIXFREE)
+
+lemma ALG_CANON_FIND_CONSTANT: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool)
+               ((op #::bool list => bool list list => bool list list) l b))
+             ((alg_prefixfree::bool list list => bool)
+               ((op #::bool list => bool list list => bool list list) l b)))
+           ((op =::bool list list => bool list list => bool)
+             ((alg_canon_find::bool list
+                               => bool list list => bool list list)
+               l b)
+             ((op #::bool list => bool list list => bool list list) l b))))"
+  by (import prob_canon ALG_CANON_FIND_CONSTANT)
+
+lemma ALG_CANON1_SORTED: "ALL x. alg_sorted (alg_canon1 x)"
+  by (import prob_canon ALG_CANON1_SORTED)
+
+lemma ALG_CANON1_PREFIXFREE: "ALL l. alg_prefixfree (alg_canon1 l)"
+  by (import prob_canon ALG_CANON1_PREFIXFREE)
+
+lemma ALG_CANON1_CONSTANT: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) ((alg_sorted::bool list list => bool) l)
+        ((alg_prefixfree::bool list list => bool) l))
+      ((op =::bool list list => bool list list => bool)
+        ((alg_canon1::bool list list => bool list list) l) l))"
+  by (import prob_canon ALG_CANON1_CONSTANT)
+
+lemma ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool)
+               ((op #::bool list => bool list list => bool list list) l b))
+             ((op &::bool => bool => bool)
+               ((alg_prefixfree::bool list list => bool)
+                 ((op #::bool list => bool list list => bool list list) l
+                   b))
+               ((alg_twinfree::bool list list => bool) b)))
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool)
+               ((alg_canon_merge::bool list
+                                  => bool list list => bool list list)
+                 l b))
+             ((op &::bool => bool => bool)
+               ((alg_prefixfree::bool list list => bool)
+                 ((alg_canon_merge::bool list
+                                    => bool list list => bool list list)
+                   l b))
+               ((alg_twinfree::bool list list => bool)
+                 ((alg_canon_merge::bool list
+                                    => bool list list => bool list list)
+                   l b))))))"
+  by (import prob_canon ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE)
+
+lemma ALG_CANON_MERGE_PREFIXFREE_PRESERVE: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (All::(bool list => bool) => bool)
+           (%h::bool list.
+               (op -->::bool => bool => bool)
+                ((All::(bool list => bool) => bool)
+                  (%x::bool list.
+                      (op -->::bool => bool => bool)
+                       ((op mem::bool list => bool list list => bool) x
+                         ((op #::bool list
+                                 => bool list list => bool list list)
+                           l b))
+                       ((op &::bool => bool => bool)
+                         ((Not::bool => bool)
+                           ((IS_PREFIX::bool list => bool list => bool) h
+                             x))
+                         ((Not::bool => bool)
+                           ((IS_PREFIX::bool list => bool list => bool) x
+                             h)))))
+                ((All::(bool list => bool) => bool)
+                  (%x::bool list.
+                      (op -->::bool => bool => bool)
+                       ((op mem::bool list => bool list list => bool) x
+                         ((alg_canon_merge::bool list
+      => bool list list => bool list list)
+                           l b))
+                       ((op &::bool => bool => bool)
+                         ((Not::bool => bool)
+                           ((IS_PREFIX::bool list => bool list => bool) h
+                             x))
+                         ((Not::bool => bool)
+                           ((IS_PREFIX::bool list => bool list => bool) x
+                             h))))))))"
+  by (import prob_canon ALG_CANON_MERGE_PREFIXFREE_PRESERVE)
+
+lemma ALG_CANON_MERGE_SHORTENS: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (All::(bool list => bool) => bool)
+           (%x::bool list.
+               (op -->::bool => bool => bool)
+                ((op mem::bool list => bool list list => bool) x
+                  ((alg_canon_merge::bool list
+                                     => bool list list => bool list list)
+                    l b))
+                ((Ex::(bool list => bool) => bool)
+                  (%y::bool list.
+                      (op &::bool => bool => bool)
+                       ((op mem::bool list => bool list list => bool) y
+                         ((op #::bool list
+                                 => bool list list => bool list list)
+                           l b))
+                       ((IS_PREFIX::bool list => bool list => bool) y
+                         x))))))"
+  by (import prob_canon ALG_CANON_MERGE_SHORTENS)
+
+lemma ALG_CANON_MERGE_CONSTANT: "(All::(bool list => bool) => bool)
+ (%l::bool list.
+     (All::(bool list list => bool) => bool)
+      (%b::bool list list.
+          (op -->::bool => bool => bool)
+           ((alg_twinfree::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) l b))
+           ((op =::bool list list => bool list list => bool)
+             ((alg_canon_merge::bool list
+                                => bool list list => bool list list)
+               l b)
+             ((op #::bool list => bool list list => bool list list) l b))))"
+  by (import prob_canon ALG_CANON_MERGE_CONSTANT)
+
+lemma ALG_CANON2_PREFIXFREE_PRESERVE: "(All::(bool list list => bool) => bool)
+ (%x::bool list list.
+     (All::(bool list => bool) => bool)
+      (%xa::bool list.
+          (op -->::bool => bool => bool)
+           ((All::(bool list => bool) => bool)
+             (%xb::bool list.
+                 (op -->::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) xb x)
+                  ((op &::bool => bool => bool)
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) xa xb))
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) xb
+                        xa)))))
+           ((All::(bool list => bool) => bool)
+             (%xb::bool list.
+                 (op -->::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) xb
+                    ((alg_canon2::bool list list => bool list list) x))
+                  ((op &::bool => bool => bool)
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) xa xb))
+                    ((Not::bool => bool)
+                      ((IS_PREFIX::bool list => bool list => bool) xb
+                        xa)))))))"
+  by (import prob_canon ALG_CANON2_PREFIXFREE_PRESERVE)
+
+lemma ALG_CANON2_SHORTENS: "(All::(bool list list => bool) => bool)
+ (%x::bool list list.
+     (All::(bool list => bool) => bool)
+      (%xa::bool list.
+          (op -->::bool => bool => bool)
+           ((op mem::bool list => bool list list => bool) xa
+             ((alg_canon2::bool list list => bool list list) x))
+           ((Ex::(bool list => bool) => bool)
+             (%y::bool list.
+                 (op &::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) y x)
+                  ((IS_PREFIX::bool list => bool list => bool) y xa)))))"
+  by (import prob_canon ALG_CANON2_SHORTENS)
+
+lemma ALG_CANON2_SORTED_PREFIXFREE_TWINFREE: "(All::(bool list list => bool) => bool)
+ (%x::bool list list.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) ((alg_sorted::bool list list => bool) x)
+        ((alg_prefixfree::bool list list => bool) x))
+      ((op &::bool => bool => bool)
+        ((alg_sorted::bool list list => bool)
+          ((alg_canon2::bool list list => bool list list) x))
+        ((op &::bool => bool => bool)
+          ((alg_prefixfree::bool list list => bool)
+            ((alg_canon2::bool list list => bool list list) x))
+          ((alg_twinfree::bool list list => bool)
+            ((alg_canon2::bool list list => bool list list) x)))))"
+  by (import prob_canon ALG_CANON2_SORTED_PREFIXFREE_TWINFREE)
+
+lemma ALG_CANON2_CONSTANT: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_twinfree::bool list list => bool) l)
+      ((op =::bool list list => bool list list => bool)
+        ((alg_canon2::bool list list => bool list list) l) l))"
+  by (import prob_canon ALG_CANON2_CONSTANT)
+
+lemma ALG_CANON_SORTED_PREFIXFREE_TWINFREE: "ALL l.
+   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::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) ((alg_sorted::bool list list => bool) l)
+        ((op &::bool => bool => bool)
+          ((alg_prefixfree::bool list list => bool) l)
+          ((alg_twinfree::bool list list => bool) l)))
+      ((op =::bool list list => bool list list => bool)
+        ((alg_canon::bool list list => bool list list) l) l))"
+  by (import prob_canon ALG_CANON_CONSTANT)
+
+lemma ALG_CANON_IDEMPOT: "ALL l. alg_canon (alg_canon l) = alg_canon l"
+  by (import prob_canon ALG_CANON_IDEMPOT)
+
+lemma ALGEBRA_CANON_DEF_ALT: "ALL l. 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. algebra_canon [x])"
+  by (import prob_canon ALGEBRA_CANON_BASIC)
+
+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::(bool list => bool) => bool)
+ (%h::bool list.
+     (All::(bool list list => bool) => bool)
+      (%t::bool list list.
+          (op -->::bool => bool => bool)
+           ((algebra_canon::bool list list => bool)
+             ((op #::bool list => bool list list => bool list list) h t))
+           ((algebra_canon::bool list list => bool) t)))"
+  by (import prob_canon ALGEBRA_CANON_TL)
+
+lemma ALGEBRA_CANON_NIL_MEM: "ALL l. (algebra_canon l & [] mem l) = (l = [[]])"
+  by (import prob_canon ALGEBRA_CANON_NIL_MEM)
+
+lemma ALGEBRA_CANON_TLS: "ALL l b. algebra_canon (map (op # b) l) = algebra_canon l"
+  by (import prob_canon ALGEBRA_CANON_TLS)
+
+lemma ALGEBRA_CANON_STEP1: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((algebra_canon::bool list list => bool)
+             ((op @::bool list list => bool list list => bool list list)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (True::bool)) l1)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (False::bool))
+                 l2)))
+           ((op &::bool => bool => bool)
+             ((algebra_canon::bool list list => bool) l1)
+             ((algebra_canon::bool list list => bool) l2))))"
+  by (import prob_canon ALGEBRA_CANON_STEP1)
+
+lemma ALGEBRA_CANON_STEP2: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op |::bool => bool => bool)
+               ((Not::bool => bool)
+                 ((op =::bool list list => bool list list => bool) l1
+                   ((op #::bool list => bool list list => bool list list)
+                     ([]::bool list) ([]::bool list list))))
+               ((Not::bool => bool)
+                 ((op =::bool list list => bool list list => bool) l2
+                   ((op #::bool list => bool list list => bool list list)
+                     ([]::bool list) ([]::bool list list)))))
+             ((op &::bool => bool => bool)
+               ((algebra_canon::bool list list => bool) l1)
+               ((algebra_canon::bool list list => bool) l2)))
+           ((algebra_canon::bool list list => bool)
+             ((op @::bool list list => bool list list => bool list list)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (True::bool)) l1)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 ((op #::bool => bool list => bool list) (False::bool))
+                 l2)))))"
+  by (import prob_canon ALGEBRA_CANON_STEP2)
+
+lemma ALGEBRA_CANON_STEP: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (op -->::bool => bool => bool)
+           ((op |::bool => bool => bool)
+             ((Not::bool => bool)
+               ((op =::bool list list => bool list list => bool) l1
+                 ((op #::bool list => bool list list => bool list list)
+                   ([]::bool list) ([]::bool list list))))
+             ((Not::bool => bool)
+               ((op =::bool list list => bool list list => bool) l2
+                 ((op #::bool list => bool list list => bool list list)
+                   ([]::bool list) ([]::bool list list)))))
+           ((op =::bool => bool => bool)
+             ((algebra_canon::bool list list => bool)
+               ((op @::bool list list => bool list list => bool list list)
+                 ((map::(bool list => bool list)
+                        => bool list list => bool list list)
+                   ((op #::bool => bool list => bool list) (True::bool)) l1)
+                 ((map::(bool list => bool list)
+                        => bool list list => bool list list)
+                   ((op #::bool => bool list => bool list) (False::bool))
+                   l2)))
+             ((op &::bool => bool => bool)
+               ((algebra_canon::bool list list => bool) l1)
+               ((algebra_canon::bool list list => bool) l2)))))"
+  by (import prob_canon ALGEBRA_CANON_STEP)
+
+lemma ALGEBRA_CANON_CASES_THM: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((op |::bool => bool => bool)
+        ((op =::bool list list => bool list list => bool) l
+          ([]::bool list list))
+        ((op |::bool => bool => bool)
+          ((op =::bool list list => bool list list => bool) l
+            ((op #::bool list => bool list list => bool list list)
+              ([]::bool list) ([]::bool list list)))
+          ((Ex::(bool list list => bool) => bool)
+            (%l1::bool list list.
+                (Ex::(bool list list => bool) => bool)
+                 (%l2::bool list list.
+                     (op &::bool => bool => bool)
+                      ((algebra_canon::bool list list => bool) l1)
+                      ((op &::bool => bool => bool)
+                        ((algebra_canon::bool list list => bool) l2)
+                        ((op =::bool list list => bool list list => bool) l
+                          ((op @::bool list list
+                                  => bool list list => bool list list)
+                            ((map::(bool list => bool list)
+                                   => bool list list => bool list list)
+                              ((op #::bool => bool list => bool list)
+                                (True::bool))
+                              l1)
+                            ((map::(bool list => bool list)
+                                   => bool list list => bool list list)
+                              ((op #::bool => bool list => bool list)
+                                (False::bool))
+                              l2))))))))))"
+  by (import prob_canon ALGEBRA_CANON_CASES_THM)
+
+lemma ALGEBRA_CANON_CASES: "(All::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) (P ([]::bool list list))
+        ((op &::bool => bool => bool)
+          (P ((op #::bool list => bool list list => bool list list)
+               ([]::bool list) ([]::bool list list)))
+          ((All::(bool list list => bool) => bool)
+            (%l1::bool list list.
+                (All::(bool list list => bool) => bool)
+                 (%l2::bool list list.
+                     (op -->::bool => bool => bool)
+                      ((op &::bool => bool => bool)
+                        ((algebra_canon::bool list list => bool) l1)
+                        ((op &::bool => bool => bool)
+                          ((algebra_canon::bool list list => bool) l2)
+                          ((algebra_canon::bool list list => bool)
+                            ((op @::bool list list
+                                    => bool list list => bool list list)
+                              ((map::(bool list => bool list)
+                                     => bool list list => bool list list)
+                                ((op #::bool => bool list => bool list)
+                                  (True::bool))
+                                l1)
+                              ((map::(bool list => bool list)
+                                     => bool list list => bool list list)
+                                ((op #::bool => bool list => bool list)
+                                  (False::bool))
+                                l2)))))
+                      (P ((op @::bool list list
+                                 => bool list list => bool list list)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (True::bool))
+                             l1)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (False::bool))
+                             l2))))))))
+      ((All::(bool list list => bool) => bool)
+        (%l::bool list list.
+            (op -->::bool => bool => bool)
+             ((algebra_canon::bool list list => bool) l) (P l))))"
+  by (import prob_canon ALGEBRA_CANON_CASES)
+
+lemma ALGEBRA_CANON_INDUCTION: "(All::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) (P ([]::bool list list))
+        ((op &::bool => bool => bool)
+          (P ((op #::bool list => bool list list => bool list list)
+               ([]::bool list) ([]::bool list list)))
+          ((All::(bool list list => bool) => bool)
+            (%l1::bool list list.
+                (All::(bool list list => bool) => bool)
+                 (%l2::bool list list.
+                     (op -->::bool => bool => bool)
+                      ((op &::bool => bool => bool)
+                        ((algebra_canon::bool list list => bool) l1)
+                        ((op &::bool => bool => bool)
+                          ((algebra_canon::bool list list => bool) l2)
+                          ((op &::bool => bool => bool) (P l1)
+                            ((op &::bool => bool => bool) (P l2)
+                              ((algebra_canon::bool list list => bool)
+                                ((op @::bool list list
+  => bool list list => bool list list)
+                                  ((map::(bool list => bool list)
+   => bool list list => bool list list)
+                                    ((op #::bool => bool list => bool list)
+(True::bool))
+                                    l1)
+                                  ((map::(bool list => bool list)
+   => bool list list => bool list list)
+                                    ((op #::bool => bool list => bool list)
+(False::bool))
+                                    l2)))))))
+                      (P ((op @::bool list list
+                                 => bool list list => bool list list)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (True::bool))
+                             l1)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (False::bool))
+                             l2))))))))
+      ((All::(bool list list => bool) => bool)
+        (%l::bool list list.
+            (op -->::bool => bool => bool)
+             ((algebra_canon::bool list list => bool) l) (P l))))"
+  by (import prob_canon ALGEBRA_CANON_INDUCTION)
+
+lemma MEM_NIL_STEP: "ALL l1 l2. ~ [] mem map (op # True) l1 @ map (op # False) l2"
+  by (import prob_canon MEM_NIL_STEP)
+
+lemma ALG_SORTED_PREFIXFREE_MEM_NIL: "ALL l. (alg_sorted l & alg_prefixfree l & [] mem l) = (l = [[]])"
+  by (import prob_canon ALG_SORTED_PREFIXFREE_MEM_NIL)
+
+lemma ALG_SORTED_PREFIXFREE_EQUALITY: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l'::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((All::(bool list => bool) => bool)
+               (%x::bool list.
+                   (op =::bool => bool => bool)
+                    ((op mem::bool list => bool list list => bool) x l)
+                    ((op mem::bool list => bool list list => bool) x l')))
+             ((op &::bool => bool => bool)
+               ((alg_sorted::bool list list => bool) l)
+               ((op &::bool => bool => bool)
+                 ((alg_sorted::bool list list => bool) l')
+                 ((op &::bool => bool => bool)
+                   ((alg_prefixfree::bool list list => bool) l)
+                   ((alg_prefixfree::bool list list => bool) l')))))
+           ((op =::bool list list => bool list list => bool) l l')))"
+  by (import prob_canon ALG_SORTED_PREFIXFREE_EQUALITY)
+
+;end_setup
+
+;setup_theory boolean_sequence
+
+consts
+  SHD :: "(nat => bool) => bool" 
+
+defs
+  SHD_primdef: "SHD == %f. f 0"
+
+lemma SHD_def: "ALL f. SHD f = f 0"
+  by (import boolean_sequence SHD_def)
+
+consts
+  STL :: "(nat => bool) => nat => bool" 
+
+defs
+  STL_primdef: "STL == %f n. f (Suc n)"
+
+lemma STL_def: "ALL f n. 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 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. (SHD s, STL s)"
+
+lemma SDEST_def: "SDEST = (%s. (SHD s, STL s))"
+  by (import boolean_sequence SDEST_def)
+
+consts
+  SCONST :: "bool => nat => bool" 
+
+defs
+  SCONST_primdef: "SCONST == K"
+
+lemma SCONST_def: "SCONST = K"
+  by (import boolean_sequence SCONST_def)
+
+consts
+  STAKE :: "nat => (nat => bool) => bool list" 
+
+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. SDROP (Suc n) = SDROP n o STL)"
+  by (import boolean_sequence SDROP_def)
+
+lemma SCONS_SURJ: "ALL x. EX xa t. x = SCONS xa t"
+  by (import boolean_sequence SCONS_SURJ)
+
+lemma SHD_STL_ISO: "ALL h t. EX x. SHD x = h & STL x = t"
+  by (import boolean_sequence SHD_STL_ISO)
+
+lemma SHD_SCONS: "ALL h t. SHD (SCONS h t) = h"
+  by (import boolean_sequence SHD_SCONS)
+
+lemma STL_SCONS: "ALL h t. STL (SCONS h t) = t"
+  by (import boolean_sequence STL_SCONS)
+
+lemma SHD_SCONST: "ALL b. SHD (SCONST b) = b"
+  by (import boolean_sequence SHD_SCONST)
+
+lemma STL_SCONST: "ALL b. STL (SCONST b) = SCONST b"
+  by (import boolean_sequence STL_SCONST)
+
+;end_setup
+
+;setup_theory prob_algebra
+
+consts
+  alg_embed :: "bool list => (nat => bool) => bool" 
+
+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 t.
+    algebra_embed (h # t) = pred_set.UNION (alg_embed h) (algebra_embed t))"
+  by (import prob_algebra algebra_embed_def)
+
+consts
+  measurable :: "((nat => bool) => bool) => bool" 
+
+defs
+  measurable_primdef: "measurable == %s. EX b. s = algebra_embed b"
+
+lemma measurable_def: "ALL s. measurable s = (EX b. s = algebra_embed b)"
+  by (import prob_algebra measurable_def)
+
+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 q. 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. 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 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. All (alg_embed c) = (c = [])"
+  by (import prob_algebra ALG_EMBED_NIL)
+
+lemma ALG_EMBED_POPULATED: "ALL b. Ex (alg_embed b)"
+  by (import prob_algebra ALG_EMBED_POPULATED)
+
+lemma ALG_EMBED_PREFIX: "(All::(bool list => bool) => bool)
+ (%b::bool list.
+     (All::(bool list => bool) => bool)
+      (%c::bool list.
+          (All::((nat => bool) => bool) => bool)
+           (%s::nat => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_embed::bool list => (nat => bool) => bool) b s)
+                  ((alg_embed::bool list => (nat => bool) => bool) c s))
+                ((op |::bool => bool => bool)
+                  ((IS_PREFIX::bool list => bool list => bool) b c)
+                  ((IS_PREFIX::bool list => bool list => bool) c b)))))"
+  by (import prob_algebra ALG_EMBED_PREFIX)
+
+lemma ALG_EMBED_PREFIX_SUBSET: "ALL b c. 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.
+   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. algebra_embed [[b]] = (%s. SHD s = b))"
+  by (import prob_algebra ALGEBRA_EMBED_BASIC)
+
+lemma ALGEBRA_EMBED_MEM: "(All::(bool list list => bool) => bool)
+ (%b::bool list list.
+     (All::((nat => bool) => bool) => bool)
+      (%x::nat => bool.
+          (op -->::bool => bool => bool)
+           ((algebra_embed::bool list list => (nat => bool) => bool) b x)
+           ((Ex::(bool list => bool) => bool)
+             (%l::bool list.
+                 (op &::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) l b)
+                  ((alg_embed::bool list => (nat => bool) => bool) l x)))))"
+  by (import prob_algebra ALGEBRA_EMBED_MEM)
+
+lemma ALGEBRA_EMBED_APPEND: "ALL l1 l2.
+   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 b.
+   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 b. 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 b. 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. algebra_embed (alg_canon1 x) = algebra_embed x"
+  by (import prob_algebra ALG_CANON1_EMBED)
+
+lemma ALG_CANON_MERGE_EMBED: "ALL l b. 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. algebra_embed (alg_canon2 x) = algebra_embed x"
+  by (import prob_algebra ALG_CANON2_EMBED)
+
+lemma ALG_CANON_EMBED: "ALL l. algebra_embed (alg_canon l) = algebra_embed l"
+  by (import prob_algebra ALG_CANON_EMBED)
+
+lemma ALGEBRA_CANON_UNIV: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((op -->::bool => bool => bool)
+        ((op =::((nat => bool) => bool) => ((nat => bool) => bool) => bool)
+          ((algebra_embed::bool list list => (nat => bool) => bool) l)
+          (pred_set.UNIV::(nat => bool) => bool))
+        ((op =::bool list list => bool list list => bool) l
+          ((op #::bool list => bool list list => bool list list)
+            ([]::bool list) ([]::bool list list)))))"
+  by (import prob_algebra ALGEBRA_CANON_UNIV)
+
+lemma ALG_CANON_REP: "ALL b c. (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::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((op =::bool => bool => bool)
+        ((All::((nat => bool) => bool) => bool)
+          (%v::nat => bool.
+              (Not::bool => bool)
+               ((algebra_embed::bool list list => (nat => bool) => bool) l
+                 v)))
+        ((op =::bool list list => bool list list => bool) l
+          ([]::bool list list))))"
+  by (import prob_algebra ALGEBRA_CANON_EMBED_EMPTY)
+
+lemma ALGEBRA_CANON_EMBED_UNIV: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((op =::bool => bool => bool)
+        ((All::((nat => bool) => bool) => bool)
+          ((algebra_embed::bool list list => (nat => bool) => bool) l))
+        ((op =::bool list list => bool list list => bool) l
+          ((op #::bool list => bool list list => bool list list)
+            ([]::bool list) ([]::bool list list)))))"
+  by (import prob_algebra ALGEBRA_CANON_EMBED_UNIV)
+
+lemma MEASURABLE_ALGEBRA: "ALL b. measurable (algebra_embed b)"
+  by (import prob_algebra MEASURABLE_ALGEBRA)
+
+lemma MEASURABLE_BASIC: "measurable EMPTY &
+measurable pred_set.UNIV & (ALL b. measurable (%s. SHD s = b))"
+  by (import prob_algebra MEASURABLE_BASIC)
+
+lemma MEASURABLE_SHD: "ALL b. measurable (%s. SHD s = b)"
+  by (import prob_algebra MEASURABLE_SHD)
+
+lemma ALGEBRA_EMBED_COMPL: "ALL l. EX l'. COMPL (algebra_embed l) = algebra_embed l'"
+  by (import prob_algebra ALGEBRA_EMBED_COMPL)
+
+lemma MEASURABLE_COMPL: "ALL s. measurable (COMPL s) = measurable s"
+  by (import prob_algebra MEASURABLE_COMPL)
+
+lemma MEASURABLE_UNION: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%t::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool) s)
+             ((measurable::((nat => bool) => bool) => bool) t))
+           ((measurable::((nat => bool) => bool) => bool)
+             ((pred_set.UNION::((nat => bool) => bool)
+                               => ((nat => bool) => bool)
+                                  => (nat => bool) => bool)
+               s t))))"
+  by (import prob_algebra MEASURABLE_UNION)
+
+lemma MEASURABLE_INTER: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%t::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool) s)
+             ((measurable::((nat => bool) => bool) => bool) t))
+           ((measurable::((nat => bool) => bool) => bool)
+             ((pred_set.INTER::((nat => bool) => bool)
+                               => ((nat => bool) => bool)
+                                  => (nat => bool) => bool)
+               s t))))"
+  by (import prob_algebra MEASURABLE_INTER)
+
+lemma MEASURABLE_STL: "ALL p. measurable (p o STL) = measurable p"
+  by (import prob_algebra MEASURABLE_STL)
+
+lemma MEASURABLE_SDROP: "ALL n p. measurable (p o SDROP n) = measurable p"
+  by (import prob_algebra MEASURABLE_SDROP)
+
+lemma MEASURABLE_INTER_HALVES: "ALL p.
+   (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 q.
+   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 p.
+   measurable (pred_set.INTER (%x. SHD x = b) (p o STL)) = measurable p"
+  by (import prob_algebra MEASURABLE_INTER_SHD)
+
+;end_setup
+
+;setup_theory prob
+
+consts
+  alg_measure :: "bool list list => real" 
+
+specification (alg_measure_primdef: alg_measure) alg_measure_def: "alg_measure [] = 0 &
+(ALL l rest. alg_measure (l # rest) = (1 / 2) ^ length l + alg_measure rest)"
+  by (import prob alg_measure_def)
+
+consts
+  algebra_measure :: "bool list list => real" 
+
+defs
+  algebra_measure_primdef: "algebra_measure ==
+%b. inf (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)"
+
+lemma algebra_measure_def: "ALL b.
+   algebra_measure b =
+   inf (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)"
+  by (import prob algebra_measure_def)
+
+consts
+  prob :: "((nat => bool) => bool) => real" 
+
+defs
+  prob_primdef: "prob ==
+%s. sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)"
+
+lemma prob_def: "ALL s.
+   prob s =
+   sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)"
+  by (import prob prob_def)
+
+lemma ALG_TWINS_MEASURE: "ALL l::bool list.
+   ((1::real) / (2::real)) ^ length (SNOC True l) +
+   ((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. alg_measure [[b]] = 1 / 2)"
+  by (import prob ALG_MEASURE_BASIC)
+
+lemma ALG_MEASURE_POS: "ALL l. 0 <= alg_measure l"
+  by (import prob ALG_MEASURE_POS)
+
+lemma ALG_MEASURE_APPEND: "ALL l1 l2. alg_measure (l1 @ l2) = alg_measure l1 + alg_measure l2"
+  by (import prob ALG_MEASURE_APPEND)
+
+lemma ALG_MEASURE_TLS: "ALL l b. 2 * alg_measure (map (op # b) l) = alg_measure l"
+  by (import prob ALG_MEASURE_TLS)
+
+lemma ALG_CANON_PREFS_MONO: "ALL l b. 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 b. alg_measure (alg_canon_find l b) <= alg_measure (l # b)"
+  by (import prob ALG_CANON_FIND_MONO)
+
+lemma ALG_CANON1_MONO: "ALL x. alg_measure (alg_canon1 x) <= alg_measure x"
+  by (import prob ALG_CANON1_MONO)
+
+lemma ALG_CANON_MERGE_MONO: "ALL l b. alg_measure (alg_canon_merge l b) <= alg_measure (l # b)"
+  by (import prob ALG_CANON_MERGE_MONO)
+
+lemma ALG_CANON2_MONO: "ALL x. alg_measure (alg_canon2 x) <= alg_measure x"
+  by (import prob ALG_CANON2_MONO)
+
+lemma ALG_CANON_MONO: "ALL l. alg_measure (alg_canon l) <= alg_measure l"
+  by (import prob ALG_CANON_MONO)
+
+lemma ALGEBRA_MEASURE_DEF_ALT: "ALL l. 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. algebra_measure [[b]] = 1 / 2)"
+  by (import prob ALGEBRA_MEASURE_BASIC)
+
+lemma ALGEBRA_CANON_MEASURE_MAX: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((op <=::real => real => bool)
+        ((alg_measure::bool list list => real) l) (1::real)))"
+  by (import prob ALGEBRA_CANON_MEASURE_MAX)
+
+lemma ALGEBRA_MEASURE_MAX: "ALL l. algebra_measure l <= 1"
+  by (import prob ALGEBRA_MEASURE_MAX)
+
+lemma ALGEBRA_MEASURE_MONO_EMBED: "(All::(bool list list => bool) => bool)
+ (%x::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%xa::bool list list.
+          (op -->::bool => bool => bool)
+           ((SUBSET::((nat => bool) => bool)
+                     => ((nat => bool) => bool) => bool)
+             ((algebra_embed::bool list list => (nat => bool) => bool) x)
+             ((algebra_embed::bool list list => (nat => bool) => bool) xa))
+           ((op <=::real => real => bool)
+             ((algebra_measure::bool list list => real) x)
+             ((algebra_measure::bool list list => real) xa))))"
+  by (import prob ALGEBRA_MEASURE_MONO_EMBED)
+
+lemma ALG_MEASURE_COMPL: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((All::(bool list list => bool) => bool)
+        (%c::bool list list.
+            (op -->::bool => bool => bool)
+             ((algebra_canon::bool list list => bool) c)
+             ((op -->::bool => bool => bool)
+               ((op =::((nat => bool) => bool)
+                       => ((nat => bool) => bool) => bool)
+                 ((COMPL::((nat => bool) => bool) => (nat => bool) => bool)
+                   ((algebra_embed::bool list list => (nat => bool) => bool)
+                     l))
+                 ((algebra_embed::bool list list => (nat => bool) => bool)
+                   c))
+               ((op =::real => real => bool)
+                 ((op +::real => real => real)
+                   ((alg_measure::bool list list => real) l)
+                   ((alg_measure::bool list list => real) c))
+                 (1::real))))))"
+  by (import prob ALG_MEASURE_COMPL)
+
+lemma ALG_MEASURE_ADDITIVE: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((algebra_canon::bool list list => bool) l)
+      ((All::(bool list list => bool) => bool)
+        (%c::bool list list.
+            (op -->::bool => bool => bool)
+             ((algebra_canon::bool list list => bool) c)
+             ((All::(bool list list => bool) => bool)
+               (%d::bool list list.
+                   (op -->::bool => bool => bool)
+                    ((algebra_canon::bool list list => bool) d)
+                    ((op -->::bool => bool => bool)
+                      ((op &::bool => bool => bool)
+                        ((op =::((nat => bool) => bool)
+                                => ((nat => bool) => bool) => bool)
+                          ((pred_set.INTER::((nat => bool) => bool)
+      => ((nat => bool) => bool) => (nat => bool) => bool)
+                            ((algebra_embed::bool list list
+       => (nat => bool) => bool)
+                              c)
+                            ((algebra_embed::bool list list
+       => (nat => bool) => bool)
+                              d))
+                          (EMPTY::(nat => bool) => bool))
+                        ((op =::((nat => bool) => bool)
+                                => ((nat => bool) => bool) => bool)
+                          ((algebra_embed::bool list list
+     => (nat => bool) => bool)
+                            l)
+                          ((pred_set.UNION::((nat => bool) => bool)
+      => ((nat => bool) => bool) => (nat => bool) => bool)
+                            ((algebra_embed::bool list list
+       => (nat => bool) => bool)
+                              c)
+                            ((algebra_embed::bool list list
+       => (nat => bool) => bool)
+                              d))))
+                      ((op =::real => real => bool)
+                        ((alg_measure::bool list list => real) l)
+                        ((op +::real => real => real)
+                          ((alg_measure::bool list list => real) c)
+                          ((alg_measure::bool list list => real) d)))))))))"
+  by (import prob ALG_MEASURE_ADDITIVE)
+
+lemma PROB_ALGEBRA: "ALL l. 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. prob (%s. SHD s = b) = 1 / 2)"
+  by (import prob PROB_BASIC)
+
+lemma PROB_ADDITIVE: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%t::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool) s)
+             ((op &::bool => bool => bool)
+               ((measurable::((nat => bool) => bool) => bool) t)
+               ((op =::((nat => bool) => bool)
+                       => ((nat => bool) => bool) => bool)
+                 ((pred_set.INTER::((nat => bool) => bool)
+                                   => ((nat => bool) => bool)
+=> (nat => bool) => bool)
+                   s t)
+                 (EMPTY::(nat => bool) => bool))))
+           ((op =::real => real => bool)
+             ((prob::((nat => bool) => bool) => real)
+               ((pred_set.UNION::((nat => bool) => bool)
+                                 => ((nat => bool) => bool)
+                                    => (nat => bool) => bool)
+                 s t))
+             ((op +::real => real => real)
+               ((prob::((nat => bool) => bool) => real) s)
+               ((prob::((nat => bool) => bool) => real) t)))))"
+  by (import prob PROB_ADDITIVE)
+
+lemma PROB_COMPL: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((measurable::((nat => bool) => bool) => bool) s)
+      ((op =::real => real => bool)
+        ((prob::((nat => bool) => bool) => real)
+          ((COMPL::((nat => bool) => bool) => (nat => bool) => bool) s))
+        ((op -::real => real => real) (1::real)
+          ((prob::((nat => bool) => bool) => real) s))))"
+  by (import prob PROB_COMPL)
+
+lemma PROB_SUP_EXISTS1: "ALL s. EX x b. algebra_measure b = x & SUBSET (algebra_embed b) s"
+  by (import prob PROB_SUP_EXISTS1)
+
+lemma PROB_SUP_EXISTS2: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (Ex::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%r::real.
+               (op -->::bool => bool => bool)
+                ((Ex::(bool list list => bool) => bool)
+                  (%b::bool list list.
+                      (op &::bool => bool => bool)
+                       ((op =::real => real => bool)
+                         ((algebra_measure::bool list list => real) b) r)
+                       ((SUBSET::((nat => bool) => bool)
+                                 => ((nat => bool) => bool) => bool)
+                         ((algebra_embed::bool list list
+    => (nat => bool) => bool)
+                           b)
+                         s)))
+                ((op <=::real => real => bool) r x))))"
+  by (import prob PROB_SUP_EXISTS2)
+
+lemma PROB_LE_X: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((All::(((nat => bool) => bool) => bool) => bool)
+             (%s'::(nat => bool) => bool.
+                 (op -->::bool => bool => bool)
+                  ((op &::bool => bool => bool)
+                    ((measurable::((nat => bool) => bool) => bool) s')
+                    ((SUBSET::((nat => bool) => bool)
+                              => ((nat => bool) => bool) => bool)
+                      s' s))
+                  ((op <=::real => real => bool)
+                    ((prob::((nat => bool) => bool) => real) s') x)))
+           ((op <=::real => real => bool)
+             ((prob::((nat => bool) => bool) => real) s) x)))"
+  by (import prob PROB_LE_X)
+
+lemma X_LE_PROB: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((Ex::(((nat => bool) => bool) => bool) => bool)
+             (%s'::(nat => bool) => bool.
+                 (op &::bool => bool => bool)
+                  ((measurable::((nat => bool) => bool) => bool) s')
+                  ((op &::bool => bool => bool)
+                    ((SUBSET::((nat => bool) => bool)
+                              => ((nat => bool) => bool) => bool)
+                      s' s)
+                    ((op <=::real => real => bool) x
+                      ((prob::((nat => bool) => bool) => real) s')))))
+           ((op <=::real => real => bool) x
+             ((prob::((nat => bool) => bool) => real) s))))"
+  by (import prob X_LE_PROB)
+
+lemma PROB_SUBSET_MONO: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%s::(nat => bool) => bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%t::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((SUBSET::((nat => bool) => bool)
+                     => ((nat => bool) => bool) => bool)
+             s t)
+           ((op <=::real => real => bool)
+             ((prob::((nat => bool) => bool) => real) s)
+             ((prob::((nat => bool) => bool) => real) t))))"
+  by (import prob PROB_SUBSET_MONO)
+
+lemma PROB_ALG: "ALL x. prob (alg_embed x) = (1 / 2) ^ length x"
+  by (import prob PROB_ALG)
+
+lemma PROB_STL: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%p::(nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((measurable::((nat => bool) => bool) => bool) p)
+      ((op =::real => real => bool)
+        ((prob::((nat => bool) => bool) => real)
+          ((op o::((nat => bool) => bool)
+                  => ((nat => bool) => nat => bool)
+                     => (nat => bool) => bool)
+            p (STL::(nat => bool) => nat => bool)))
+        ((prob::((nat => bool) => bool) => real) p)))"
+  by (import prob PROB_STL)
+
+lemma PROB_SDROP: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%p::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((measurable::((nat => bool) => bool) => bool) p)
+           ((op =::real => real => bool)
+             ((prob::((nat => bool) => bool) => real)
+               ((op o::((nat => bool) => bool)
+                       => ((nat => bool) => nat => bool)
+                          => (nat => bool) => bool)
+                 p ((SDROP::nat => (nat => bool) => nat => bool) n)))
+             ((prob::((nat => bool) => bool) => real) p))))"
+  by (import prob PROB_SDROP)
+
+lemma PROB_INTER_HALVES: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%p::(nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((measurable::((nat => bool) => bool) => bool) p)
+      ((op =::real => real => bool)
+        ((op +::real => real => real)
+          ((prob::((nat => bool) => bool) => real)
+            ((pred_set.INTER::((nat => bool) => bool)
+                              => ((nat => bool) => bool)
+                                 => (nat => bool) => bool)
+              (%x::nat => bool.
+                  (op =::bool => bool => bool)
+                   ((SHD::(nat => bool) => bool) x) (True::bool))
+              p))
+          ((prob::((nat => bool) => bool) => real)
+            ((pred_set.INTER::((nat => bool) => bool)
+                              => ((nat => bool) => bool)
+                                 => (nat => bool) => bool)
+              (%x::nat => bool.
+                  (op =::bool => bool => bool)
+                   ((SHD::(nat => bool) => bool) x) (False::bool))
+              p)))
+        ((prob::((nat => bool) => bool) => real) p)))"
+  by (import prob PROB_INTER_HALVES)
+
+lemma PROB_INTER_SHD: "(All::(bool => bool) => bool)
+ (%b::bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%p::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((measurable::((nat => bool) => bool) => bool) p)
+           ((op =::real => real => bool)
+             ((prob::((nat => bool) => bool) => real)
+               ((pred_set.INTER::((nat => bool) => bool)
+                                 => ((nat => bool) => bool)
+                                    => (nat => bool) => bool)
+                 (%x::nat => bool.
+                     (op =::bool => bool => bool)
+                      ((SHD::(nat => bool) => bool) x) b)
+                 ((op o::((nat => bool) => bool)
+                         => ((nat => bool) => nat => bool)
+                            => (nat => bool) => bool)
+                   p (STL::(nat => bool) => nat => bool))))
+             ((op *::real => real => real)
+               ((op /::real => real => real) (1::real)
+                 ((number_of::bin => real)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool))))
+               ((prob::((nat => bool) => bool) => real) p)))))"
+  by (import prob PROB_INTER_SHD)
+
+lemma ALGEBRA_MEASURE_POS: "ALL l. 0 <= algebra_measure l"
+  by (import prob ALGEBRA_MEASURE_POS)
+
+lemma ALGEBRA_MEASURE_RANGE: "ALL l. 0 <= algebra_measure l & algebra_measure l <= 1"
+  by (import prob ALGEBRA_MEASURE_RANGE)
+
+lemma PROB_POS: "ALL p. 0 <= prob p"
+  by (import prob PROB_POS)
+
+lemma PROB_MAX: "ALL p. prob p <= 1"
+  by (import prob PROB_MAX)
+
+lemma PROB_RANGE: "ALL p. 0 <= prob p & prob p <= 1"
+  by (import prob PROB_RANGE)
+
+lemma ABS_PROB: "ALL p. abs (prob p) = prob p"
+  by (import prob ABS_PROB)
+
+lemma PROB_SHD: "ALL b. prob (%s. SHD s = b) = 1 / 2"
+  by (import prob PROB_SHD)
+
+lemma PROB_COMPL_LE1: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%p::(nat => bool) => bool.
+     (All::(real => bool) => bool)
+      (%r::real.
+          (op -->::bool => bool => bool)
+           ((measurable::((nat => bool) => bool) => bool) p)
+           ((op =::bool => bool => bool)
+             ((op <=::real => real => bool)
+               ((prob::((nat => bool) => bool) => real)
+                 ((COMPL::((nat => bool) => bool) => (nat => bool) => bool)
+                   p))
+               r)
+             ((op <=::real => real => bool)
+               ((op -::real => real => real) (1::real) r)
+               ((prob::((nat => bool) => bool) => real) p)))))"
+  by (import prob PROB_COMPL_LE1)
+
+;end_setup
+
+;setup_theory prob_pseudo
+
+consts
+  pseudo_linear_hd :: "nat => bool" 
+
+defs
+  pseudo_linear_hd_primdef: "pseudo_linear_hd == EVEN"
+
+lemma pseudo_linear_hd_def: "pseudo_linear_hd = EVEN"
+  by (import prob_pseudo pseudo_linear_hd_def)
+
+consts
+  pseudo_linear_tl :: "nat => nat => nat => nat => nat" 
+
+defs
+  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 b n x. 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. (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. SHD (pseudo_linear1 x) = pseudo_linear_hd x) &
+(ALL x.
+    STL (pseudo_linear1 x) =
+    pseudo_linear1
+     (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)))))))
+       x))"
+  by (import prob_pseudo pseudo_linear1_def)
+
+consts
+  pseudo :: "nat => nat => bool" 
+
+defs
+  pseudo_primdef: "pseudo == pseudo_linear1"
+
+lemma pseudo_def: "pseudo = pseudo_linear1"
+  by (import prob_pseudo pseudo_def)
+
+;end_setup
+
+;setup_theory prob_indep
+
+consts
+  indep_set :: "((nat => bool) => bool) => ((nat => bool) => bool) => bool" 
+
+defs
+  indep_set_primdef: "indep_set ==
+%p q. measurable p &
+      measurable q & prob (pred_set.INTER p q) = prob p * prob q"
+
+lemma indep_set_def: "ALL p q.
+   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
+  alg_cover_set :: "bool list list => bool" 
+
+defs
+  alg_cover_set_primdef: "alg_cover_set ==
+%l. alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV"
+
+lemma alg_cover_set_def: "ALL l.
+   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 x. SOME b. b mem l & alg_embed b x"
+
+lemma alg_cover_def: "ALL l x. alg_cover l x = (SOME b. b mem l & alg_embed b x)"
+  by (import prob_indep alg_cover_def)
+
+consts
+  indep :: "((nat => bool) => 'a * (nat => bool)) => bool" 
+
+defs
+  indep_primdef: "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)))"
+
+lemma indep_def: "ALL f.
+   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::(((nat => bool) => bool) => bool) => bool)
+ (%p::(nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((measurable::((nat => bool) => bool) => bool) p)
+      ((op &::bool => bool => bool)
+        ((indep_set::((nat => bool) => bool)
+                     => ((nat => bool) => bool) => bool)
+          (EMPTY::(nat => bool) => bool) p)
+        ((indep_set::((nat => bool) => bool)
+                     => ((nat => bool) => bool) => bool)
+          (pred_set.UNIV::(nat => bool) => bool) p)))"
+  by (import prob_indep INDEP_SET_BASIC)
+
+lemma INDEP_SET_SYM: "ALL p q. indep_set p q = indep_set p q"
+  by (import prob_indep INDEP_SET_SYM)
+
+lemma INDEP_SET_DISJOINT_DECOMP: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%p::(nat => bool) => bool.
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%q::(nat => bool) => bool.
+          (All::(((nat => bool) => bool) => bool) => bool)
+           (%r::(nat => bool) => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((indep_set::((nat => bool) => bool)
+                               => ((nat => bool) => bool) => bool)
+                    p r)
+                  ((op &::bool => bool => bool)
+                    ((indep_set::((nat => bool) => bool)
+                                 => ((nat => bool) => bool) => bool)
+                      q r)
+                    ((op =::((nat => bool) => bool)
+                            => ((nat => bool) => bool) => bool)
+                      ((pred_set.INTER::((nat => bool) => bool)
+  => ((nat => bool) => bool) => (nat => bool) => bool)
+                        p q)
+                      (EMPTY::(nat => bool) => bool))))
+                ((indep_set::((nat => bool) => bool)
+                             => ((nat => bool) => bool) => bool)
+                  ((pred_set.UNION::((nat => bool) => bool)
+                                    => ((nat => bool) => bool)
+ => (nat => bool) => bool)
+                    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::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (All::((nat => bool) => bool) => bool)
+      (%x::nat => bool.
+          (op -->::bool => bool => bool)
+           ((alg_cover_set::bool list list => bool) l)
+           ((op &::bool => bool => bool)
+             ((op mem::bool list => bool list list => bool)
+               ((alg_cover::bool list list => (nat => bool) => bool list) l
+                 x)
+               l)
+             ((alg_embed::bool list => (nat => bool) => bool)
+               ((alg_cover::bool list list => (nat => bool) => bool list) 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::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (All::(bool => bool) => bool)
+      (%b::bool.
+          (op -->::bool => bool => bool)
+           ((Not::bool => bool)
+             ((op mem::bool list => bool list list => bool) ([]::bool list)
+               l))
+           ((op =::bool list list => bool list list => bool)
+             ((map::(bool list => bool list)
+                    => bool list list => bool list list)
+               ((op #::bool => bool list => bool list) b)
+               ((map::(bool list => bool list)
+                      => bool list list => bool list list)
+                 (tl::bool list => bool list)
+                 ((filter::(bool list => bool)
+                           => bool list list => bool list list)
+                   (%x::bool list.
+                       (op =::bool => bool => bool)
+                        ((hd::bool list => bool) x) b)
+                   l)))
+             ((filter::(bool list => bool)
+                       => bool list list => bool list list)
+               (%x::bool list.
+                   (op =::bool => bool => bool) ((hd::bool list => bool) x)
+                    b)
+               l))))"
+  by (import prob_indep MAP_CONS_TL_FILTER)
+
+lemma ALG_COVER_SET_CASES_THM: "ALL l.
+   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::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        (P ((op #::bool list => bool list list => bool list list)
+             ([]::bool list) ([]::bool list list)))
+        ((All::(bool list list => bool) => bool)
+          (%l1::bool list list.
+              (All::(bool list list => bool) => bool)
+               (%l2::bool list list.
+                   (op -->::bool => bool => bool)
+                    ((op &::bool => bool => bool)
+                      ((alg_cover_set::bool list list => bool) l1)
+                      ((op &::bool => bool => bool)
+                        ((alg_cover_set::bool list list => bool) l2)
+                        ((alg_cover_set::bool list list => bool)
+                          ((op @::bool list list
+                                  => bool list list => bool list list)
+                            ((map::(bool list => bool list)
+                                   => bool list list => bool list list)
+                              ((op #::bool => bool list => bool list)
+                                (True::bool))
+                              l1)
+                            ((map::(bool list => bool list)
+                                   => bool list list => bool list list)
+                              ((op #::bool => bool list => bool list)
+                                (False::bool))
+                              l2)))))
+                    (P ((op @::bool list list
+                               => bool list list => bool list list)
+                         ((map::(bool list => bool list)
+                                => bool list list => bool list list)
+                           ((op #::bool => bool list => bool list)
+                             (True::bool))
+                           l1)
+                         ((map::(bool list => bool list)
+                                => bool list list => bool list list)
+                           ((op #::bool => bool list => bool list)
+                             (False::bool))
+                           l2)))))))
+      ((All::(bool list list => bool) => bool)
+        (%l::bool list list.
+            (op -->::bool => bool => bool)
+             ((alg_cover_set::bool list list => bool) l) (P l))))"
+  by (import prob_indep ALG_COVER_SET_CASES)
+
+lemma ALG_COVER_SET_INDUCTION: "(All::((bool list list => bool) => bool) => bool)
+ (%P::bool list list => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        (P ((op #::bool list => bool list list => bool list list)
+             ([]::bool list) ([]::bool list list)))
+        ((All::(bool list list => bool) => bool)
+          (%l1::bool list list.
+              (All::(bool list list => bool) => bool)
+               (%l2::bool list list.
+                   (op -->::bool => bool => bool)
+                    ((op &::bool => bool => bool)
+                      ((alg_cover_set::bool list list => bool) l1)
+                      ((op &::bool => bool => bool)
+                        ((alg_cover_set::bool list list => bool) l2)
+                        ((op &::bool => bool => bool) (P l1)
+                          ((op &::bool => bool => bool) (P l2)
+                            ((alg_cover_set::bool list list => bool)
+                              ((op @::bool list list
+=> bool list list => bool list list)
+                                ((map::(bool list => bool list)
+ => bool list list => bool list list)
+                                  ((op #::bool => bool list => bool list)
+                                    (True::bool))
+                                  l1)
+                                ((map::(bool list => bool list)
+ => bool list list => bool list list)
+                                  ((op #::bool => bool list => bool list)
+                                    (False::bool))
+                                  l2)))))))
+                    (P ((op @::bool list list
+                               => bool list list => bool list list)
+                         ((map::(bool list => bool list)
+                                => bool list list => bool list list)
+                           ((op #::bool => bool list => bool list)
+                             (True::bool))
+                           l1)
+                         ((map::(bool list => bool list)
+                                => bool list list => bool list list)
+                           ((op #::bool => bool list => bool list)
+                             (False::bool))
+                           l2)))))))
+      ((All::(bool list list => bool) => bool)
+        (%l::bool list list.
+            (op -->::bool => bool => bool)
+             ((alg_cover_set::bool list list => bool) l) (P l))))"
+  by (import prob_indep ALG_COVER_SET_INDUCTION)
+
+lemma ALG_COVER_EXISTS_UNIQUE: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_cover_set::bool list list => bool) l)
+      ((All::((nat => bool) => bool) => bool)
+        (%s::nat => bool.
+            (Ex1::(bool list => bool) => bool)
+             (%x::bool list.
+                 (op &::bool => bool => bool)
+                  ((op mem::bool list => bool list list => bool) x l)
+                  ((alg_embed::bool list => (nat => bool) => bool) x s)))))"
+  by (import prob_indep ALG_COVER_EXISTS_UNIQUE)
+
+lemma ALG_COVER_UNIQUE: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (All::(bool list => bool) => bool)
+      (%x::bool list.
+          (All::((nat => bool) => bool) => bool)
+           (%s::nat => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_cover_set::bool list list => bool) l)
+                  ((op &::bool => bool => bool)
+                    ((op mem::bool list => bool list list => bool) x l)
+                    ((alg_embed::bool list => (nat => bool) => bool) x s)))
+                ((op =::bool list => bool list => bool)
+                  ((alg_cover::bool list list => (nat => bool) => bool list)
+                    l s)
+                  x))))"
+  by (import prob_indep ALG_COVER_UNIQUE)
+
+lemma ALG_COVER_STEP: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (All::(bool => bool) => bool)
+           (%h::bool.
+               (All::((nat => bool) => bool) => bool)
+                (%t::nat => bool.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((alg_cover_set::bool list list => bool) l1)
+                       ((alg_cover_set::bool list list => bool) l2))
+                     ((op =::bool list => bool list => bool)
+                       ((alg_cover::bool list list
+                                    => (nat => bool) => bool list)
+                         ((op @::bool list list
+                                 => bool list list => bool list list)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (True::bool))
+                             l1)
+                           ((map::(bool list => bool list)
+                                  => bool list list => bool list list)
+                             ((op #::bool => bool list => bool list)
+                               (False::bool))
+                             l2))
+                         ((SCONS::bool => (nat => bool) => nat => bool) h
+                           t))
+                       ((If::bool => bool list => bool list => bool list) h
+                         ((op #::bool => bool list => bool list)
+                           (True::bool)
+                           ((alg_cover::bool list list
+  => (nat => bool) => bool list)
+                             l1 t))
+                         ((op #::bool => bool list => bool list)
+                           (False::bool)
+                           ((alg_cover::bool list list
+  => (nat => bool) => bool list)
+                             l2 t))))))))"
+  by (import prob_indep ALG_COVER_STEP)
+
+lemma ALG_COVER_HEAD: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_cover_set::bool list list => bool) l)
+      ((All::((bool list => bool) => bool) => bool)
+        (%f::bool list => bool.
+            (op =::((nat => bool) => bool)
+                   => ((nat => bool) => bool) => bool)
+             ((op o::(bool list => bool)
+                     => ((nat => bool) => bool list)
+                        => (nat => bool) => bool)
+               f ((alg_cover::bool list list => (nat => bool) => bool list)
+                   l))
+             ((algebra_embed::bool list list => (nat => bool) => bool)
+               ((filter::(bool list => bool)
+                         => bool list list => bool list list)
+                 f l)))))"
+  by (import prob_indep ALG_COVER_HEAD)
+
+lemma ALG_COVER_TAIL_STEP: "(All::(bool list list => bool) => bool)
+ (%l1::bool list list.
+     (All::(bool list list => bool) => bool)
+      (%l2::bool list list.
+          (All::(((nat => bool) => bool) => bool) => bool)
+           (%q::(nat => bool) => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((alg_cover_set::bool list list => bool) l1)
+                  ((alg_cover_set::bool list list => bool) l2))
+                ((op =::((nat => bool) => bool)
+                        => ((nat => bool) => bool) => bool)
+                  ((op o::((nat => bool) => bool)
+                          => ((nat => bool) => nat => bool)
+                             => (nat => bool) => bool)
+                    q (%x::nat => bool.
+                          (SDROP::nat => (nat => bool) => nat => bool)
+                           ((size::bool list => nat)
+                             ((alg_cover::bool list list
+    => (nat => bool) => bool list)
+                               ((op @::bool list list
+ => bool list list => bool list list)
+                                 ((map::(bool list => bool list)
+  => bool list list => bool list list)
+                                   ((op #::bool => bool list => bool list)
+                                     (True::bool))
+                                   l1)
+                                 ((map::(bool list => bool list)
+  => bool list list => bool list list)
+                                   ((op #::bool => bool list => bool list)
+                                     (False::bool))
+                                   l2))
+                               x))
+                           x))
+                  ((pred_set.UNION::((nat => bool) => bool)
+                                    => ((nat => bool) => bool)
+ => (nat => bool) => bool)
+                    ((pred_set.INTER::((nat => bool) => bool)
+=> ((nat => bool) => bool) => (nat => bool) => bool)
+                      (%x::nat => bool.
+                          (op =::bool => bool => bool)
+                           ((SHD::(nat => bool) => bool) x) (True::bool))
+                      ((op o::((nat => bool) => bool)
+                              => ((nat => bool) => nat => bool)
+                                 => (nat => bool) => bool)
+                        q ((op o::((nat => bool) => nat => bool)
+                                  => ((nat => bool) => nat => bool)
+                                     => (nat => bool) => nat => bool)
+                            (%x::nat => bool.
+                                (SDROP::nat => (nat => bool) => nat => bool)
+                                 ((size::bool list => nat)
+                                   ((alg_cover::bool list list
+          => (nat => bool) => bool list)
+                                     l1 x))
+                                 x)
+                            (STL::(nat => bool) => nat => bool))))
+                    ((pred_set.INTER::((nat => bool) => bool)
+=> ((nat => bool) => bool) => (nat => bool) => bool)
+                      (%x::nat => bool.
+                          (op =::bool => bool => bool)
+                           ((SHD::(nat => bool) => bool) x) (False::bool))
+                      ((op o::((nat => bool) => bool)
+                              => ((nat => bool) => nat => bool)
+                                 => (nat => bool) => bool)
+                        q ((op o::((nat => bool) => nat => bool)
+                                  => ((nat => bool) => nat => bool)
+                                     => (nat => bool) => nat => bool)
+                            (%x::nat => bool.
+                                (SDROP::nat => (nat => bool) => nat => bool)
+                                 ((size::bool list => nat)
+                                   ((alg_cover::bool list list
+          => (nat => bool) => bool list)
+                                     l2 x))
+                                 x)
+                            (STL::(nat => bool) => nat => bool)))))))))"
+  by (import prob_indep ALG_COVER_TAIL_STEP)
+
+lemma ALG_COVER_TAIL_MEASURABLE: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_cover_set::bool list list => bool) l)
+      ((All::(((nat => bool) => bool) => bool) => bool)
+        (%q::(nat => bool) => bool.
+            (op =::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool)
+               ((op o::((nat => bool) => bool)
+                       => ((nat => bool) => nat => bool)
+                          => (nat => bool) => bool)
+                 q (%x::nat => bool.
+                       (SDROP::nat => (nat => bool) => nat => bool)
+                        ((size::bool list => nat)
+                          ((alg_cover::bool list list
+ => (nat => bool) => bool list)
+                            l x))
+                        x)))
+             ((measurable::((nat => bool) => bool) => bool) q))))"
+  by (import prob_indep ALG_COVER_TAIL_MEASURABLE)
+
+lemma ALG_COVER_TAIL_PROB: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_cover_set::bool list list => bool) l)
+      ((All::(((nat => bool) => bool) => bool) => bool)
+        (%q::(nat => bool) => bool.
+            (op -->::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool) q)
+             ((op =::real => real => bool)
+               ((prob::((nat => bool) => bool) => real)
+                 ((op o::((nat => bool) => bool)
+                         => ((nat => bool) => nat => bool)
+                            => (nat => bool) => bool)
+                   q (%x::nat => bool.
+                         (SDROP::nat => (nat => bool) => nat => bool)
+                          ((size::bool list => nat)
+                            ((alg_cover::bool list list
+   => (nat => bool) => bool list)
+                              l x))
+                          x)))
+               ((prob::((nat => bool) => bool) => real) q)))))"
+  by (import prob_indep ALG_COVER_TAIL_PROB)
+
+lemma INDEP_INDEP_SET_LEMMA: "(All::(bool list list => bool) => bool)
+ (%l::bool list list.
+     (op -->::bool => bool => bool)
+      ((alg_cover_set::bool list list => bool) l)
+      ((All::(((nat => bool) => bool) => bool) => bool)
+        (%q::(nat => bool) => bool.
+            (op -->::bool => bool => bool)
+             ((measurable::((nat => bool) => bool) => bool) q)
+             ((All::(bool list => bool) => bool)
+               (%x::bool list.
+                   (op -->::bool => bool => bool)
+                    ((op mem::bool list => bool list list => bool) x l)
+                    ((op =::real => real => bool)
+                      ((prob::((nat => bool) => bool) => real)
+                        ((pred_set.INTER::((nat => bool) => bool)
+    => ((nat => bool) => bool) => (nat => bool) => bool)
+                          ((alg_embed::bool list => (nat => bool) => bool)
+                            x)
+                          ((op o::((nat => bool) => bool)
+                                  => ((nat => bool) => nat => bool)
+                                     => (nat => bool) => bool)
+                            q (%x::nat => bool.
+                                  (SDROP::nat
+    => (nat => bool) => nat => bool)
+                                   ((size::bool list => nat)
+                                     ((alg_cover::bool list list
+            => (nat => bool) => bool list)
+ l x))
+                                   x))))
+                      ((op *::real => real => real)
+                        ((op ^::real => nat => real)
+                          ((op /::real => real => real) (1::real)
+                            ((number_of::bin => real)
+                              ((op BIT::bin => bool => bin)
+                                ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                                  (True::bool))
+                                (False::bool))))
+                          ((size::bool list => nat) x))
+                        ((prob::((nat => bool) => bool) => real) q))))))))"
+  by (import prob_indep INDEP_INDEP_SET_LEMMA)
+
+lemma INDEP_SET_LIST: "(All::(((nat => bool) => bool) => bool) => bool)
+ (%q::(nat => bool) => bool.
+     (All::(bool list list => bool) => bool)
+      (%l::bool list list.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((alg_sorted::bool list list => bool) l)
+             ((op &::bool => bool => bool)
+               ((alg_prefixfree::bool list list => bool) l)
+               ((op &::bool => bool => bool)
+                 ((measurable::((nat => bool) => bool) => bool) q)
+                 ((All::(bool list => bool) => bool)
+                   (%x::bool list.
+                       (op -->::bool => bool => bool)
+                        ((op mem::bool list => bool list list => bool) x l)
+                        ((indep_set::((nat => bool) => bool)
+                                     => ((nat => bool) => bool) => bool)
+                          ((alg_embed::bool list => (nat => bool) => bool)
+                            x)
+                          q))))))
+           ((indep_set::((nat => bool) => bool)
+                        => ((nat => bool) => bool) => bool)
+             ((algebra_embed::bool list list => (nat => bool) => bool) l)
+             q)))"
+  by (import prob_indep INDEP_SET_LIST)
+
+lemma INDEP_INDEP_SET: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => 'a * (nat => bool).
+     (All::(('a => bool) => bool) => bool)
+      (%p::'a => bool.
+          (All::(((nat => bool) => bool) => bool) => bool)
+           (%q::(nat => bool) => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f)
+                  ((measurable::((nat => bool) => bool) => bool) q))
+                ((indep_set::((nat => bool) => bool)
+                             => ((nat => bool) => bool) => bool)
+                  ((op o::('a => bool)
+                          => ((nat => bool) => 'a) => (nat => bool) => bool)
+                    p ((op o::('a * (nat => bool) => 'a)
+                              => ((nat => bool) => 'a * (nat => bool))
+                                 => (nat => bool) => 'a)
+                        (fst::'a * (nat => bool) => 'a) f))
+                  ((op o::((nat => bool) => bool)
+                          => ((nat => bool) => nat => bool)
+                             => (nat => bool) => bool)
+                    q ((op o::('a * (nat => bool) => nat => bool)
+                              => ((nat => bool) => 'a * (nat => bool))
+                                 => (nat => bool) => nat => bool)
+                        (snd::'a * (nat => bool) => nat => bool) f))))))"
+  by (import prob_indep INDEP_INDEP_SET)
+
+lemma INDEP_UNIT: "ALL x. 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. BIND SDEST (%k. f o SCONS k) = f"
+  by (import prob_indep BIND_STEP)
+
+lemma INDEP_BIND_SDEST: "(All::((bool => (nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::bool => (nat => bool) => 'a * (nat => bool).
+     (op -->::bool => bool => bool)
+      ((All::(bool => bool) => bool)
+        (%x::bool.
+            (indep::((nat => bool) => 'a * (nat => bool)) => bool) (f x)))
+      ((indep::((nat => bool) => 'a * (nat => bool)) => bool)
+        ((BIND::((nat => bool) => bool * (nat => bool))
+                => (bool => (nat => bool) => 'a * (nat => bool))
+                   => (nat => bool) => 'a * (nat => bool))
+          (SDEST::(nat => bool) => bool * (nat => bool)) f)))"
+  by (import prob_indep INDEP_BIND_SDEST)
+
+lemma INDEP_BIND: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => 'a * (nat => bool).
+     (All::(('a => (nat => bool) => 'b * (nat => bool)) => bool) => bool)
+      (%g::'a => (nat => bool) => 'b * (nat => bool).
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f)
+             ((All::('a => bool) => bool)
+               (%x::'a.
+                   (indep::((nat => bool) => 'b * (nat => bool)) => bool)
+                    (g x))))
+           ((indep::((nat => bool) => 'b * (nat => bool)) => bool)
+             ((BIND::((nat => bool) => 'a * (nat => bool))
+                     => ('a => (nat => bool) => 'b * (nat => bool))
+                        => (nat => bool) => 'b * (nat => bool))
+               f g))))"
+  by (import prob_indep INDEP_BIND)
+
+lemma INDEP_PROB: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => 'a * (nat => bool).
+     (All::(('a => bool) => bool) => bool)
+      (%p::'a => bool.
+          (All::(((nat => bool) => bool) => bool) => bool)
+           (%q::(nat => bool) => bool.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f)
+                  ((measurable::((nat => bool) => bool) => bool) q))
+                ((op =::real => real => bool)
+                  ((prob::((nat => bool) => bool) => real)
+                    ((pred_set.INTER::((nat => bool) => bool)
+=> ((nat => bool) => bool) => (nat => bool) => bool)
+                      ((op o::('a => bool)
+                              => ((nat => bool) => 'a)
+                                 => (nat => bool) => bool)
+                        p ((op o::('a * (nat => bool) => 'a)
+                                  => ((nat => bool) => 'a * (nat => bool))
+                                     => (nat => bool) => 'a)
+                            (fst::'a * (nat => bool) => 'a) f))
+                      ((op o::((nat => bool) => bool)
+                              => ((nat => bool) => nat => bool)
+                                 => (nat => bool) => bool)
+                        q ((op o::('a * (nat => bool) => nat => bool)
+                                  => ((nat => bool) => 'a * (nat => bool))
+                                     => (nat => bool) => nat => bool)
+                            (snd::'a * (nat => bool) => nat => bool) f))))
+                  ((op *::real => real => real)
+                    ((prob::((nat => bool) => bool) => real)
+                      ((op o::('a => bool)
+                              => ((nat => bool) => 'a)
+                                 => (nat => bool) => bool)
+                        p ((op o::('a * (nat => bool) => 'a)
+                                  => ((nat => bool) => 'a * (nat => bool))
+                                     => (nat => bool) => 'a)
+                            (fst::'a * (nat => bool) => 'a) f)))
+                    ((prob::((nat => bool) => bool) => real) q))))))"
+  by (import prob_indep INDEP_PROB)
+
+lemma INDEP_MEASURABLE1: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => 'a * (nat => bool).
+     (All::(('a => bool) => bool) => bool)
+      (%p::'a => bool.
+          (op -->::bool => bool => bool)
+           ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f)
+           ((measurable::((nat => bool) => bool) => bool)
+             ((op o::('a => bool)
+                     => ((nat => bool) => 'a) => (nat => bool) => bool)
+               p ((op o::('a * (nat => bool) => 'a)
+                         => ((nat => bool) => 'a * (nat => bool))
+                            => (nat => bool) => 'a)
+                   (fst::'a * (nat => bool) => 'a) f)))))"
+  by (import prob_indep INDEP_MEASURABLE1)
+
+lemma INDEP_MEASURABLE2: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => 'a * (nat => bool).
+     (All::(((nat => bool) => bool) => bool) => bool)
+      (%q::(nat => bool) => bool.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f)
+             ((measurable::((nat => bool) => bool) => bool) q))
+           ((measurable::((nat => bool) => bool) => bool)
+             ((op o::((nat => bool) => bool)
+                     => ((nat => bool) => nat => bool)
+                        => (nat => bool) => bool)
+               q ((op o::('a * (nat => bool) => nat => bool)
+                         => ((nat => bool) => 'a * (nat => bool))
+                            => (nat => bool) => nat => bool)
+                   (snd::'a * (nat => bool) => nat => bool) f)))))"
+  by (import prob_indep INDEP_MEASURABLE2)
+
+lemma PROB_INDEP_BOUND: "(All::(((nat => bool) => nat * (nat => bool)) => bool) => bool)
+ (%f::(nat => bool) => nat * (nat => bool).
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((indep::((nat => bool) => nat * (nat => bool)) => bool) f)
+           ((op =::real => real => bool)
+             ((prob::((nat => bool) => bool) => real)
+               (%s::nat => bool.
+                   (op <::nat => nat => bool)
+                    ((fst::nat * (nat => bool) => nat) (f s))
+                    ((Suc::nat => nat) n)))
+             ((op +::real => real => real)
+               ((prob::((nat => bool) => bool) => real)
+                 (%s::nat => bool.
+                     (op <::nat => nat => bool)
+                      ((fst::nat * (nat => bool) => nat) (f s)) n))
+               ((prob::((nat => bool) => bool) => real)
+                 (%s::nat => bool.
+                     (op =::nat => nat => bool)
+                      ((fst::nat * (nat => bool) => nat) (f s)) n))))))"
+  by (import prob_indep PROB_INDEP_BOUND)
+
+;end_setup
+
+;setup_theory prob_uniform
+
+consts
+  unif_bound :: "nat => nat" 
+
+defs
+  unif_bound_primdef: "unif_bound ==
+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. 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) = Suc (unif_bound (Suc v div 2))"
+  by (import prob_uniform unif_bound_def)
+
+lemma unif_bound_ind: "(All::((nat => bool) => bool) => bool)
+ (%P::nat => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) (P (0::nat))
+        ((All::(nat => bool) => bool)
+          (%v::nat.
+              (op -->::bool => bool => bool)
+               (P ((op div::nat => nat => nat) ((Suc::nat => nat) v)
+                    ((number_of::bin => nat)
+                      ((op BIT::bin => bool => bin)
+                        ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                          (True::bool))
+                        (False::bool)))))
+               (P ((Suc::nat => nat) v)))))
+      ((All::(nat => bool) => bool) P))"
+  by (import prob_uniform unif_bound_ind)
+
+constdefs
+  unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" 
+  "unif_tupled ==
+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 =>
+         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. 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 =>
+         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)
+
+consts
+  unif :: "nat => (nat => bool) => nat * (nat => bool)" 
+
+defs
+  unif_primdef: "unif == %x x1. unif_tupled (x, x1)"
+
+lemma unif_curried_def: "ALL x x1. unif x x1 = unif_tupled (x, x1)"
+  by (import prob_uniform unif_curried_def)
+
+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::((nat => (nat => bool) => bool) => bool) => bool)
+ (%P::nat => (nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::((nat => bool) => bool) => bool) (P (0::nat)))
+        ((All::(nat => bool) => bool)
+          (%v2::nat.
+              (All::((nat => bool) => bool) => bool)
+               (%s::nat => bool.
+                   (op -->::bool => bool => bool)
+                    (P ((op div::nat => nat => nat) ((Suc::nat => nat) v2)
+                         ((number_of::bin => nat)
+                           ((op BIT::bin => bool => bin)
+                             ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                               (True::bool))
+                             (False::bool))))
+                      s)
+                    (P ((Suc::nat => nat) v2) s)))))
+      ((All::(nat => bool) => bool)
+        (%v::nat. (All::((nat => bool) => bool) => bool) (P v))))"
+  by (import prob_uniform unif_ind)
+
+constdefs
+  uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" 
+  "(op ==::(nat * nat * (nat => bool) => nat * (nat => bool))
+        => (nat * nat * (nat => bool) => nat * (nat => bool)) => prop)
+ (uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool))
+ ((WFREC::(nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+          => ((nat * nat * (nat => bool) => nat * (nat => bool))
+              => nat * nat * (nat => bool) => nat * (nat => bool))
+             => nat * nat * (nat => bool) => nat * (nat => bool))
+   ((Eps::((nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+           => bool)
+          => nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+     (%R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool.
+         (op &::bool => bool => bool)
+          ((WF::(nat * nat * (nat => bool)
+                 => nat * nat * (nat => bool) => bool)
+                => bool)
+            R)
+          ((All::(nat => bool) => bool)
+            (%t::nat.
+                (All::((nat => bool) => bool) => bool)
+                 (%s::nat => bool.
+                     (All::(nat => bool) => bool)
+                      (%n::nat.
+                          (All::(nat => bool) => bool)
+                           (%res::nat.
+                               (All::((nat => bool) => bool) => bool)
+                                (%s'::nat => bool.
+                                    (op -->::bool => bool => bool)
+                                     ((op &::bool => bool => bool)
+ ((op =::nat * (nat => bool) => nat * (nat => bool) => bool)
+   ((Pair::nat => (nat => bool) => nat * (nat => bool)) res s')
+   ((unif::nat => (nat => bool) => nat * (nat => bool)) n s))
+ ((Not::bool => bool)
+   ((op <::nat => nat => bool) res ((Suc::nat => nat) n))))
+                                     (R
+ ((Pair::nat => nat * (nat => bool) => nat * nat * (nat => bool)) t
+   ((Pair::nat => (nat => bool) => nat * (nat => bool))
+     ((Suc::nat => nat) n) s'))
+ ((Pair::nat => nat * (nat => bool) => nat * nat * (nat => bool))
+   ((Suc::nat => nat) t)
+   ((Pair::nat => (nat => bool) => nat * (nat => bool))
+     ((Suc::nat => nat) n) s)))))))))))
+   (%uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool).
+       (split::(nat => nat * (nat => bool) => nat * (nat => bool))
+               => nat * nat * (nat => bool) => nat * (nat => bool))
+        (%(v::nat) v1::nat * (nat => bool).
+            (nat_case::nat * (nat => bool)
+                       => (nat => nat * (nat => bool))
+                          => nat => nat * (nat => bool))
+             ((split::(nat => (nat => bool) => nat * (nat => bool))
+                      => nat * (nat => bool) => nat * (nat => bool))
+               (%(v3::nat) v4::nat => bool.
+                   (nat_case::nat * (nat => bool)
+                              => (nat => nat * (nat => bool))
+                                 => nat => nat * (nat => bool))
+                    (ARB::nat * (nat => bool))
+                    (%v5::nat.
+                        (Pair::nat => (nat => bool) => nat * (nat => bool))
+                         (0::nat) v4)
+                    v3)
+               v1)
+             (%v2::nat.
+                 (split::(nat => (nat => bool) => nat * (nat => bool))
+                         => nat * (nat => bool) => nat * (nat => bool))
+                  (%(v7::nat) v8::nat => bool.
+                      (nat_case::nat * (nat => bool)
+                                 => (nat => nat * (nat => bool))
+                                    => nat => nat * (nat => bool))
+                       (ARB::nat * (nat => bool))
+                       (%v9::nat.
+                           (Let::nat * (nat => bool)
+                                 => (nat * (nat => bool)
+                                     => nat * (nat => bool))
+                                    => nat * (nat => bool))
+                            ((unif::nat
+                                    => (nat => bool) => nat * (nat => bool))
+                              v9 v8)
+                            ((split::(nat
+=> (nat => bool) => nat * (nat => bool))
+                                     => nat * (nat => bool)
+  => nat * (nat => bool))
+                              (%(res::nat) s'::nat => bool.
+                                  (If::bool
+ => nat * (nat => bool) => nat * (nat => bool) => nat * (nat => bool))
+                                   ((op <::nat => nat => bool) res
+                                     ((Suc::nat => nat) v9))
+                                   ((Pair::nat
+     => (nat => bool) => nat * (nat => bool))
+                                     res s')
+                                   (uniform_tupled
+                                     ((Pair::nat
+       => nat * (nat => bool) => nat * nat * (nat => bool))
+ v2 ((Pair::nat => (nat => bool) => nat * (nat => bool))
+      ((Suc::nat => nat) v9) s'))))))
+                       v7)
+                  v1)
+             v)))"
+
+lemma uniform_tupled_primitive_def: "(op =::(nat * nat * (nat => bool) => nat * (nat => bool))
+       => (nat * nat * (nat => bool) => nat * (nat => bool)) => bool)
+ (uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool))
+ ((WFREC::(nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+          => ((nat * nat * (nat => bool) => nat * (nat => bool))
+              => nat * nat * (nat => bool) => nat * (nat => bool))
+             => nat * nat * (nat => bool) => nat * (nat => bool))
+   ((Eps::((nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+           => bool)
+          => nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool)
+     (%R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool.
+         (op &::bool => bool => bool)
+          ((WF::(nat * nat * (nat => bool)
+                 => nat * nat * (nat => bool) => bool)
+                => bool)
+            R)
+          ((All::(nat => bool) => bool)
+            (%t::nat.
+                (All::((nat => bool) => bool) => bool)
+                 (%s::nat => bool.
+                     (All::(nat => bool) => bool)
+                      (%n::nat.
+                          (All::(nat => bool) => bool)
+                           (%res::nat.
+                               (All::((nat => bool) => bool) => bool)
+                                (%s'::nat => bool.
+                                    (op -->::bool => bool => bool)
+                                     ((op &::bool => bool => bool)
+ ((op =::nat * (nat => bool) => nat * (nat => bool) => bool)
+   ((Pair::nat => (nat => bool) => nat * (nat => bool)) res s')
+   ((unif::nat => (nat => bool) => nat * (nat => bool)) n s))
+ ((Not::bool => bool)
+   ((op <::nat => nat => bool) res ((Suc::nat => nat) n))))
+                                     (R
+ ((Pair::nat => nat * (nat => bool) => nat * nat * (nat => bool)) t
+   ((Pair::nat => (nat => bool) => nat * (nat => bool))
+     ((Suc::nat => nat) n) s'))
+ ((Pair::nat => nat * (nat => bool) => nat * nat * (nat => bool))
+   ((Suc::nat => nat) t)
+   ((Pair::nat => (nat => bool) => nat * (nat => bool))
+     ((Suc::nat => nat) n) s)))))))))))
+   (%uniform_tupled::nat * nat * (nat => bool) => nat * (nat => bool).
+       (split::(nat => nat * (nat => bool) => nat * (nat => bool))
+               => nat * nat * (nat => bool) => nat * (nat => bool))
+        (%(v::nat) v1::nat * (nat => bool).
+            (nat_case::nat * (nat => bool)
+                       => (nat => nat * (nat => bool))
+                          => nat => nat * (nat => bool))
+             ((split::(nat => (nat => bool) => nat * (nat => bool))
+                      => nat * (nat => bool) => nat * (nat => bool))
+               (%(v3::nat) v4::nat => bool.
+                   (nat_case::nat * (nat => bool)
+                              => (nat => nat * (nat => bool))
+                                 => nat => nat * (nat => bool))
+                    (ARB::nat * (nat => bool))
+                    (%v5::nat.
+                        (Pair::nat => (nat => bool) => nat * (nat => bool))
+                         (0::nat) v4)
+                    v3)
+               v1)
+             (%v2::nat.
+                 (split::(nat => (nat => bool) => nat * (nat => bool))
+                         => nat * (nat => bool) => nat * (nat => bool))
+                  (%(v7::nat) v8::nat => bool.
+                      (nat_case::nat * (nat => bool)
+                                 => (nat => nat * (nat => bool))
+                                    => nat => nat * (nat => bool))
+                       (ARB::nat * (nat => bool))
+                       (%v9::nat.
+                           (Let::nat * (nat => bool)
+                                 => (nat * (nat => bool)
+                                     => nat * (nat => bool))
+                                    => nat * (nat => bool))
+                            ((unif::nat
+                                    => (nat => bool) => nat * (nat => bool))
+                              v9 v8)
+                            ((split::(nat
+=> (nat => bool) => nat * (nat => bool))
+                                     => nat * (nat => bool)
+  => nat * (nat => bool))
+                              (%(res::nat) s'::nat => bool.
+                                  (If::bool
+ => nat * (nat => bool) => nat * (nat => bool) => nat * (nat => bool))
+                                   ((op <::nat => nat => bool) res
+                                     ((Suc::nat => nat) v9))
+                                   ((Pair::nat
+     => (nat => bool) => nat * (nat => bool))
+                                     res s')
+                                   (uniform_tupled
+                                     ((Pair::nat
+       => nat * (nat => bool) => nat * nat * (nat => bool))
+ v2 ((Pair::nat => (nat => bool) => nat * (nat => bool))
+      ((Suc::nat => nat) v9) s'))))))
+                       v7)
+                  v1)
+             v)))"
+  by (import prob_uniform uniform_tupled_primitive_def)
+
+consts
+  uniform :: "nat => nat => (nat => bool) => nat * (nat => bool)" 
+
+defs
+  uniform_primdef: "uniform == %x x1 x2. uniform_tupled (x, x1, x2)"
+
+lemma uniform_curried_def: "ALL x x1 x2. uniform x x1 x2 = uniform_tupled (x, x1, x2)"
+  by (import prob_uniform uniform_curried_def)
+
+lemma uniform_ind: "(All::((nat => nat => (nat => bool) => bool) => bool) => bool)
+ (%P::nat => nat => (nat => bool) => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(nat => bool) => bool)
+          (%x::nat.
+              (All::((nat => bool) => bool) => bool)
+               (P ((Suc::nat => nat) x) (0::nat))))
+        ((op &::bool => bool => bool)
+          ((All::((nat => bool) => bool) => bool) (P (0::nat) (0::nat)))
+          ((op &::bool => bool => bool)
+            ((All::(nat => bool) => bool)
+              (%x::nat.
+                  (All::((nat => bool) => bool) => bool)
+                   (P (0::nat) ((Suc::nat => nat) x))))
+            ((All::(nat => bool) => bool)
+              (%x::nat.
+                  (All::(nat => bool) => bool)
+                   (%xa::nat.
+                       (All::((nat => bool) => bool) => bool)
+                        (%xb::nat => bool.
+                            (op -->::bool => bool => bool)
+                             ((All::(nat => bool) => bool)
+                               (%xc::nat.
+                                   (All::((nat => bool) => bool) => bool)
+                                    (%xd::nat => bool.
+  (op -->::bool => bool => bool)
+   ((op &::bool => bool => bool)
+     ((op =::nat * (nat => bool) => nat * (nat => bool) => bool)
+       ((Pair::nat => (nat => bool) => nat * (nat => bool)) xc xd)
+       ((unif::nat => (nat => bool) => nat * (nat => bool)) xa xb))
+     ((Not::bool => bool)
+       ((op <::nat => nat => bool) xc ((Suc::nat => nat) xa))))
+   (P x ((Suc::nat => nat) xa) xd))))
+                             (P ((Suc::nat => nat) x) ((Suc::nat => nat) xa)
+                               xb))))))))
+      ((All::(nat => bool) => bool)
+        (%x::nat.
+            (All::(nat => bool) => bool)
+             (%xa::nat. (All::((nat => bool) => bool) => bool) (P x xa)))))"
+  by (import prob_uniform uniform_ind)
+
+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. (Suc n div 2 = 0) = (n = 0)"
+  by (import prob_uniform SUC_DIV_TWO_ZERO)
+
+lemma UNIF_BOUND_LOWER: "ALL n. n < 2 ^ unif_bound n"
+  by (import prob_uniform UNIF_BOUND_LOWER)
+
+lemma UNIF_BOUND_LOWER_SUC: "ALL n. Suc n <= 2 ^ unif_bound n"
+  by (import prob_uniform UNIF_BOUND_LOWER_SUC)
+
+lemma UNIF_BOUND_UPPER: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (op -->::bool => bool => bool)
+      ((Not::bool => bool) ((op =::nat => nat => bool) n (0::nat)))
+      ((op <=::nat => nat => bool)
+        ((op ^::nat => nat => nat)
+          ((number_of::bin => nat)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool)))
+          ((unif_bound::nat => nat) n))
+        ((op *::nat => nat => nat)
+          ((number_of::bin => nat)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool)))
+          n)))"
+  by (import prob_uniform UNIF_BOUND_UPPER)
+
+lemma UNIF_BOUND_UPPER_SUC: "ALL n. 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.
+    unif (Suc n) =
+    BIND (unif (Suc n div 2))
+     (%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. uniform 0 (Suc x) = UNIT 0) &
+(ALL x xa.
+    uniform (Suc 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. indep (unif n)"
+  by (import prob_uniform INDEP_UNIF)
+
+lemma INDEP_UNIFORM: "ALL t n. indep (uniform t (Suc n))"
+  by (import prob_uniform INDEP_UNIFORM)
+
+lemma PROB_UNIF: "ALL n k.
+   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 s. fst (unif n s) < 2 ^ unif_bound n"
+  by (import prob_uniform UNIF_RANGE)
+
+lemma PROB_UNIF_PAIR: "ALL n k k'.
+   (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::(nat => bool) => bool)
+ (%n::nat.
+     (All::(nat => bool) => bool)
+      (%k::nat.
+          (op -->::bool => bool => bool)
+           ((op <=::nat => nat => bool) k
+             ((op ^::nat => nat => nat)
+               ((number_of::bin => nat)
+                 ((op BIT::bin => bool => bin)
+                   ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                     (True::bool))
+                   (False::bool)))
+               ((unif_bound::nat => nat) n)))
+           ((op =::real => real => bool)
+             ((prob::((nat => bool) => bool) => real)
+               (%s::nat => bool.
+                   (op <::nat => nat => bool)
+                    ((fst::nat * (nat => bool) => nat)
+                      ((unif::nat => (nat => bool) => nat * (nat => bool)) n
+                        s))
+                    k))
+             ((op *::real => real => real) ((real::nat => real) k)
+               ((op ^::real => nat => real)
+                 ((op /::real => real => real) (1::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool))))
+                 ((unif_bound::nat => nat) n))))))"
+  by (import prob_uniform PROB_UNIF_BOUND)
+
+lemma PROB_UNIF_GOOD: "ALL n. 1 / 2 <= prob (%s. fst (unif n s) < Suc n)"
+  by (import prob_uniform PROB_UNIF_GOOD)
+
+lemma UNIFORM_RANGE: "ALL t n s. 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)))))"
+  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)))))))"
+  by (import prob_uniform PROB_UNIFORM_UPPER_BOUND)
+
+lemma PROB_UNIFORM_PAIR_SUC: "(All::(nat => bool) => bool)
+ (%t::nat.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (All::(nat => bool) => bool)
+           (%k::nat.
+               (All::(nat => bool) => bool)
+                (%k'::nat.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::nat => nat => bool) k ((Suc::nat => nat) n))
+                       ((op <::nat => nat => bool) k'
+                         ((Suc::nat => nat) n)))
+                     ((op <=::real => real => bool)
+                       ((abs::real => real)
+                         ((op -::real => real => real)
+                           ((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))
+                                  k))
+                           ((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))
+                                  k'))))
+                       ((op ^::real => nat => real)
+                         ((op /::real => real => real) (1::real)
+                           ((number_of::bin => real)
+                             ((op BIT::bin => bool => bin)
+                               ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                                 (True::bool))
+                               (False::bool))))
+                         t))))))"
+  by (import prob_uniform PROB_UNIFORM_PAIR_SUC)
+
+lemma PROB_UNIFORM_SUC: "(All::(nat => bool) => bool)
+ (%t::nat.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (All::(nat => bool) => bool)
+           (%k::nat.
+               (op -->::bool => bool => bool)
+                ((op <::nat => nat => bool) k ((Suc::nat => nat) n))
+                ((op <=::real => real => bool)
+                  ((abs::real => real)
+                    ((op -::real => real => real)
+                      ((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))
+                             k))
+                      ((op /::real => real => real) (1::real)
+                        ((real::nat => real) ((Suc::nat => nat) n)))))
+                  ((op ^::real => nat => real)
+                    ((op /::real => real => real) (1::real)
+                      ((number_of::bin => real)
+                        ((op BIT::bin => bool => bin)
+                          ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                            (True::bool))
+                          (False::bool))))
+                    t)))))"
+  by (import prob_uniform PROB_UNIFORM_SUC)
+
+lemma PROB_UNIFORM: "(All::(nat => bool) => bool)
+ (%t::nat.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (All::(nat => bool) => bool)
+           (%k::nat.
+               (op -->::bool => bool => bool)
+                ((op <::nat => nat => bool) k n)
+                ((op <=::real => real => bool)
+                  ((abs::real => real)
+                    ((op -::real => real => real)
+                      ((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 n s))
+                             k))
+                      ((op /::real => real => real) (1::real)
+                        ((real::nat => real) n))))
+                  ((op ^::real => nat => real)
+                    ((op /::real => real => real) (1::real)
+                      ((number_of::bin => real)
+                        ((op BIT::bin => bool => bin)
+                          ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                            (True::bool))
+                          (False::bool))))
+                    t)))))"
+  by (import prob_uniform PROB_UNIFORM)
+
+;end_setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4Real.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,7644 @@
+theory HOL4Real = HOL4Base:
+
+;setup_theory realax
+
+lemma HREAL_RDISTRIB: "ALL x y z.
+   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 y. x ~= hreal_add x y"
+  by (import realax HREAL_EQ_ADDL)
+
+lemma HREAL_EQ_LADD: "ALL x y z. (hreal_add x y = hreal_add x z) = (y = z)"
+  by (import realax HREAL_EQ_LADD)
+
+lemma HREAL_LT_REFL: "ALL x. ~ hreal_lt x x"
+  by (import realax HREAL_LT_REFL)
+
+lemma HREAL_LT_ADDL: "ALL x y. hreal_lt x (hreal_add x y)"
+  by (import realax HREAL_LT_ADDL)
+
+lemma HREAL_LT_NE: "ALL x y. hreal_lt x y --> x ~= y"
+  by (import realax HREAL_LT_NE)
+
+lemma HREAL_LT_ADDR: "ALL x y. ~ hreal_lt (hreal_add x y) x"
+  by (import realax HREAL_LT_ADDR)
+
+lemma HREAL_LT_GT: "ALL x y. hreal_lt x y --> ~ hreal_lt y x"
+  by (import realax HREAL_LT_GT)
+
+lemma HREAL_LT_ADD2: "ALL x1 x2 y1 y2.
+   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 y z. hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z"
+  by (import realax HREAL_LT_LADD)
+
+constdefs
+  treal_0 :: "hreal * hreal" 
+  "treal_0 == (hreal_1, hreal_1)"
+
+lemma treal_0: "treal_0 = (hreal_1, hreal_1)"
+  by (import realax treal_0)
+
+constdefs
+  treal_1 :: "hreal * hreal" 
+  "treal_1 == (hreal_add hreal_1 hreal_1, hreal_1)"
+
+lemma treal_1: "treal_1 = (hreal_add hreal_1 hreal_1, hreal_1)"
+  by (import realax treal_1)
+
+constdefs
+  treal_neg :: "hreal * hreal => hreal * hreal" 
+  "treal_neg == %(x, y). (y, x)"
+
+lemma treal_neg: "ALL x y. treal_neg (x, y) = (y, x)"
+  by (import realax treal_neg)
+
+constdefs
+  treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+  "treal_add == %(x1, y1) (x2, y2). (hreal_add x1 x2, hreal_add y1 y2)"
+
+lemma treal_add: "ALL x1 y1 x2 y2.
+   treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)"
+  by (import realax treal_add)
+
+constdefs
+  treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" 
+  "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: "ALL x1 y1 x2 y2.
+   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)
+
+constdefs
+  treal_lt :: "hreal * hreal => hreal * hreal => bool" 
+  "treal_lt == %(x1, y1) (x2, y2). hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
+
+lemma treal_lt: "ALL x1 y1 x2 y2.
+   treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)"
+  by (import realax treal_lt)
+
+constdefs
+  treal_inv :: "hreal * hreal => hreal * hreal" 
+  "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: "ALL x y.
+   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)
+
+constdefs
+  treal_eq :: "hreal * hreal => hreal * hreal => bool" 
+  "treal_eq == %(x1, y1) (x2, y2). hreal_add x1 y2 = hreal_add x2 y1"
+
+lemma treal_eq: "ALL x1 y1 x2 y2.
+   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. treal_eq x x"
+  by (import realax TREAL_EQ_REFL)
+
+lemma TREAL_EQ_SYM: "ALL x y. treal_eq x y = treal_eq y x"
+  by (import realax TREAL_EQ_SYM)
+
+lemma TREAL_EQ_TRANS: "ALL x y z. treal_eq x y & treal_eq y z --> treal_eq x z"
+  by (import realax TREAL_EQ_TRANS)
+
+lemma TREAL_EQ_EQUIV: "ALL p q. treal_eq p q = (treal_eq p = treal_eq q)"
+  by (import realax TREAL_EQ_EQUIV)
+
+lemma TREAL_EQ_AP: "ALL p q. 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 y. treal_add x y = treal_add y x"
+  by (import realax TREAL_ADD_SYM)
+
+lemma TREAL_MUL_SYM: "ALL x y. treal_mul x y = treal_mul y x"
+  by (import realax TREAL_MUL_SYM)
+
+lemma TREAL_ADD_ASSOC: "ALL x y z. 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 y z. 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 y z.
+   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. treal_eq (treal_add treal_0 x) x"
+  by (import realax TREAL_ADD_LID)
+
+lemma TREAL_MUL_LID: "ALL x. treal_eq (treal_mul treal_1 x) x"
+  by (import realax TREAL_MUL_LID)
+
+lemma TREAL_ADD_LINV: "ALL x. 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. ~ 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 y. treal_eq x y | treal_lt x y | treal_lt y x"
+  by (import realax TREAL_LT_TOTAL)
+
+lemma TREAL_LT_REFL: "ALL x. ~ treal_lt x x"
+  by (import realax TREAL_LT_REFL)
+
+lemma TREAL_LT_TRANS: "ALL x y z. treal_lt x y & treal_lt y z --> treal_lt x z"
+  by (import realax TREAL_LT_TRANS)
+
+lemma TREAL_LT_ADD: "ALL x y z. 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 y.
+   treal_lt treal_0 x & treal_lt treal_0 y -->
+   treal_lt treal_0 (treal_mul x y)"
+  by (import realax TREAL_LT_MUL)
+
+constdefs
+  treal_of_hreal :: "hreal => hreal * hreal" 
+  "treal_of_hreal == %x. (hreal_add x hreal_1, hreal_1)"
+
+lemma treal_of_hreal: "ALL x. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)"
+  by (import realax treal_of_hreal)
+
+constdefs
+  hreal_of_treal :: "hreal * hreal => hreal" 
+  "hreal_of_treal == %(x, y). SOME d. x = hreal_add y d"
+
+lemma hreal_of_treal: "ALL x y. hreal_of_treal (x, y) = (SOME d. x = hreal_add y d)"
+  by (import realax hreal_of_treal)
+
+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 i. 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 i. treal_eq h i --> hreal_of_treal h = hreal_of_treal i"
+  by (import realax TREAL_BIJ_WELLDEF)
+
+lemma TREAL_NEG_WELLDEF: "ALL x1 x2. treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)"
+  by (import realax TREAL_NEG_WELLDEF)
+
+lemma TREAL_ADD_WELLDEFR: "ALL x1 x2 y. 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 x2 y1 y2.
+   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 x2 y. 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 x2 y1 y2.
+   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 x2 y. treal_eq x1 x2 --> treal_lt x1 y = treal_lt x2 y"
+  by (import realax TREAL_LT_WELLDEFR)
+
+lemma TREAL_LT_WELLDEFL: "ALL x y1 y2. treal_eq y1 y2 --> treal_lt x y1 = treal_lt x y2"
+  by (import realax TREAL_LT_WELLDEFL)
+
+lemma TREAL_LT_WELLDEF: "ALL x1 x2 y1 y2.
+   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 x2. 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: "(0::real) = (0::real)"
+  by (import real REAL_0)
+
+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::real))"
+  by (import real REAL_ADD_LID_UNIQ)
+
+lemma REAL_ADD_RID_UNIQ: "ALL (x::real) y::real. (x + y = x) = (y = (0::real))"
+  by (import real REAL_ADD_RID_UNIQ)
+
+lemma REAL_LNEG_UNIQ: "ALL (x::real) y::real. (x + y = (0::real)) = (x = - y)"
+  by (import real REAL_LNEG_UNIQ)
+
+lemma REAL_LT_ANTISYM: "ALL (x::real) y::real. ~ (x < y & y < x)"
+  by (import real REAL_LT_ANTISYM)
+
+lemma REAL_LET_TOTAL: "ALL (x::real) y::real. x <= y | y < x"
+  by (import real REAL_LET_TOTAL)
+
+lemma REAL_LTE_TOTAL: "ALL (x::real) y::real. x < y | y <= x"
+  by (import real REAL_LTE_TOTAL)
+
+lemma REAL_LET_ANTISYM: "ALL (x::real) y::real. ~ (x < y & y <= x)"
+  by (import real REAL_LET_ANTISYM)
+
+lemma REAL_LTE_ANTSYM: "ALL (x::real) y::real. ~ (x <= y & y < x)"
+  by (import real REAL_LTE_ANTSYM)
+
+lemma REAL_LT_NEGTOTAL: "ALL x::real. x = (0::real) | (0::real) < x | (0::real) < - x"
+  by (import real REAL_LT_NEGTOTAL)
+
+lemma REAL_LE_NEGTOTAL: "ALL x::real. (0::real) <= x | (0::real) <= - x"
+  by (import real REAL_LE_NEGTOTAL)
+
+lemma REAL_LE_MUL: "ALL (x::real) y::real.
+   (0::real) <= x & (0::real) <= y --> (0::real) <= x * y"
+  by (import real REAL_LE_MUL)
+
+lemma REAL_LT_ADDNEG: "ALL (x::real) (y::real) z::real. (y < x + - z) = (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)"
+  by (import real REAL_LT_ADDNEG2)
+
+lemma REAL_LT_ADD1: "ALL (x::real) y::real. x <= y --> x < y + (1::real)"
+  by (import real REAL_LT_ADD1)
+
+lemma REAL_LE_DOUBLE: "ALL x::real. ((0::real) <= x + x) = ((0::real) <= x)"
+  by (import real REAL_LE_DOUBLE)
+
+lemma REAL_NEG_EQ: "ALL (x::real) y::real. (- x = y) = (x = - y)"
+  by (import real REAL_NEG_EQ)
+
+lemma REAL_NEG_MINUS1: "ALL x::real. - x = - (1::real) * x"
+  by (import real REAL_NEG_MINUS1)
+
+lemma REAL_LT_LMUL_0: "ALL (x::real) y::real.
+   (0::real) < x --> ((0::real) < x * y) = ((0::real) < y)"
+  by (import real REAL_LT_LMUL_0)
+
+lemma REAL_LT_RMUL_0: "ALL (x::real) y::real.
+   (0::real) < y --> ((0::real) < x * y) = ((0::real) < x)"
+  by (import real REAL_LT_RMUL_0)
+
+lemma REAL_LT_LMUL: "ALL (x::real) (y::real) z::real. (0::real) < 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::real) --> x = inverse y"
+  by (import real REAL_LINV_UNIQ)
+
+lemma REAL_LE_INV: "ALL x::real. (0::real) <= x --> (0::real) <= inverse x"
+  by (import real REAL_LE_INV)
+
+lemma REAL_LE_ADDR: "ALL (x::real) y::real. (x <= x + y) = ((0::real) <= y)"
+  by (import real REAL_LE_ADDR)
+
+lemma REAL_LE_ADDL: "ALL (x::real) y::real. (y <= x + y) = ((0::real) <= x)"
+  by (import real REAL_LE_ADDL)
+
+lemma REAL_LT_ADDR: "ALL (x::real) y::real. (x < x + y) = ((0::real) < y)"
+  by (import real REAL_LT_ADDR)
+
+lemma REAL_LT_ADDL: "ALL (x::real) y::real. (y < x + y) = ((0::real) < x)"
+  by (import real REAL_LT_ADDL)
+
+lemma REAL_LT_NZ: "ALL n::nat. (real n ~= (0::real)) = ((0::real) < real n)"
+  by (import real REAL_LT_NZ)
+
+lemma REAL_NZ_IMP_LT: "ALL n::nat. n ~= (0::nat) --> (0::real) < real n"
+  by (import real REAL_NZ_IMP_LT)
+
+lemma REAL_LT_RDIV_0: "ALL (y::real) z::real.
+   (0::real) < z --> ((0::real) < y / z) = ((0::real) < y)"
+  by (import real REAL_LT_RDIV_0)
+
+lemma REAL_LT_RDIV: "ALL (x::real) (y::real) z::real. (0::real) < z --> (x / z < y / z) = (x < y)"
+  by (import real REAL_LT_RDIV)
+
+lemma REAL_LT_FRACTION_0: "ALL (n::nat) d::real.
+   n ~= (0::nat) --> ((0::real) < d / real n) = ((0::real) < d)"
+  by (import real REAL_LT_FRACTION_0)
+
+lemma REAL_LT_MULTIPLE: "ALL (x::nat) xa::real.
+   (1::nat) < x --> (xa < real x * xa) = ((0::real) < xa)"
+  by (import real REAL_LT_MULTIPLE)
+
+lemma REAL_LT_FRACTION: "ALL (n::nat) d::real. (1::nat) < n --> (d / real n < d) = ((0::real) < d)"
+  by (import real REAL_LT_FRACTION)
+
+lemma REAL_LT_HALF2: "ALL d::real. (d / (2::real) < d) = ((0::real) < d)"
+  by (import real REAL_LT_HALF2)
+
+lemma REAL_DIV_LMUL: "ALL (x::real) y::real. y ~= (0::real) --> y * (x / y) = x"
+  by (import real REAL_DIV_LMUL)
+
+lemma REAL_DIV_RMUL: "ALL (x::real) y::real. y ~= (0::real) --> x / y * y = x"
+  by (import real REAL_DIV_RMUL)
+
+lemma REAL_DOWN: "ALL x::real. (0::real) < x --> (EX xa::real. (0::real) < xa & xa < x)"
+  by (import real REAL_DOWN)
+
+lemma REAL_SUB_SUB: "ALL (x::real) y::real. x - y - 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_LET_ADD: "ALL (x::real) y::real. (0::real) <= x & (0::real) < y --> (0::real) < x + y"
+  by (import real REAL_LET_ADD)
+
+lemma REAL_LTE_ADD: "ALL (x::real) y::real. (0::real) < x & (0::real) <= y --> (0::real) < x + y"
+  by (import real REAL_LTE_ADD)
+
+lemma REAL_SUB_LNEG: "ALL (x::real) y::real. - x - y = - (x + y)"
+  by (import real REAL_SUB_LNEG)
+
+lemma REAL_SUB_NEG2: "ALL (x::real) y::real. - x - - y = 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"
+  by (import real REAL_SUB_TRIANGLE)
+
+lemma REAL_INV_MUL: "ALL (x::real) y::real.
+   x ~= (0::real) & y ~= (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::real) & y ~= (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"
+  by (import real REAL_SUB_SUB2)
+
+lemma REAL_ADD_SUB2: "ALL (x::real) y::real. x - (x + y) = - y"
+  by (import real REAL_ADD_SUB2)
+
+lemma REAL_LE_MUL2: "ALL (x1::real) (x2::real) (y1::real) y2::real.
+   (0::real) <= x1 & (0::real) <= y1 & x1 <= x2 & y1 <= y2 -->
+   x1 * y1 <= x2 * y2"
+  by (import real REAL_LE_MUL2)
+
+lemma REAL_LE_LDIV: "ALL (x::real) (y::real) z::real. (0::real) < x & y <= z * x --> y / x <= z"
+  by (import real REAL_LE_LDIV)
+
+lemma REAL_LE_RDIV: "ALL (x::real) (y::real) z::real. (0::real) < x & y * x <= z --> y <= z / x"
+  by (import real REAL_LE_RDIV)
+
+lemma REAL_LT_DIV: "ALL (x::real) xa::real.
+   (0::real) < x & (0::real) < xa --> (0::real) < x / xa"
+  by (import real REAL_LT_DIV)
+
+lemma REAL_LE_DIV: "ALL (x::real) xa::real.
+   (0::real) <= x & (0::real) <= xa --> (0::real) <= x / xa"
+  by (import real REAL_LE_DIV)
+
+lemma REAL_LT_1: "ALL (x::real) y::real. (0::real) <= x & x < y --> x / y < (1::real)"
+  by (import real REAL_LT_1)
+
+lemma REAL_POS_NZ: "ALL x::real. (0::real) < x --> x ~= (0::real)"
+  by (import real REAL_POS_NZ)
+
+lemma REAL_EQ_LMUL_IMP: "ALL (x::real) (xa::real) xb::real.
+   x ~= (0::real) & x * xa = x * xb --> xa = xb"
+  by (import real REAL_EQ_LMUL_IMP)
+
+lemma REAL_FACT_NZ: "ALL n. 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::real) < x * x) = (x ~= (0::real))"
+  by (import real REAL_POASQ)
+
+lemma REAL_SUMSQ: "ALL (x::real) y::real.
+   (x * x + y * y = (0::real)) = (x = (0::real) & y = (0::real))"
+  by (import real REAL_SUMSQ)
+
+lemma REAL_MIDDLE1: "ALL (a::real) b::real. a <= b --> a <= (a + b) / (2::real)"
+  by (import real REAL_MIDDLE1)
+
+lemma REAL_MIDDLE2: "ALL (a::real) b::real. a <= b --> (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"
+  by (import real ABS_LT_MUL2)
+
+lemma ABS_SUB: "ALL (x::real) y::real. abs (x - y) = abs (y - x)"
+  by (import real ABS_SUB)
+
+lemma ABS_REFL: "ALL x::real. (abs x = x) = ((0::real) <= x)"
+  by (import real ABS_REFL)
+
+lemma ABS_BETWEEN: "ALL (x::real) (y::real) d::real.
+   ((0::real) < d & x - d < y & 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"
+  by (import real ABS_BOUND)
+
+lemma ABS_STILLNZ: "ALL (x::real) y::real. abs (x - y) < abs y --> x ~= (0::real)"
+  by (import real ABS_STILLNZ)
+
+lemma ABS_CASES: "ALL x::real. x = (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"
+  by (import real ABS_BETWEEN1)
+
+lemma ABS_SIGN: "ALL (x::real) y::real. abs (x - y) < y --> (0::real) < x"
+  by (import real ABS_SIGN)
+
+lemma ABS_SIGN2: "ALL (x::real) y::real. abs (x - y) < - 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"
+  by (import real ABS_CIRCLE)
+
+lemma REAL_SUB_ABS: "ALL (x::real) y::real. abs x - abs y <= abs (x - y)"
+  by (import real REAL_SUB_ABS)
+
+lemma ABS_SUB_ABS: "ALL (x::real) y::real. abs (abs x - abs y) <= abs (x - y)"
+  by (import real ABS_SUB_ABS)
+
+lemma ABS_BETWEEN2: "ALL (x0::real) (x::real) (y0::real) y::real.
+   x0 < y0 &
+   abs (x - x0) < (y0 - x0) / (2::real) &
+   abs (y - y0) < (y0 - x0) / (2::real) -->
+   x < y"
+  by (import real ABS_BETWEEN2)
+
+lemma POW_PLUS1: "ALL e. 0 < e --> (ALL n. 1 + real n * e <= (1 + e) ^ n)"
+  by (import real POW_PLUS1)
+
+lemma POW_M1: "ALL n::nat. abs ((- (1::real)) ^ n) = (1::real)"
+  by (import real POW_M1)
+
+lemma REAL_LE1_POW2: "ALL x::real. (1::real) <= x --> (1::real) <= x ^ 2"
+  by (import real REAL_LE1_POW2)
+
+lemma REAL_LT1_POW2: "ALL x::real. (1::real) < x --> (1::real) < x ^ 2"
+  by (import real REAL_LT1_POW2)
+
+lemma POW_POS_LT: "ALL (x::real) n::nat. (0::real) < x --> (0::real) < x ^ Suc n"
+  by (import real POW_POS_LT)
+
+lemma POW_LT: "ALL (n::nat) (x::real) y::real.
+   (0::real) <= x & x < y --> x ^ Suc n < y ^ Suc n"
+  by (import real POW_LT)
+
+lemma POW_ZERO_EQ: "ALL (n::nat) x::real. (x ^ Suc n = (0::real)) = (x = (0::real))"
+  by (import real POW_ZERO_EQ)
+
+lemma REAL_POW_LT2: "ALL (n::nat) (x::real) y::real.
+   n ~= (0::nat) & (0::real) <= x & x < y --> x ^ n < y ^ n"
+  by (import real REAL_POW_LT2)
+
+lemma REAL_SUP_SOMEPOS: "ALL P::real => bool.
+   (EX x::real. P 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))"
+  by (import real SUP_LEMMA1)
+
+lemma SUP_LEMMA2: "ALL P::real => bool.
+   Ex P --> (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)"
+  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))"
+  by (import real REAL_SUP_EXISTS)
+
+constdefs
+  sup :: "(real => bool) => real" 
+  "sup == %P. SOME s. ALL y. (EX x. P x & y < x) = (y < s)"
+
+lemma sup: "ALL P. sup P = (SOME s. ALL y. (EX x. P x & y < x) = (y < s))"
+  by (import real sup)
+
+lemma REAL_SUP: "ALL P.
+   Ex P & (EX z. ALL x. P x --> x < z) -->
+   (ALL y. (EX x. P x & y < x) = (y < sup P))"
+  by (import real REAL_SUP)
+
+lemma REAL_SUP_UBOUND: "ALL P. Ex P & (EX z. ALL x. P x --> x < z) --> (ALL y. P y --> y <= 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))"
+  by (import real SETOK_LE_LT)
+
+lemma REAL_SUP_LE: "ALL P.
+   Ex P & (EX z. ALL x. P x --> x <= z) -->
+   (ALL y. (EX x. P x & y < x) = (y < sup P))"
+  by (import real REAL_SUP_LE)
+
+lemma REAL_SUP_UBOUND_LE: "ALL P. Ex P & (EX z. ALL x. P x --> x <= z) --> (ALL y. P y --> y <= sup P)"
+  by (import real REAL_SUP_UBOUND_LE)
+
+lemma REAL_ARCH_LEAST: "ALL y.
+   0 < y -->
+   (ALL x. 0 <= x --> (EX n. 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 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)
+
+constdefs
+  sum :: "nat * nat => (nat => real) => real" 
+  "real.sum == split sumc"
+
+lemma SUM_DEF: "ALL m n f. real.sum (m, n) f = sumc m n f"
+  by (import real SUM_DEF)
+
+lemma sum: "ALL x xa xb.
+   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 n p. 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 m n. 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 m n. abs (real.sum (m, n) f) <= real.sum (m, n) (%n. abs (f n))"
+  by (import real ABS_SUM)
+
+lemma SUM_LE: "ALL f g m n.
+   (ALL 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 g m n.
+   (ALL 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. (ALL n. 0 <= f n) --> (ALL m n. 0 <= real.sum (m, n) f)"
+  by (import real SUM_POS)
+
+lemma SUM_POS_GEN: "ALL f m. (ALL n. m <= n --> 0 <= f n) --> (ALL n. 0 <= real.sum (m, n) f)"
+  by (import real SUM_POS_GEN)
+
+lemma SUM_ABS: "ALL f m x.
+   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 m n. 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 N.
+   (ALL n. N <= n --> f n = 0) -->
+   (ALL m n. N <= m --> real.sum (m, n) f = 0)"
+  by (import real SUM_ZERO)
+
+lemma SUM_ADD: "ALL f g m n.
+   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 c m n. real.sum (m, n) (%n. c * f n) = c * real.sum (m, n) f"
+  by (import real SUM_CMUL)
+
+lemma SUM_NEG: "ALL f n d. real.sum (n, d) (%n. - f n) = - real.sum (n, d) f"
+  by (import real SUM_NEG)
+
+lemma SUM_SUB: "ALL f g m n.
+   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 g m n.
+   (ALL 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 f c. 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 k m n.
+   (ALL 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 k f.
+   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 n. real.sum (n, 1) f = f n"
+  by (import real SUM_1)
+
+lemma SUM_2: "ALL f n. real.sum (n, 2) f = f n + f (n + 1)"
+  by (import real SUM_2)
+
+lemma SUM_OFFSET: "ALL f n k.
+   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 m k n. real.sum (m + k, n) f = real.sum (m, n) (%r. f (r + k))"
+  by (import real SUM_REINDEX)
+
+lemma SUM_0: "ALL m n. real.sum (m, n) (%r. 0) = 0"
+  by (import real SUM_0)
+
+lemma SUM_PERMUTE_0: "ALL n p.
+   (ALL y<n. EX! x. x < n & p x = y) -->
+   (ALL f. real.sum (0, n) (%n. f (p n)) = real.sum (0, n) f)"
+  by (import real SUM_PERMUTE_0)
+
+lemma SUM_CANCEL: "ALL f n d. real.sum (n, d) (%n. f (Suc n) - f n) = f (n + d) - f n"
+  by (import real SUM_CANCEL)
+
+lemma REAL_LE_RNEG: "ALL (x::real) y::real. (x <= - y) = (x + y <= (0::real))"
+  by (import real REAL_LE_RNEG)
+
+lemma REAL_EQ_RDIV_EQ: "ALL (x::real) (xa::real) xb::real.
+   (0::real) < xb --> (x = xa / xb) = (x * xb = xa)"
+  by (import real REAL_EQ_RDIV_EQ)
+
+lemma REAL_EQ_LDIV_EQ: "ALL (x::real) (xa::real) xb::real.
+   (0::real) < xb --> (x / xb = xa) = (x = xa * xb)"
+  by (import real REAL_EQ_LDIV_EQ)
+
+;end_setup
+
+;setup_theory topology
+
+constdefs
+  re_Union :: "(('a => bool) => bool) => 'a => bool" 
+  "re_Union == %P x. EX s. P s & s x"
+
+lemma re_Union: "ALL P. re_Union P = (%x. EX s. P s & s x)"
+  by (import topology re_Union)
+
+constdefs
+  re_union :: "('a => bool) => ('a => bool) => 'a => bool" 
+  "re_union == %P Q x. P x | Q x"
+
+lemma re_union: "ALL P Q. re_union P Q = (%x. P x | Q x)"
+  by (import topology re_union)
+
+constdefs
+  re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" 
+  "re_intersect == %P Q x. P x & Q x"
+
+lemma re_intersect: "ALL P Q. re_intersect P Q = (%x. P x & Q x)"
+  by (import topology re_intersect)
+
+constdefs
+  re_null :: "'a => bool" 
+  "re_null == %x. False"
+
+lemma re_null: "re_null = (%x. False)"
+  by (import topology re_null)
+
+constdefs
+  re_universe :: "'a => bool" 
+  "re_universe == %x. True"
+
+lemma re_universe: "re_universe = (%x. True)"
+  by (import topology re_universe)
+
+constdefs
+  re_subset :: "('a => bool) => ('a => bool) => bool" 
+  "re_subset == %P Q. ALL x. P x --> Q x"
+
+lemma re_subset: "ALL P Q. re_subset P Q = (ALL x. P x --> Q x)"
+  by (import topology re_subset)
+
+constdefs
+  re_compl :: "('a => bool) => 'a => bool" 
+  "re_compl == %P x. ~ P x"
+
+lemma re_compl: "ALL P. re_compl P = (%x. ~ P x)"
+  by (import topology re_compl)
+
+lemma SUBSET_REFL: "ALL P. re_subset P P"
+  by (import topology SUBSET_REFL)
+
+lemma COMPL_MEM: "ALL P x. P x = (~ re_compl P x)"
+  by (import topology COMPL_MEM)
+
+lemma SUBSET_ANTISYM: "ALL P Q. (re_subset P Q & re_subset Q P) = (P = Q)"
+  by (import topology SUBSET_ANTISYM)
+
+lemma SUBSET_TRANS: "ALL P Q R. re_subset P Q & re_subset Q R --> re_subset P R"
+  by (import topology SUBSET_TRANS)
+
+constdefs
+  istopology :: "(('a => bool) => bool) => bool" 
+  "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))"
+
+lemma istopology: "ALL L.
+   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 => bool) => bool) => bool) => (('a => bool) => bool) set)
+ (istopology::(('a => bool) => bool) => bool)" 
+  by (rule typedef_helper,import topology topology_TY_DEF)
+
+lemmas topology_TY_DEF = typedef_hol2hol4 [OF type_definition_topology]
+
+consts
+  topology :: "(('a => bool) => bool) => 'a topology" 
+  "open" :: "'a topology => ('a => bool) => bool" 
+
+specification ("open" topology) topology_tybij: "(ALL a::'a topology. topology (open a) = a) &
+(ALL r::('a => bool) => bool. istopology r = (open (topology r) = r))"
+  by (import topology topology_tybij)
+
+lemma TOPOLOGY: "ALL L.
+   open L re_null &
+   open L re_universe &
+   (ALL a b. open L a & open L b --> open L (re_intersect a b)) &
+   (ALL P. re_subset P (open L) --> open L (re_Union P))"
+  by (import topology TOPOLOGY)
+
+lemma TOPOLOGY_UNION: "ALL x xa. re_subset xa (open x) --> open x (re_Union xa)"
+  by (import topology TOPOLOGY_UNION)
+
+constdefs
+  neigh :: "'a topology => ('a => bool) * 'a => bool" 
+  "neigh == %top (N, x). EX P. open top P & re_subset P N & P x"
+
+lemma neigh: "ALL top N x. neigh top (N, x) = (EX P. open top P & re_subset P N & P x)"
+  by (import topology neigh)
+
+lemma OPEN_OWN_NEIGH: "ALL S' top x. open top S' & S' x --> neigh top (S', x)"
+  by (import topology OPEN_OWN_NEIGH)
+
+lemma OPEN_UNOPEN: "ALL S' top. open top S' = (re_Union (%P. open top P & re_subset P S') = S')"
+  by (import topology OPEN_UNOPEN)
+
+lemma OPEN_SUBOPEN: "ALL S' top.
+   open top S' = (ALL x. S' x --> (EX P. P x & open top P & re_subset P S'))"
+  by (import topology OPEN_SUBOPEN)
+
+lemma OPEN_NEIGH: "ALL S' top.
+   open top S' = (ALL x. S' x --> (EX N. neigh top (N, x) & re_subset N S'))"
+  by (import topology OPEN_NEIGH)
+
+constdefs
+  closed :: "'a topology => ('a => bool) => bool" 
+  "closed == %L S'. open L (re_compl S')"
+
+lemma closed: "ALL L S'. closed L S' = open L (re_compl S')"
+  by (import topology closed)
+
+constdefs
+  limpt :: "'a topology => 'a => ('a => bool) => bool" 
+  "limpt == %top x S'. ALL N. neigh top (N, x) --> (EX y. x ~= y & S' y & N y)"
+
+lemma limpt: "ALL top x S'.
+   limpt top x S' =
+   (ALL N. neigh top (N, x) --> (EX y. x ~= y & S' y & N y))"
+  by (import topology limpt)
+
+lemma CLOSED_LIMPT: "ALL top S'. closed top S' = (ALL x. limpt top x S' --> S' x)"
+  by (import topology CLOSED_LIMPT)
+
+constdefs
+  ismet :: "('a * 'a => real) => bool" 
+  "ismet ==
+%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: "ALL m.
+   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 * 'a => real) => bool) => ('a * 'a => real) set)
+ (ismet::('a * 'a => real) => bool)" 
+  by (rule typedef_helper,import topology metric_TY_DEF)
+
+lemmas metric_TY_DEF = typedef_hol2hol4 [OF type_definition_metric]
+
+consts
+  metric :: "('a * 'a => real) => 'a metric" 
+  dist :: "'a metric => 'a * 'a => real" 
+
+specification (dist metric) metric_tybij: "(ALL a::'a metric. metric (dist a) = a) &
+(ALL r::'a * 'a => real. ismet r = (dist (metric r) = r))"
+  by (import topology metric_tybij)
+
+lemma METRIC_ISMET: "ALL m. ismet (dist m)"
+  by (import topology METRIC_ISMET)
+
+lemma METRIC_ZERO: "ALL m x y. (dist m (x, y) = 0) = (x = y)"
+  by (import topology METRIC_ZERO)
+
+lemma METRIC_SAME: "ALL m x. dist m (x, x) = 0"
+  by (import topology METRIC_SAME)
+
+lemma METRIC_POS: "ALL m x y. 0 <= dist m (x, y)"
+  by (import topology METRIC_POS)
+
+lemma METRIC_SYM: "ALL m x y. dist m (x, y) = dist m (y, x)"
+  by (import topology METRIC_SYM)
+
+lemma METRIC_TRIANGLE: "ALL m x y z. dist m (x, z) <= dist m (x, y) + dist m (y, z)"
+  by (import topology METRIC_TRIANGLE)
+
+lemma METRIC_NZ: "ALL m x y. x ~= y --> 0 < dist m (x, y)"
+  by (import topology METRIC_NZ)
+
+constdefs
+  mtop :: "'a metric => 'a topology" 
+  "mtop ==
+%m. topology
+     (%S'. ALL x.
+              S' x --> (EX e. 0 < e & (ALL y. dist m (x, y) < e --> S' y)))"
+
+lemma mtop: "ALL m.
+   mtop m =
+   topology
+    (%S'. ALL x.
+             S' x --> (EX e. 0 < e & (ALL y. dist m (x, y) < e --> S' y)))"
+  by (import topology mtop)
+
+lemma mtop_istopology: "ALL m.
+   istopology
+    (%S'. ALL x.
+             S' x --> (EX e. 0 < e & (ALL y. dist m (x, y) < e --> S' y)))"
+  by (import topology mtop_istopology)
+
+lemma MTOP_OPEN: "ALL S' x.
+   open (mtop x) S' =
+   (ALL xa. S' xa --> (EX e. 0 < e & (ALL y. dist x (xa, y) < e --> S' y)))"
+  by (import topology MTOP_OPEN)
+
+constdefs
+  B :: "'a metric => 'a * real => 'a => bool" 
+  "B == %m (x, e) y. dist m (x, y) < e"
+
+lemma ball: "ALL m x e. B m (x, e) = (%y. dist m (x, y) < e)"
+  by (import topology ball)
+
+lemma BALL_OPEN: "ALL m x e. 0 < e --> open (mtop m) (B m (x, e))"
+  by (import topology BALL_OPEN)
+
+lemma BALL_NEIGH: "ALL m x e. 0 < e --> neigh (mtop m) (B m (x, e), x)"
+  by (import topology BALL_NEIGH)
+
+lemma MTOP_LIMPT: "ALL m x S'.
+   limpt (mtop m) x S' =
+   (ALL e. 0 < e --> (EX y. x ~= y & S' y & dist m (x, y) < e))"
+  by (import topology MTOP_LIMPT)
+
+lemma ISMET_R1: "ismet (%(x, y). abs (y - x))"
+  by (import topology ISMET_R1)
+
+constdefs
+  mr1 :: "real metric" 
+  "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 y. dist mr1 (x, y) = abs (y - x)"
+  by (import topology MR1_DEF)
+
+lemma MR1_ADD: "ALL x d. dist mr1 (x, x + d) = abs d"
+  by (import topology MR1_ADD)
+
+lemma MR1_SUB: "ALL x d. dist mr1 (x, x - d) = abs d"
+  by (import topology MR1_SUB)
+
+lemma MR1_ADD_POS: "ALL x d. 0 <= d --> dist mr1 (x, x + d) = d"
+  by (import topology MR1_ADD_POS)
+
+lemma MR1_SUB_LE: "ALL x d. 0 <= d --> dist mr1 (x, x - d) = d"
+  by (import topology MR1_SUB_LE)
+
+lemma MR1_ADD_LT: "ALL x d. 0 < d --> dist mr1 (x, x + d) = d"
+  by (import topology MR1_ADD_LT)
+
+lemma MR1_SUB_LT: "ALL x d. 0 < d --> dist mr1 (x, x - d) = d"
+  by (import topology MR1_SUB_LT)
+
+lemma MR1_BETWEEN1: "ALL x y z. x < z & dist mr1 (x, y) < z - x --> y < z"
+  by (import topology MR1_BETWEEN1)
+
+lemma MR1_LIMPT: "ALL x. limpt (mtop mr1) x re_universe"
+  by (import topology MR1_LIMPT)
+
+;end_setup
+
+;setup_theory nets
+
+constdefs
+  dorder :: "('a => 'a => bool) => bool" 
+  "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))"
+
+lemma dorder: "ALL g.
+   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)
+
+constdefs
+  tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" 
+  "tends ==
+%(s::'b => 'a) (l::'a) (top::'a topology, g::'b => 'b => bool).
+   ALL N::'a => bool.
+      neigh top (N, l) -->
+      (EX n::'b. g n n & (ALL m::'b. g m n --> N (s m)))"
+
+lemma tends: "ALL (s::'b => 'a) (l::'a) (top::'a topology) g::'b => 'b => bool.
+   tends s l (top, g) =
+   (ALL N::'a => bool.
+       neigh top (N, l) -->
+       (EX n::'b. g n n & (ALL m::'b. g m n --> N (s m))))"
+  by (import nets tends)
+
+constdefs
+  bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" 
+  "bounded ==
+%(m, g) f. EX k x N. g N N & (ALL n. g n N --> dist m (f n, x) < k)"
+
+lemma bounded: "ALL m g f.
+   bounded (m, g) f =
+   (EX k x N. g N N & (ALL n. g n N --> dist m (f n, x) < k))"
+  by (import nets bounded)
+
+constdefs
+  tendsto :: "'a metric * 'a => 'a => 'a => bool" 
+  "tendsto == %(m, x) y z. 0 < dist m (x, y) & dist m (x, y) <= dist m (x, z)"
+
+lemma tendsto: "ALL m x y z.
+   tendsto (m, x) y z = (0 < dist m (x, y) & dist m (x, y) <= dist m (x, z))"
+  by (import nets tendsto)
+
+lemma DORDER_LEMMA: "ALL g.
+   dorder g -->
+   (ALL P Q.
+       (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 x. dorder (tendsto (m, x))"
+  by (import nets DORDER_TENDSTO)
+
+lemma MTOP_TENDS: "ALL d g x x0.
+   tends x x0 (mtop d, g) =
+   (ALL e.
+       0 < e --> (EX n. g n n & (ALL m. g m n --> dist d (x m, x0) < e)))"
+  by (import nets MTOP_TENDS)
+
+lemma MTOP_TENDS_UNIQ: "ALL (g::'b => 'b => bool) d::'a metric.
+   dorder g -->
+   tends (x::'b => 'a) (x0::'a) (mtop d, g) &
+   tends x (x1::'a) (mtop d, g) -->
+   x0 = x1"
+  by (import nets MTOP_TENDS_UNIQ)
+
+lemma SEQ_TENDS: "ALL d x x0.
+   tends x x0 (mtop d, nat_ge) =
+   (ALL xa. 0 < xa --> (EX xb. ALL xc. xb <= xc --> dist d (x xc, x0) < xa))"
+  by (import nets SEQ_TENDS)
+
+lemma LIM_TENDS: "ALL m1 m2 f x0 y0.
+   limpt (mtop m1) x0 re_universe -->
+   tends f y0 (mtop m2, tendsto (m1, x0)) =
+   (ALL e.
+       0 < e -->
+       (EX d. 0 < d &
+              (ALL x.
+                  0 < dist m1 (x, x0) & dist m1 (x, x0) <= d -->
+                  dist m2 (f x, y0) < e)))"
+  by (import nets LIM_TENDS)
+
+lemma LIM_TENDS2: "ALL m1 m2 f x0 y0.
+   limpt (mtop m1) x0 re_universe -->
+   tends f y0 (mtop m2, tendsto (m1, x0)) =
+   (ALL e.
+       0 < e -->
+       (EX d. 0 < d &
+              (ALL x.
+                  0 < dist m1 (x, x0) & dist m1 (x, x0) < d -->
+                  dist m2 (f x, y0) < e)))"
+  by (import nets LIM_TENDS2)
+
+lemma MR1_BOUNDED: "ALL g f.
+   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 x x0. 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 x x0. tends x x0 (mtop mr1, g) --> bounded (mr1, g) x"
+  by (import nets NET_CONV_BOUNDED)
+
+lemma NET_CONV_NZ: "ALL g x x0.
+   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 x x0.
+   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.
+   dorder g -->
+   (ALL x y.
+       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.
+   dorder g -->
+   (ALL x y.
+       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 k x. 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.
+   dorder g -->
+   (ALL x x0 y y0.
+       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.
+   dorder g -->
+   (ALL x x0.
+       tends x x0 (mtop mr1, g) = tends (%n. - x n) (- x0) (mtop mr1, g))"
+  by (import nets NET_NEG)
+
+lemma NET_SUB: "ALL g.
+   dorder g -->
+   (ALL x x0 y y0.
+       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.
+   dorder g -->
+   (ALL x y x0 y0.
+       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.
+   dorder g -->
+   (ALL x x0.
+       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.
+   dorder g -->
+   (ALL x x0 y y0.
+       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 x x0.
+   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.
+   dorder g -->
+   (ALL x x0 y y0.
+       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
+
+;setup_theory seq
+
+constdefs
+  "-->" :: "(nat => real) => real => bool" ("-->")
+  "--> == %x x0. tends x x0 (mtop mr1, nat_ge)"
+
+lemma tends_num_real: "ALL x x0. --> x x0 = tends x x0 (mtop mr1, nat_ge)"
+  by (import seq tends_num_real)
+
+lemma SEQ: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (op =::bool => bool => bool)
+           ((-->::(nat => real) => real => bool) x x0)
+           ((All::(real => bool) => bool)
+             (%e::real.
+                 (op -->::bool => bool => bool)
+                  ((op <::real => real => bool) (0::real) e)
+                  ((Ex::(nat => bool) => bool)
+                    (%N::nat.
+                        (All::(nat => bool) => bool)
+                         (%n::nat.
+                             (op -->::bool => bool => bool)
+                              ((op <=::nat => nat => bool) N n)
+                              ((op <::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) (x n) x0))
+                                e))))))))"
+  by (import seq SEQ)
+
+lemma SEQ_CONST: "ALL k. --> (%x. k) k"
+  by (import seq SEQ_CONST)
+
+lemma SEQ_ADD: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((-->::(nat => real) => real => bool) x x0)
+                       ((-->::(nat => real) => real => bool) y y0))
+                     ((-->::(nat => real) => real => bool)
+                       (%n::nat. (op +::real => real => real) (x n) (y n))
+                       ((op +::real => real => real) x0 y0))))))"
+  by (import seq SEQ_ADD)
+
+lemma SEQ_MUL: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((-->::(nat => real) => real => bool) x x0)
+                       ((-->::(nat => real) => real => bool) y y0))
+                     ((-->::(nat => real) => real => bool)
+                       (%n::nat. (op *::real => real => real) (x n) (y n))
+                       ((op *::real => real => real) x0 y0))))))"
+  by (import seq SEQ_MUL)
+
+lemma SEQ_NEG: "ALL x x0. --> x x0 = --> (%n. - x n) (- x0)"
+  by (import seq SEQ_NEG)
+
+lemma SEQ_INV: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((-->::(nat => real) => real => bool) x x0)
+             ((Not::bool => bool)
+               ((op =::real => real => bool) x0 (0::real))))
+           ((-->::(nat => real) => real => bool)
+             (%n::nat. (inverse::real => real) (x n))
+             ((inverse::real => real) x0))))"
+  by (import seq SEQ_INV)
+
+lemma SEQ_SUB: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((-->::(nat => real) => real => bool) x x0)
+                       ((-->::(nat => real) => real => bool) y y0))
+                     ((-->::(nat => real) => real => bool)
+                       (%n::nat. (op -::real => real => real) (x n) (y n))
+                       ((op -::real => real => real) x0 y0))))))"
+  by (import seq SEQ_SUB)
+
+lemma SEQ_DIV: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((-->::(nat => real) => real => bool) x x0)
+                       ((op &::bool => bool => bool)
+                         ((-->::(nat => real) => real => bool) y y0)
+                         ((Not::bool => bool)
+                           ((op =::real => real => bool) y0 (0::real)))))
+                     ((-->::(nat => real) => real => bool)
+                       (%n::nat. (op /::real => real => real) (x n) (y n))
+                       ((op /::real => real => real) x0 y0))))))"
+  by (import seq SEQ_DIV)
+
+lemma SEQ_UNIQ: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x1::real.
+          (All::(real => bool) => bool)
+           (%x2::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((-->::(nat => real) => real => bool) x x1)
+                  ((-->::(nat => real) => real => bool) x x2))
+                ((op =::real => real => bool) x1 x2))))"
+  by (import seq SEQ_UNIQ)
+
+constdefs
+  convergent :: "(nat => real) => bool" 
+  "convergent == %f. Ex (--> f)"
+
+lemma convergent: "ALL f. convergent f = Ex (--> f)"
+  by (import seq convergent)
+
+constdefs
+  cauchy :: "(nat => real) => bool" 
+  "(op ==::((nat => real) => bool) => ((nat => real) => bool) => prop)
+ (cauchy::(nat => real) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%e::real.
+          (op -->::bool => bool => bool)
+           ((op <::real => real => bool) (0::real) e)
+           ((Ex::(nat => bool) => bool)
+             (%N::nat.
+                 (All::(nat => bool) => bool)
+                  (%m::nat.
+                      (All::(nat => bool) => bool)
+                       (%n::nat.
+                           (op -->::bool => bool => bool)
+                            ((op &::bool => bool => bool)
+                              ((op <=::nat => nat => bool) N m)
+                              ((op <=::nat => nat => bool) N n))
+                            ((op <::real => real => bool)
+                              ((abs::real => real)
+                                ((op -::real => real => real) (f m) (f n)))
+                              e)))))))"
+
+lemma cauchy: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op =::bool => bool => bool) ((cauchy::(nat => real) => bool) f)
+      ((All::(real => bool) => bool)
+        (%e::real.
+            (op -->::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) e)
+             ((Ex::(nat => bool) => bool)
+               (%N::nat.
+                   (All::(nat => bool) => bool)
+                    (%m::nat.
+                        (All::(nat => bool) => bool)
+                         (%n::nat.
+                             (op -->::bool => bool => bool)
+                              ((op &::bool => bool => bool)
+                                ((op <=::nat => nat => bool) N m)
+                                ((op <=::nat => nat => bool) N n))
+                              ((op <::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) (f m)
+                                    (f n)))
+                                e))))))))"
+  by (import seq cauchy)
+
+constdefs
+  lim :: "(nat => real) => real" 
+  "lim == %f. Eps (--> f)"
+
+lemma lim: "ALL f. lim f = Eps (--> f)"
+  by (import seq lim)
+
+lemma SEQ_LIM: "ALL f. convergent f = --> f (lim f)"
+  by (import seq SEQ_LIM)
+
+constdefs
+  subseq :: "(nat => nat) => bool" 
+  "(op ==::((nat => nat) => bool) => ((nat => nat) => bool) => prop)
+ (subseq::(nat => nat) => bool)
+ (%f::nat => nat.
+     (All::(nat => bool) => bool)
+      (%m::nat.
+          (All::(nat => bool) => bool)
+           (%n::nat.
+               (op -->::bool => bool => bool)
+                ((op <::nat => nat => bool) m n)
+                ((op <::nat => nat => bool) (f m) (f n)))))"
+
+lemma subseq: "(All::((nat => nat) => bool) => bool)
+ (%f::nat => nat.
+     (op =::bool => bool => bool) ((subseq::(nat => nat) => bool) f)
+      ((All::(nat => bool) => bool)
+        (%m::nat.
+            (All::(nat => bool) => bool)
+             (%n::nat.
+                 (op -->::bool => bool => bool)
+                  ((op <::nat => nat => bool) m n)
+                  ((op <::nat => nat => bool) (f m) (f n))))))"
+  by (import seq subseq)
+
+lemma SUBSEQ_SUC: "ALL f. subseq f = (ALL n. f n < f (Suc n))"
+  by (import seq SUBSEQ_SUC)
+
+constdefs
+  mono :: "(nat => real) => bool" 
+  "(op ==::((nat => real) => bool) => ((nat => real) => bool) => prop)
+ (seq.mono::(nat => real) => bool)
+ (%f::nat => real.
+     (op |::bool => bool => bool)
+      ((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) (f m) (f n)))))
+      ((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) (f n) (f m))))))"
+
+lemma mono: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op =::bool => bool => bool) ((seq.mono::(nat => real) => bool) f)
+      ((op |::bool => bool => bool)
+        ((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) (f m) (f n)))))
+        ((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) (f n) (f m)))))))"
+  by (import seq mono)
+
+lemma MONO_SUC: "ALL f. 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 (s::nat => real) N::nat. EX k::real. ALL n<N. abs (s n) < k"
+  by (import seq MAX_LEMMA)
+
+lemma SEQ_BOUNDED: "ALL s. bounded (mr1, nat_ge) s = (EX k. ALL n. abs (s n) < k)"
+  by (import seq SEQ_BOUNDED)
+
+lemma SEQ_BOUNDED_2: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%k::real.
+          (All::(real => bool) => bool)
+           (%k'::real.
+               (op -->::bool => bool => bool)
+                ((All::(nat => bool) => bool)
+                  (%n::nat.
+                      (op &::bool => bool => bool)
+                       ((op <=::real => real => bool) k (f n))
+                       ((op <=::real => real => bool) (f n) k')))
+                ((bounded::real metric * (nat => nat => bool)
+                           => (nat => real) => bool)
+                  ((Pair::real metric
+                          => (nat => nat => bool)
+                             => real metric * (nat => nat => bool))
+                    (mr1::real metric) (nat_ge::nat => nat => bool))
+                  f))))"
+  by (import seq SEQ_BOUNDED_2)
+
+lemma SEQ_CBOUNDED: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool) ((cauchy::(nat => real) => bool) f)
+      ((bounded::real metric * (nat => nat => bool)
+                 => (nat => real) => bool)
+        ((Pair::real metric
+                => (nat => nat => bool)
+                   => real metric * (nat => nat => bool))
+          (mr1::real metric) (nat_ge::nat => nat => bool))
+        f))"
+  by (import seq SEQ_CBOUNDED)
+
+lemma SEQ_ICONV: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((bounded::real metric * (nat => nat => bool)
+                   => (nat => real) => bool)
+          ((Pair::real metric
+                  => (nat => nat => bool)
+                     => real metric * (nat => nat => bool))
+            (mr1::real metric) (nat_ge::nat => nat => bool))
+          f)
+        ((All::(nat => bool) => bool)
+          (%m::nat.
+              (All::(nat => bool) => bool)
+               (%n::nat.
+                   (op -->::bool => bool => bool)
+                    ((op <=::nat => nat => bool) n m)
+                    ((op <=::real => real => bool) (f n) (f m))))))
+      ((convergent::(nat => real) => bool) f))"
+  by (import seq SEQ_ICONV)
+
+lemma SEQ_NEG_CONV: "ALL f. convergent f = convergent (%n. - f n)"
+  by (import seq SEQ_NEG_CONV)
+
+lemma SEQ_NEG_BOUNDED: "ALL f. bounded (mr1, nat_ge) (%n. - f n) = bounded (mr1, nat_ge) f"
+  by (import seq SEQ_NEG_BOUNDED)
+
+lemma SEQ_BCONV: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((bounded::real metric * (nat => nat => bool)
+                   => (nat => real) => bool)
+          ((Pair::real metric
+                  => (nat => nat => bool)
+                     => real metric * (nat => nat => bool))
+            (mr1::real metric) (nat_ge::nat => nat => bool))
+          f)
+        ((seq.mono::(nat => real) => bool) f))
+      ((convergent::(nat => real) => bool) f))"
+  by (import seq SEQ_BCONV)
+
+lemma SEQ_MONOSUB: "ALL s. EX f. subseq f & seq.mono (%n. s (f n))"
+  by (import seq SEQ_MONOSUB)
+
+lemma SEQ_SBOUNDED: "(All::((nat => real) => bool) => bool)
+ (%s::nat => real.
+     (All::((nat => nat) => bool) => bool)
+      (%f::nat => nat.
+          (op -->::bool => bool => bool)
+           ((bounded::real metric * (nat => nat => bool)
+                      => (nat => real) => bool)
+             ((Pair::real metric
+                     => (nat => nat => bool)
+                        => real metric * (nat => nat => bool))
+               (mr1::real metric) (nat_ge::nat => nat => bool))
+             s)
+           ((bounded::real metric * (nat => nat => bool)
+                      => (nat => real) => bool)
+             ((Pair::real metric
+                     => (nat => nat => bool)
+                        => real metric * (nat => nat => bool))
+               (mr1::real metric) (nat_ge::nat => nat => bool))
+             (%n::nat. s (f n)))))"
+  by (import seq SEQ_SBOUNDED)
+
+lemma SEQ_SUBLE: "(All::((nat => nat) => bool) => bool)
+ (%f::nat => nat.
+     (op -->::bool => bool => bool) ((subseq::(nat => nat) => bool) f)
+      ((All::(nat => bool) => bool)
+        (%n::nat. (op <=::nat => nat => bool) n (f n))))"
+  by (import seq SEQ_SUBLE)
+
+lemma SEQ_DIRECT: "(All::((nat => nat) => bool) => bool)
+ (%f::nat => nat.
+     (op -->::bool => bool => bool) ((subseq::(nat => nat) => bool) f)
+      ((All::(nat => bool) => bool)
+        (%N1::nat.
+            (All::(nat => bool) => bool)
+             (%N2::nat.
+                 (Ex::(nat => bool) => bool)
+                  (%x::nat.
+                      (op &::bool => bool => bool)
+                       ((op <=::nat => nat => bool) N1 x)
+                       ((op <=::nat => nat => bool) N2 (f x)))))))"
+  by (import seq SEQ_DIRECT)
+
+lemma SEQ_CAUCHY: "ALL f. cauchy f = convergent f"
+  by (import seq SEQ_CAUCHY)
+
+lemma SEQ_LE: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((-->::(nat => real) => real => bool) f l)
+                       ((op &::bool => bool => bool)
+                         ((-->::(nat => real) => real => bool) g m)
+                         ((Ex::(nat => bool) => bool)
+                           (%x::nat.
+                               (All::(nat => bool) => bool)
+                                (%xa::nat.
+                                    (op -->::bool => bool => bool)
+                                     ((op <=::nat => nat => bool) x xa)
+                                     ((op <=::real => real => bool) (f xa)
+ (g xa)))))))
+                     ((op <=::real => real => bool) l m)))))"
+  by (import seq SEQ_LE)
+
+lemma SEQ_SUC: "ALL f l. --> f l = --> (%n. f (Suc n)) l"
+  by (import seq SEQ_SUC)
+
+lemma SEQ_ABS: "ALL f. --> (%n. abs (f n)) 0 = --> f 0"
+  by (import seq SEQ_ABS)
+
+lemma SEQ_ABS_IMP: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (op -->::bool => bool => bool)
+           ((-->::(nat => real) => real => bool) f l)
+           ((-->::(nat => real) => real => bool)
+             (%n::nat. (abs::real => real) (f n)) ((abs::real => real) l))))"
+  by (import seq SEQ_ABS_IMP)
+
+lemma SEQ_INV0: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool)
+      ((All::(real => bool) => bool)
+        (%y::real.
+            (Ex::(nat => bool) => bool)
+             (%N::nat.
+                 (All::(nat => bool) => bool)
+                  (%n::nat.
+                      (op -->::bool => bool => bool)
+                       ((op <=::nat => nat => bool) N n)
+                       ((op <::real => real => bool) y (f n))))))
+      ((-->::(nat => real) => real => bool)
+        (%n::nat. (inverse::real => real) (f n)) (0::real)))"
+  by (import seq SEQ_INV0)
+
+lemma SEQ_POWER_ABS: "(All::(real => bool) => bool)
+ (%c::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) ((abs::real => real) c) (1::real))
+      ((-->::(nat => real) => real => bool)
+        ((op ^::real => nat => real) ((abs::real => real) c)) (0::real)))"
+  by (import seq SEQ_POWER_ABS)
+
+lemma SEQ_POWER: "(All::(real => bool) => bool)
+ (%c::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) ((abs::real => real) c) (1::real))
+      ((-->::(nat => real) => real => bool) ((op ^::real => nat => real) c)
+        (0::real)))"
+  by (import seq SEQ_POWER)
+
+lemma NEST_LEMMA: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((All::(nat => bool) => bool)
+               (%n::nat.
+                   (op <=::real => real => bool) (f n)
+                    (f ((Suc::nat => nat) n))))
+             ((op &::bool => bool => bool)
+               ((All::(nat => bool) => bool)
+                 (%n::nat.
+                     (op <=::real => real => bool) (g ((Suc::nat => nat) n))
+                      (g n)))
+               ((All::(nat => bool) => bool)
+                 (%n::nat. (op <=::real => real => bool) (f n) (g n)))))
+           ((Ex::(real => bool) => bool)
+             (%l::real.
+                 (Ex::(real => bool) => bool)
+                  (%m::real.
+                      (op &::bool => bool => bool)
+                       ((op <=::real => real => bool) l m)
+                       ((op &::bool => bool => bool)
+                         ((op &::bool => bool => bool)
+                           ((All::(nat => bool) => bool)
+                             (%n::nat.
+                                 (op <=::real => real => bool) (f n) l))
+                           ((-->::(nat => real) => real => bool) f l))
+                         ((op &::bool => bool => bool)
+                           ((All::(nat => bool) => bool)
+                             (%n::nat.
+                                 (op <=::real => real => bool) m (g n)))
+                           ((-->::(nat => real) => real => bool) g m))))))))"
+  by (import seq NEST_LEMMA)
+
+lemma NEST_LEMMA_UNIQ: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((All::(nat => bool) => bool)
+               (%n::nat.
+                   (op <=::real => real => bool) (f n)
+                    (f ((Suc::nat => nat) n))))
+             ((op &::bool => bool => bool)
+               ((All::(nat => bool) => bool)
+                 (%n::nat.
+                     (op <=::real => real => bool) (g ((Suc::nat => nat) n))
+                      (g n)))
+               ((op &::bool => bool => bool)
+                 ((All::(nat => bool) => bool)
+                   (%n::nat. (op <=::real => real => bool) (f n) (g n)))
+                 ((-->::(nat => real) => real => bool)
+                   (%n::nat. (op -::real => real => real) (f n) (g n))
+                   (0::real)))))
+           ((Ex::(real => bool) => bool)
+             (%x::real.
+                 (op &::bool => bool => bool)
+                  ((op &::bool => bool => bool)
+                    ((All::(nat => bool) => bool)
+                      (%n::nat. (op <=::real => real => bool) (f n) x))
+                    ((-->::(nat => real) => real => bool) f x))
+                  ((op &::bool => bool => bool)
+                    ((All::(nat => bool) => bool)
+                      (%n::nat. (op <=::real => real => bool) x (g n)))
+                    ((-->::(nat => real) => real => bool) g x))))))"
+  by (import seq NEST_LEMMA_UNIQ)
+
+lemma BOLZANO_LEMMA: "(All::((real * real => bool) => bool) => bool)
+ (%P::real * real => bool.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((All::(real => bool) => bool)
+          (%a::real.
+              (All::(real => bool) => bool)
+               (%b::real.
+                   (All::(real => bool) => bool)
+                    (%c::real.
+                        (op -->::bool => bool => bool)
+                         ((op &::bool => bool => bool)
+                           ((op <=::real => real => bool) a b)
+                           ((op &::bool => bool => bool)
+                             ((op <=::real => real => bool) b c)
+                             ((op &::bool => bool => bool)
+                               (P ((Pair::real => real => real * real) a b))
+                               (P ((Pair::real => real => real * real) b
+                                    c)))))
+                         (P ((Pair::real => real => real * real) a c))))))
+        ((All::(real => bool) => bool)
+          (%x::real.
+              (Ex::(real => bool) => bool)
+               (%d::real.
+                   (op &::bool => bool => bool)
+                    ((op <::real => real => bool) (0::real) d)
+                    ((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 x)
+                                  ((op &::bool => bool => bool)
+                                    ((op <=::real => real => bool) x b)
+                                    ((op <::real => real => bool)
+((op -::real => real => real) b a) d)))
+                                (P ((Pair::real => real => real * real) a
+                                     b)))))))))
+      ((All::(real => bool) => bool)
+        (%a::real.
+            (All::(real => bool) => bool)
+             (%b::real.
+                 (op -->::bool => bool => bool)
+                  ((op <=::real => real => bool) a b)
+                  (P ((Pair::real => real => real * real) a b))))))"
+  by (import seq BOLZANO_LEMMA)
+
+constdefs
+  sums :: "(nat => real) => real => bool" 
+  "sums == %f. --> (%n. real.sum (0, n) f)"
+
+lemma sums: "ALL f s. sums f s = --> (%n. real.sum (0, n) f) s"
+  by (import seq sums)
+
+constdefs
+  summable :: "(nat => real) => bool" 
+  "summable == %f. Ex (sums f)"
+
+lemma summable: "ALL f. summable f = Ex (sums f)"
+  by (import seq summable)
+
+constdefs
+  suminf :: "(nat => real) => real" 
+  "suminf == %f. Eps (sums f)"
+
+lemma suminf: "ALL f. suminf f = Eps (sums f)"
+  by (import seq suminf)
+
+lemma SUM_SUMMABLE: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (op -->::bool => bool => bool)
+           ((sums::(nat => real) => real => bool) f l)
+           ((summable::(nat => real) => bool) f)))"
+  by (import seq SUM_SUMMABLE)
+
+lemma SUMMABLE_SUM: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool) ((summable::(nat => real) => bool) f)
+      ((sums::(nat => real) => real => bool) f
+        ((suminf::(nat => real) => real) f)))"
+  by (import seq SUMMABLE_SUM)
+
+lemma SUM_UNIQ: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((sums::(nat => real) => real => bool) f x)
+           ((op =::real => real => bool) x
+             ((suminf::(nat => real) => real) f))))"
+  by (import seq SUM_UNIQ)
+
+lemma SER_0: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((All::(nat => bool) => bool)
+             (%m::nat.
+                 (op -->::bool => bool => bool)
+                  ((op <=::nat => nat => bool) n m)
+                  ((op =::real => real => bool) (f m) (0::real))))
+           ((sums::(nat => real) => real => bool) f
+             ((real.sum::nat * nat => (nat => real) => real)
+               ((Pair::nat => nat => nat * nat) (0::nat) n) f))))"
+  by (import seq SER_0)
+
+lemma SER_POS_LE: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((summable::(nat => real) => bool) f)
+             ((All::(nat => bool) => bool)
+               (%m::nat.
+                   (op -->::bool => bool => bool)
+                    ((op <=::nat => nat => bool) n m)
+                    ((op <=::real => real => bool) (0::real) (f m)))))
+           ((op <=::real => real => bool)
+             ((real.sum::nat * nat => (nat => real) => real)
+               ((Pair::nat => nat => nat * nat) (0::nat) n) f)
+             ((suminf::(nat => real) => real) f))))"
+  by (import seq SER_POS_LE)
+
+lemma SER_POS_LT: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((summable::(nat => real) => bool) f)
+             ((All::(nat => bool) => bool)
+               (%m::nat.
+                   (op -->::bool => bool => bool)
+                    ((op <=::nat => nat => bool) n m)
+                    ((op <::real => real => bool) (0::real) (f m)))))
+           ((op <::real => real => bool)
+             ((real.sum::nat * nat => (nat => real) => real)
+               ((Pair::nat => nat => nat * nat) (0::nat) n) f)
+             ((suminf::(nat => real) => real) f))))"
+  by (import seq SER_POS_LT)
+
+lemma SER_GROUP: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(nat => bool) => bool)
+      (%k::nat.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((summable::(nat => real) => bool) f)
+             ((op <::nat => nat => bool) (0::nat) k))
+           ((sums::(nat => real) => real => bool)
+             (%n::nat.
+                 (real.sum::nat * nat => (nat => real) => real)
+                  ((Pair::nat => nat => nat * nat)
+                    ((op *::nat => nat => nat) n k) k)
+                  f)
+             ((suminf::(nat => real) => real) f))))"
+  by (import seq SER_GROUP)
+
+lemma SER_PAIR: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool) ((summable::(nat => real) => bool) f)
+      ((sums::(nat => real) => real => bool)
+        (%n::nat.
+            (real.sum::nat * nat => (nat => real) => real)
+             ((Pair::nat => nat => nat * nat)
+               ((op *::nat => nat => nat)
+                 ((number_of::bin => nat)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool)))
+                 n)
+               ((number_of::bin => nat)
+                 ((op BIT::bin => bool => bin)
+                   ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                     (True::bool))
+                   (False::bool))))
+             f)
+        ((suminf::(nat => real) => real) f)))"
+  by (import seq SER_PAIR)
+
+lemma SER_OFFSET: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool) ((summable::(nat => real) => bool) f)
+      ((All::(nat => bool) => bool)
+        (%k::nat.
+            (sums::(nat => real) => real => bool)
+             (%n::nat. f ((op +::nat => nat => nat) n k))
+             ((op -::real => real => real)
+               ((suminf::(nat => real) => real) f)
+               ((real.sum::nat * nat => (nat => real) => real)
+                 ((Pair::nat => nat => nat * nat) (0::nat) k) f)))))"
+  by (import seq SER_OFFSET)
+
+lemma SER_POS_LT_PAIR: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((summable::(nat => real) => bool) f)
+             ((All::(nat => bool) => bool)
+               (%d::nat.
+                   (op <::real => real => bool) (0::real)
+                    ((op +::real => real => real)
+                      (f ((op +::nat => nat => nat) n
+                           ((op *::nat => nat => nat)
+                             ((number_of::bin => nat)
+                               ((op BIT::bin => bool => bin)
+                                 ((op BIT::bin => bool => bin)
+                                   (bin.Pls::bin) (True::bool))
+                                 (False::bool)))
+                             d)))
+                      (f ((op +::nat => nat => nat) n
+                           ((op +::nat => nat => nat)
+                             ((op *::nat => nat => nat)
+                               ((number_of::bin => nat)
+                                 ((op BIT::bin => bool => bin)
+                                   ((op BIT::bin => bool => bin)
+                                     (bin.Pls::bin) (True::bool))
+                                   (False::bool)))
+                               d)
+                             (1::nat))))))))
+           ((op <::real => real => bool)
+             ((real.sum::nat * nat => (nat => real) => real)
+               ((Pair::nat => nat => nat * nat) (0::nat) n) f)
+             ((suminf::(nat => real) => real) f))))"
+  by (import seq SER_POS_LT_PAIR)
+
+lemma SER_ADD: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((sums::(nat => real) => real => bool) x x0)
+                       ((sums::(nat => real) => real => bool) y y0))
+                     ((sums::(nat => real) => real => bool)
+                       (%n::nat. (op +::real => real => real) (x n) (y n))
+                       ((op +::real => real => real) x0 y0))))))"
+  by (import seq SER_ADD)
+
+lemma SER_CMUL: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::(real => bool) => bool)
+           (%c::real.
+               (op -->::bool => bool => bool)
+                ((sums::(nat => real) => real => bool) x x0)
+                ((sums::(nat => real) => real => bool)
+                  (%n::nat. (op *::real => real => real) c (x n))
+                  ((op *::real => real => real) c x0)))))"
+  by (import seq SER_CMUL)
+
+lemma SER_NEG: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (op -->::bool => bool => bool)
+           ((sums::(nat => real) => real => bool) x x0)
+           ((sums::(nat => real) => real => bool)
+             (%xa::nat. (uminus::real => real) (x xa))
+             ((uminus::real => real) x0))))"
+  by (import seq SER_NEG)
+
+lemma SER_SUB: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::((nat => real) => bool) => bool)
+           (%y::nat => real.
+               (All::(real => bool) => bool)
+                (%y0::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((sums::(nat => real) => real => bool) x x0)
+                       ((sums::(nat => real) => real => bool) y y0))
+                     ((sums::(nat => real) => real => bool)
+                       (%xa::nat.
+                           (op -::real => real => real) (x xa) (y xa))
+                       ((op -::real => real => real) x0 y0))))))"
+  by (import seq SER_SUB)
+
+lemma SER_CDIV: "(All::((nat => real) => bool) => bool)
+ (%x::nat => real.
+     (All::(real => bool) => bool)
+      (%x0::real.
+          (All::(real => bool) => bool)
+           (%c::real.
+               (op -->::bool => bool => bool)
+                ((sums::(nat => real) => real => bool) x x0)
+                ((sums::(nat => real) => real => bool)
+                  (%xa::nat. (op /::real => real => real) (x xa) c)
+                  ((op /::real => real => real) x0 c)))))"
+  by (import seq SER_CDIV)
+
+lemma SER_CAUCHY: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op =::bool => bool => bool) ((summable::(nat => real) => bool) f)
+      ((All::(real => bool) => bool)
+        (%e::real.
+            (op -->::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) e)
+             ((Ex::(nat => bool) => bool)
+               (%N::nat.
+                   (All::(nat => bool) => bool)
+                    (%m::nat.
+                        (All::(nat => bool) => bool)
+                         (%n::nat.
+                             (op -->::bool => bool => bool)
+                              ((op <=::nat => nat => bool) N m)
+                              ((op <::real => real => bool)
+                                ((abs::real => real)
+                                  ((real.sum::nat * nat
+        => (nat => real) => real)
+                                    ((Pair::nat => nat => nat * nat) m n)
+                                    f))
+                                e))))))))"
+  by (import seq SER_CAUCHY)
+
+lemma SER_ZERO: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool) ((summable::(nat => real) => bool) f)
+      ((-->::(nat => real) => real => bool) f (0::real)))"
+  by (import seq SER_ZERO)
+
+lemma SER_COMPAR: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((Ex::(nat => bool) => bool)
+               (%x::nat.
+                   (All::(nat => bool) => bool)
+                    (%xa::nat.
+                        (op -->::bool => bool => bool)
+                         ((op <=::nat => nat => bool) x xa)
+                         ((op <=::real => real => bool)
+                           ((abs::real => real) (f xa)) (g xa)))))
+             ((summable::(nat => real) => bool) g))
+           ((summable::(nat => real) => bool) f)))"
+  by (import seq SER_COMPAR)
+
+lemma SER_COMPARA: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((Ex::(nat => bool) => bool)
+               (%x::nat.
+                   (All::(nat => bool) => bool)
+                    (%xa::nat.
+                        (op -->::bool => bool => bool)
+                         ((op <=::nat => nat => bool) x xa)
+                         ((op <=::real => real => bool)
+                           ((abs::real => real) (f xa)) (g xa)))))
+             ((summable::(nat => real) => bool) g))
+           ((summable::(nat => real) => bool)
+             (%k::nat. (abs::real => real) (f k)))))"
+  by (import seq SER_COMPARA)
+
+lemma SER_LE: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((All::(nat => bool) => bool)
+               (%n::nat. (op <=::real => real => bool) (f n) (g n)))
+             ((op &::bool => bool => bool)
+               ((summable::(nat => real) => bool) f)
+               ((summable::(nat => real) => bool) g)))
+           ((op <=::real => real => bool)
+             ((suminf::(nat => real) => real) f)
+             ((suminf::(nat => real) => real) g))))"
+  by (import seq SER_LE)
+
+lemma SER_LE2: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((nat => real) => bool) => bool)
+      (%g::nat => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((All::(nat => bool) => bool)
+               (%n::nat.
+                   (op <=::real => real => bool) ((abs::real => real) (f n))
+                    (g n)))
+             ((summable::(nat => real) => bool) g))
+           ((op &::bool => bool => bool)
+             ((summable::(nat => real) => bool) f)
+             ((op <=::real => real => bool)
+               ((suminf::(nat => real) => real) f)
+               ((suminf::(nat => real) => real) g)))))"
+  by (import seq SER_LE2)
+
+lemma SER_ACONV: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool)
+      ((summable::(nat => real) => bool)
+        (%n::nat. (abs::real => real) (f n)))
+      ((summable::(nat => real) => bool) f))"
+  by (import seq SER_ACONV)
+
+lemma SER_ABS: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (op -->::bool => bool => bool)
+      ((summable::(nat => real) => bool)
+        (%n::nat. (abs::real => real) (f n)))
+      ((op <=::real => real => bool)
+        ((abs::real => real) ((suminf::(nat => real) => real) f))
+        ((suminf::(nat => real) => real)
+          (%n::nat. (abs::real => real) (f n)))))"
+  by (import seq SER_ABS)
+
+lemma GP_FINITE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((Not::bool => bool) ((op =::real => real => bool) x (1::real)))
+      ((All::(nat => bool) => bool)
+        (%n::nat.
+            (op =::real => real => bool)
+             ((real.sum::nat * nat => (nat => real) => real)
+               ((Pair::nat => nat => nat * nat) (0::nat) n)
+               ((op ^::real => nat => real) x))
+             ((op /::real => real => real)
+               ((op -::real => real => real)
+                 ((op ^::real => nat => real) x n) (1::real))
+               ((op -::real => real => real) x (1::real))))))"
+  by (import seq GP_FINITE)
+
+lemma GP: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) ((abs::real => real) x) (1::real))
+      ((sums::(nat => real) => real => bool) ((op ^::real => nat => real) x)
+        ((inverse::real => real)
+          ((op -::real => real => real) (1::real) 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::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%c::real.
+          (All::(nat => bool) => bool)
+           (%N::nat.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <::real => real => bool) c (1::real))
+                  ((All::(nat => bool) => bool)
+                    (%n::nat.
+                        (op -->::bool => bool => bool)
+                         ((op <=::nat => nat => bool) N n)
+                         ((op <=::real => real => bool)
+                           ((abs::real => real) (f ((Suc::nat => nat) n)))
+                           ((op *::real => real => real) c
+                             ((abs::real => real) (f n)))))))
+                ((summable::(nat => real) => bool) f))))"
+  by (import seq SER_RATIO)
+
+;end_setup
+
+;setup_theory lim
+
+constdefs
+  tends_real_real :: "(real => real) => real => real => bool" 
+  "tends_real_real == %f l x0. tends f l (mtop mr1, tendsto (mr1, x0))"
+
+lemma tends_real_real: "ALL f l x0. tends_real_real f l x0 = tends f l (mtop mr1, tendsto (mr1, x0))"
+  by (import lim tends_real_real)
+
+lemma LIM: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%y0::real.
+          (All::(real => bool) => bool)
+           (%x0::real.
+               (op =::bool => bool => bool)
+                ((tends_real_real::(real => real) => real => real => bool) f
+                  y0 x0)
+                ((All::(real => bool) => bool)
+                  (%e::real.
+                      (op -->::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) e)
+                       ((Ex::(real => bool) => bool)
+                         (%d::real.
+                             (op &::bool => bool => bool)
+                              ((op <::real => real => bool) (0::real) d)
+                              ((All::(real => bool) => bool)
+                                (%x::real.
+                                    (op -->::bool => bool => bool)
+                                     ((op &::bool => bool => bool)
+ ((op <::real => real => bool) (0::real)
+   ((abs::real => real) ((op -::real => real => real) x x0)))
+ ((op <::real => real => bool)
+   ((abs::real => real) ((op -::real => real => real) x x0)) d))
+                                     ((op <::real => real => bool)
+ ((abs::real => real) ((op -::real => real => real) (f x) y0)) e))))))))))"
+  by (import lim LIM)
+
+lemma LIM_CONST: "ALL k. All (tends_real_real (%x. k) k)"
+  by (import lim LIM_CONST)
+
+lemma LIM_ADD: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              f l x)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              g m x))
+                          ((tends_real_real::(real => real)
+       => real => real => bool)
+                            (%x::real.
+                                (op +::real => real => real) (f x) (g x))
+                            ((op +::real => real => real) l m) x))))))"
+  by (import lim LIM_ADD)
+
+lemma LIM_MUL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              f l x)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              g m x))
+                          ((tends_real_real::(real => real)
+       => real => real => bool)
+                            (%x::real.
+                                (op *::real => real => real) (f x) (g x))
+                            ((op *::real => real => real) l m) x))))))"
+  by (import lim LIM_MUL)
+
+lemma LIM_NEG: "ALL f l x. tends_real_real f l x = tends_real_real (%x. - f x) (- l) x"
+  by (import lim LIM_NEG)
+
+lemma LIM_INV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((tends_real_real::(real => real) => real => real => bool)
+                    f l x)
+                  ((Not::bool => bool)
+                    ((op =::real => real => bool) l (0::real))))
+                ((tends_real_real::(real => real) => real => real => bool)
+                  (%x::real. (inverse::real => real) (f x))
+                  ((inverse::real => real) l) x))))"
+  by (import lim LIM_INV)
+
+lemma LIM_SUB: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              f l x)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              g m x))
+                          ((tends_real_real::(real => real)
+       => real => real => bool)
+                            (%x::real.
+                                (op -::real => real => real) (f x) (g x))
+                            ((op -::real => real => real) l m) x))))))"
+  by (import lim LIM_SUB)
+
+lemma LIM_DIV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((tends_real_real::(real => real)
+         => real => real => bool)
+                              f l x)
+                            ((op &::bool => bool => bool)
+                              ((tends_real_real::(real => real)
+           => real => real => bool)
+                                g m x)
+                              ((Not::bool => bool)
+                                ((op =::real => real => bool) m
+                                  (0::real)))))
+                          ((tends_real_real::(real => real)
+       => real => real => bool)
+                            (%x::real.
+                                (op /::real => real => real) (f x) (g x))
+                            ((op /::real => real => real) l m) x))))))"
+  by (import lim LIM_DIV)
+
+lemma LIM_NULL: "ALL f l x. 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. tends_real_real (%x. x) x0 x0"
+  by (import lim LIM_X)
+
+lemma LIM_UNIQ: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%m::real.
+               (All::(real => bool) => bool)
+                (%x::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         f l x)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         f m x))
+                     ((op =::real => real => bool) l m)))))"
+  by (import lim LIM_UNIQ)
+
+lemma LIM_EQUAL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%x0::real.
+                    (op -->::bool => bool => bool)
+                     ((All::(real => bool) => bool)
+                       (%x::real.
+                           (op -->::bool => bool => bool)
+                            ((Not::bool => bool)
+                              ((op =::real => real => bool) x x0))
+                            ((op =::real => real => bool) (f x) (g x))))
+                     ((op =::bool => bool => bool)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         f l x0)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         g l x0))))))"
+  by (import lim LIM_EQUAL)
+
+lemma LIM_TRANSFORM: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x0::real.
+               (All::(real => bool) => bool)
+                (%l::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         (%x::real.
+                             (op -::real => real => real) (f x) (g x))
+                         (0::real) x0)
+                       ((tends_real_real::(real => real)
+    => real => real => bool)
+                         g l x0))
+                     ((tends_real_real::(real => real)
+  => real => real => bool)
+                       f l x0)))))"
+  by (import lim LIM_TRANSFORM)
+
+constdefs
+  diffl :: "(real => real) => real => real => bool" 
+  "diffl == %f l x. tends_real_real (%h. (f (x + h) - f x) / h) l 0"
+
+lemma diffl: "ALL f l x. diffl f l x = tends_real_real (%h. (f (x + h) - f x) / h) l 0"
+  by (import lim diffl)
+
+constdefs
+  contl :: "(real => real) => real => bool" 
+  "contl == %f x. tends_real_real (%h. f (x + h)) (f x) 0"
+
+lemma contl: "ALL f x. contl f x = tends_real_real (%h. f (x + h)) (f x) 0"
+  by (import lim contl)
+
+constdefs
+  differentiable :: "(real => real) => real => bool" 
+  "differentiable == %f x. EX l. diffl f l x"
+
+lemma differentiable: "ALL f x. differentiable f x = (EX l. diffl f l x)"
+  by (import lim differentiable)
+
+lemma DIFF_UNIQ: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%m::real.
+               (All::(real => bool) => bool)
+                (%x::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((diffl::(real => real) => real => real => bool) f l
+                         x)
+                       ((diffl::(real => real) => real => real => bool) f m
+                         x))
+                     ((op =::real => real => bool) l m)))))"
+  by (import lim DIFF_UNIQ)
+
+lemma DIFF_CONT: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((diffl::(real => real) => real => real => bool) f l x)
+                ((contl::(real => real) => real => bool) f x))))"
+  by (import lim DIFF_CONT)
+
+lemma CONTL_LIM: "ALL f x. contl f x = tends_real_real f (f x) x"
+  by (import lim CONTL_LIM)
+
+lemma DIFF_CARAT: "ALL f l x.
+   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. All (contl (%x. k))"
+  by (import lim CONT_CONST)
+
+lemma CONT_ADD: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((contl::(real => real) => real => bool) f x)
+                  ((contl::(real => real) => real => bool) g x))
+                ((contl::(real => real) => real => bool)
+                  (%x::real. (op +::real => real => real) (f x) (g x)) x))))"
+  by (import lim CONT_ADD)
+
+lemma CONT_MUL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((contl::(real => real) => real => bool) f x)
+                  ((contl::(real => real) => real => bool) g x))
+                ((contl::(real => real) => real => bool)
+                  (%x::real. (op *::real => real => real) (f x) (g x)) x))))"
+  by (import lim CONT_MUL)
+
+lemma CONT_NEG: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((contl::(real => real) => real => bool) f x)
+           ((contl::(real => real) => real => bool)
+             (%x::real. (uminus::real => real) (f x)) x)))"
+  by (import lim CONT_NEG)
+
+lemma CONT_INV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((contl::(real => real) => real => bool) f x)
+             ((Not::bool => bool)
+               ((op =::real => real => bool) (f x) (0::real))))
+           ((contl::(real => real) => real => bool)
+             (%x::real. (inverse::real => real) (f x)) x)))"
+  by (import lim CONT_INV)
+
+lemma CONT_SUB: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((contl::(real => real) => real => bool) f x)
+                  ((contl::(real => real) => real => bool) g x))
+                ((contl::(real => real) => real => bool)
+                  (%x::real. (op -::real => real => real) (f x) (g x)) x))))"
+  by (import lim CONT_SUB)
+
+lemma CONT_DIV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((contl::(real => real) => real => bool) f x)
+                  ((op &::bool => bool => bool)
+                    ((contl::(real => real) => real => bool) g x)
+                    ((Not::bool => bool)
+                      ((op =::real => real => bool) (g x) (0::real)))))
+                ((contl::(real => real) => real => bool)
+                  (%x::real. (op /::real => real => real) (f x) (g x)) x))))"
+  by (import lim CONT_DIV)
+
+lemma CONT_COMPOSE: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((contl::(real => real) => real => bool) f x)
+                  ((contl::(real => real) => real => bool) g (f x)))
+                ((contl::(real => real) => real => bool) (%x::real. g (f x))
+                  x))))"
+  by (import lim CONT_COMPOSE)
+
+lemma IVT: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(real => bool) => bool)
+           (%b::real.
+               (All::(real => bool) => bool)
+                (%y::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <=::real => real => bool) a b)
+                       ((op &::bool => bool => bool)
+                         ((op &::bool => bool => bool)
+                           ((op <=::real => real => bool) (f a) y)
+                           ((op <=::real => real => bool) y (f 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)
+                       (%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) (f x) y))))))))"
+  by (import lim IVT)
+
+lemma IVT2: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(real => bool) => bool)
+           (%b::real.
+               (All::(real => bool) => bool)
+                (%y::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <=::real => real => bool) a b)
+                       ((op &::bool => bool => bool)
+                         ((op &::bool => bool => bool)
+                           ((op <=::real => real => bool) (f b) y)
+                           ((op <=::real => real => bool) y (f a)))
+                         ((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)
+                       (%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) (f x) y))))))))"
+  by (import lim IVT2)
+
+lemma DIFF_CONST: "ALL k. All (diffl (%x. k) 0)"
+  by (import lim DIFF_CONST)
+
+lemma DIFF_ADD: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((diffl::(real => real) => real => real => bool)
+                              f l x)
+                            ((diffl::(real => real) => real => real => bool)
+                              g m x))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real.
+                                (op +::real => real => real) (f x) (g x))
+                            ((op +::real => real => real) l m) x))))))"
+  by (import lim DIFF_ADD)
+
+lemma DIFF_MUL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((diffl::(real => real) => real => real => bool)
+                              f l x)
+                            ((diffl::(real => real) => real => real => bool)
+                              g m x))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real.
+                                (op *::real => real => real) (f x) (g x))
+                            ((op +::real => real => real)
+                              ((op *::real => real => real) l (g x))
+                              ((op *::real => real => real) m (f x)))
+                            x))))))"
+  by (import lim DIFF_MUL)
+
+lemma DIFF_CMUL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%c::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%x::real.
+                    (op -->::bool => bool => bool)
+                     ((diffl::(real => real) => real => real => bool) f l x)
+                     ((diffl::(real => real) => real => real => bool)
+                       (%x::real. (op *::real => real => real) c (f x))
+                       ((op *::real => real => real) c l) x)))))"
+  by (import lim DIFF_CMUL)
+
+lemma DIFF_NEG: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((diffl::(real => real) => real => real => bool) f l x)
+                ((diffl::(real => real) => real => real => bool)
+                  (%x::real. (uminus::real => real) (f x))
+                  ((uminus::real => real) l) x))))"
+  by (import lim DIFF_NEG)
+
+lemma DIFF_SUB: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((diffl::(real => real) => real => real => bool)
+                              f l x)
+                            ((diffl::(real => real) => real => real => bool)
+                              g m x))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real.
+                                (op -::real => real => real) (f x) (g x))
+                            ((op -::real => real => real) l m) x))))))"
+  by (import lim DIFF_SUB)
+
+lemma DIFF_CHAIN: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((diffl::(real => real) => real => real => bool)
+                              f l (g x))
+                            ((diffl::(real => real) => real => real => bool)
+                              g m x))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real. f (g x))
+                            ((op *::real => real => real) l m) x))))))"
+  by (import lim DIFF_CHAIN)
+
+lemma DIFF_X: "All (diffl (%x. x) 1)"
+  by (import lim DIFF_X)
+
+lemma DIFF_POW: "ALL n x. diffl (%x. x ^ n) (real n * x ^ (n - 1)) x"
+  by (import lim DIFF_POW)
+
+lemma DIFF_XM1: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((Not::bool => bool) ((op =::real => real => bool) x (0::real)))
+      ((diffl::(real => real) => real => real => bool)
+        (inverse::real => real)
+        ((uminus::real => real)
+          ((op ^::real => nat => real) ((inverse::real => real) x)
+            ((number_of::bin => nat)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))
+        x))"
+  by (import lim DIFF_XM1)
+
+lemma DIFF_INV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%l::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((Not::bool => bool)
+                    ((op =::real => real => bool) (f x) (0::real))))
+                ((diffl::(real => real) => real => real => bool)
+                  (%x::real. (inverse::real => real) (f x))
+                  ((uminus::real => real)
+                    ((op /::real => real => real) l
+                      ((op ^::real => nat => real) (f x)
+                        ((number_of::bin => nat)
+                          ((op BIT::bin => bool => bin)
+                            ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                              (True::bool))
+                            (False::bool))))))
+                  x))))"
+  by (import lim DIFF_INV)
+
+lemma DIFF_DIV: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%m::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((diffl::(real => real) => real => real => bool)
+                              f l x)
+                            ((op &::bool => bool => bool)
+                              ((diffl::(real => real)
+ => real => real => bool)
+                                g m x)
+                              ((Not::bool => bool)
+                                ((op =::real => real => bool) (g x)
+                                  (0::real)))))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real.
+                                (op /::real => real => real) (f x) (g x))
+                            ((op /::real => real => real)
+                              ((op -::real => real => real)
+                                ((op *::real => real => real) l (g x))
+                                ((op *::real => real => real) m (f x)))
+                              ((op ^::real => nat => real) (g x)
+                                ((number_of::bin => nat)
+                                  ((op BIT::bin => bool => bin)
+                                    ((op BIT::bin => bool => bin)
+(bin.Pls::bin) (True::bool))
+                                    (False::bool)))))
+                            x))))))"
+  by (import lim DIFF_DIV)
+
+lemma DIFF_SUM: "(All::((nat => real => real) => bool) => bool)
+ (%f::nat => real => real.
+     (All::((nat => real => real) => bool) => bool)
+      (%f'::nat => real => real.
+          (All::(nat => bool) => bool)
+           (%m::nat.
+               (All::(nat => bool) => bool)
+                (%n::nat.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (op -->::bool => bool => bool)
+                          ((All::(nat => bool) => bool)
+                            (%r::nat.
+                                (op -->::bool => bool => bool)
+                                 ((op &::bool => bool => bool)
+                                   ((op <=::nat => nat => bool) m r)
+                                   ((op <::nat => nat => bool) r
+                                     ((op +::nat => nat => nat) m n)))
+                                 ((diffl::(real => real)
+    => real => real => bool)
+                                   (f r) (f' r x) x)))
+                          ((diffl::(real => real) => real => real => bool)
+                            (%x::real.
+                                (real.sum::nat * nat
+     => (nat => real) => real)
+                                 ((Pair::nat => nat => nat * nat) m n)
+                                 (%n::nat. f n x))
+                            ((real.sum::nat * nat => (nat => real) => real)
+                              ((Pair::nat => nat => nat * nat) m n)
+                              (%r::nat. f' r x))
+                            x))))))"
+  by (import lim DIFF_SUM)
+
+lemma CONT_BOUNDED: "(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.
+                      (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)))))))"
+  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))))))))))))"
+  by (import lim CONT_HASSUP)
+
+lemma CONT_ATTAINS: "(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)
+                  (%x::real.
+                      (op &::bool => bool => bool)
+                       ((All::(real => bool) => bool)
+                         (%xa::real.
+                             (op -->::bool => bool => bool)
+                              ((op &::bool => bool => bool)
+                                ((op <=::real => real => bool) a xa)
+                                ((op <=::real => real => bool) xa b))
+                              ((op <=::real => real => bool) (f xa) x)))
+                       ((Ex::(real => bool) => bool)
+                         (%xa::real.
+                             (op &::bool => bool => bool)
+                              ((op <=::real => real => bool) a xa)
+                              ((op &::bool => bool => bool)
+                                ((op <=::real => real => bool) xa b)
+                                ((op =::real => real => bool) (f xa)
+                                  x)))))))))"
+  by (import lim CONT_ATTAINS)
+
+lemma CONT_ATTAINS2: "(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)
+                  (%x::real.
+                      (op &::bool => bool => bool)
+                       ((All::(real => bool) => bool)
+                         (%xa::real.
+                             (op -->::bool => bool => bool)
+                              ((op &::bool => bool => bool)
+                                ((op <=::real => real => bool) a xa)
+                                ((op <=::real => real => bool) xa b))
+                              ((op <=::real => real => bool) x (f xa))))
+                       ((Ex::(real => bool) => bool)
+                         (%xa::real.
+                             (op &::bool => bool => bool)
+                              ((op <=::real => real => bool) a xa)
+                              ((op &::bool => bool => bool)
+                                ((op <=::real => real => bool) xa b)
+                                ((op =::real => real => bool) (f xa)
+                                  x)))))))))"
+  by (import lim CONT_ATTAINS2)
+
+lemma CONT_ATTAINS_ALL: "(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)
+                  (%x::real.
+                      (Ex::(real => bool) => bool)
+                       (%M::real.
+                           (op &::bool => bool => bool)
+                            ((op <=::real => real => bool) x M)
+                            ((op &::bool => bool => bool)
+                              ((All::(real => bool) => bool)
+                                (%y::real.
+                                    (op -->::bool => bool => bool)
+                                     ((op &::bool => bool => bool)
+ ((op <=::real => real => bool) x y) ((op <=::real => real => bool) y 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) (f x) y))))))
+                              ((All::(real => bool) => bool)
+                                (%xa::real.
+                                    (op -->::bool => bool => bool)
+                                     ((op &::bool => bool => bool)
+ ((op <=::real => real => bool) a xa) ((op <=::real => real => bool) xa b))
+                                     ((op &::bool => bool => bool)
+ ((op <=::real => real => bool) x (f xa))
+ ((op <=::real => real => bool) (f xa) M)))))))))))"
+  by (import lim CONT_ATTAINS_ALL)
+
+lemma DIFF_LINC: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((op <::real => real => bool) (0::real) l))
+                ((Ex::(real => bool) => bool)
+                  (%d::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((All::(real => bool) => bool)
+                         (%h::real.
+                             (op -->::bool => bool => bool)
+                              ((op &::bool => bool => bool)
+                                ((op <::real => real => bool) (0::real) h)
+                                ((op <::real => real => bool) h d))
+                              ((op <::real => real => bool) (f x)
+                                (f ((op +::real => real => real) x
+                                     h))))))))))"
+  by (import lim DIFF_LINC)
+
+lemma DIFF_LDEC: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((op <::real => real => bool) l (0::real)))
+                ((Ex::(real => bool) => bool)
+                  (%d::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((All::(real => bool) => bool)
+                         (%h::real.
+                             (op -->::bool => bool => bool)
+                              ((op &::bool => bool => bool)
+                                ((op <::real => real => bool) (0::real) h)
+                                ((op <::real => real => bool) h d))
+                              ((op <::real => real => bool) (f x)
+                                (f ((op -::real => real => real) x
+                                     h))))))))))"
+  by (import lim DIFF_LDEC)
+
+lemma DIFF_LMAX: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((Ex::(real => bool) => bool)
+                    (%d::real.
+                        (op &::bool => bool => bool)
+                         ((op <::real => real => bool) (0::real) d)
+                         ((All::(real => bool) => bool)
+                           (%y::real.
+                               (op -->::bool => bool => bool)
+                                ((op <::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) x y))
+                                  d)
+                                ((op <=::real => real => bool) (f y)
+                                  (f x)))))))
+                ((op =::real => real => bool) l (0::real)))))"
+  by (import lim DIFF_LMAX)
+
+lemma DIFF_LMIN: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((Ex::(real => bool) => bool)
+                    (%d::real.
+                        (op &::bool => bool => bool)
+                         ((op <::real => real => bool) (0::real) d)
+                         ((All::(real => bool) => bool)
+                           (%y::real.
+                               (op -->::bool => bool => bool)
+                                ((op <::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) x y))
+                                  d)
+                                ((op <=::real => real => bool) (f x)
+                                  (f y)))))))
+                ((op =::real => real => bool) l (0::real)))))"
+  by (import lim DIFF_LMIN)
+
+lemma DIFF_LCONST: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) f l x)
+                  ((Ex::(real => bool) => bool)
+                    (%d::real.
+                        (op &::bool => bool => bool)
+                         ((op <::real => real => bool) (0::real) d)
+                         ((All::(real => bool) => bool)
+                           (%y::real.
+                               (op -->::bool => bool => bool)
+                                ((op <::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) x y))
+                                  d)
+                                ((op =::real => real => bool) (f y)
+                                  (f x)))))))
+                ((op =::real => real => bool) l (0::real)))))"
+  by (import lim DIFF_LCONST)
+
+lemma INTERVAL_LEMMA: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (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))
+                ((Ex::(real => bool) => bool)
+                  (%d::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((All::(real => bool) => bool)
+                         (%y::real.
+                             (op -->::bool => bool => bool)
+                              ((op <::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) x y))
+                                d)
+                              ((op &::bool => bool => bool)
+                                ((op <=::real => real => bool) a y)
+                                ((op <=::real => real => bool) y b)))))))))"
+  by (import lim INTERVAL_LEMMA)
+
+lemma ROLLE: "(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)
+                  ((op &::bool => bool => bool)
+                    ((op =::real => real => bool) (f a) (f b))
+                    ((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))
+                             ((contl::(real => real) => real => bool) f x)))
+                      ((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))
+                             ((differentiable::(real => real)
+         => real => bool)
+                               f x))))))
+                ((Ex::(real => bool) => bool)
+                  (%z::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) a z)
+                       ((op &::bool => bool => bool)
+                         ((op <::real => real => bool) z b)
+                         ((diffl::(real => real) => real => real => bool) f
+                           (0::real) 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::((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)
+                  ((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))
+                           ((contl::(real => real) => real => bool) f x)))
+                    ((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))
+                           ((differentiable::(real => real) => real => bool)
+                             f x)))))
+                ((Ex::(real => bool) => bool)
+                  (%l::real.
+                      (Ex::(real => bool) => bool)
+                       (%z::real.
+                           (op &::bool => bool => bool)
+                            ((op <::real => real => bool) a z)
+                            ((op &::bool => bool => bool)
+                              ((op <::real => real => bool) z b)
+                              ((op &::bool => bool => bool)
+                                ((diffl::(real => real)
+   => real => real => bool)
+                                  f l z)
+                                ((op =::real => real => bool)
+                                  ((op -::real => real => real) (f b) (f a))
+                                  ((op *::real => real => real)
+                                    ((op -::real => real => real) b a)
+                                    l))))))))))"
+  by (import lim MVT)
+
+lemma DIFF_ISCONST_END: "(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)
+                  ((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))
+                           ((contl::(real => real) => real => bool) f x)))
+                    ((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))
+                           ((diffl::(real => real) => real => real => bool)
+                             f (0::real) x)))))
+                ((op =::real => real => bool) (f b) (f a)))))"
+  by (import lim DIFF_ISCONST_END)
+
+lemma DIFF_ISCONST: "(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)
+                  ((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))
+                           ((contl::(real => real) => real => bool) f x)))
+                    ((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))
+                           ((diffl::(real => real) => real => real => bool)
+                             f (0::real) x)))))
+                ((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) (f a)))))))"
+  by (import lim DIFF_ISCONST)
+
+lemma DIFF_ISCONST_ALL: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (op -->::bool => bool => bool)
+      ((All::(real => bool) => bool)
+        ((diffl::(real => real) => real => real => bool) f (0::real)))
+      ((All::(real => bool) => bool)
+        (%x::real.
+            (All::(real => bool) => bool)
+             (%y::real. (op =::real => real => bool) (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)"
+  by (import lim INTERVAL_ABS)
+
+lemma CONT_INJ_LEMMA: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (All::(real => bool) => bool)
+                (%d::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((op &::bool => bool => bool)
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((op =::real => real => bool) (g (f z)) z)))
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((contl::(real => real) => real => bool) f
+                                  z)))))
+                     ((Not::bool => bool)
+                       ((All::(real => bool) => bool)
+                         (%z::real.
+                             (op -->::bool => bool => bool)
+                              ((op <=::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) z x))
+                                d)
+                              ((op <=::real => real => bool) (f z)
+                                (f x)))))))))"
+  by (import lim CONT_INJ_LEMMA)
+
+lemma CONT_INJ_LEMMA2: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (All::(real => bool) => bool)
+                (%d::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((op &::bool => bool => bool)
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((op =::real => real => bool) (g (f z)) z)))
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((contl::(real => real) => real => bool) f
+                                  z)))))
+                     ((Not::bool => bool)
+                       ((All::(real => bool) => bool)
+                         (%z::real.
+                             (op -->::bool => bool => bool)
+                              ((op <=::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) z x))
+                                d)
+                              ((op <=::real => real => bool) (f x)
+                                (f z)))))))))"
+  by (import lim CONT_INJ_LEMMA2)
+
+lemma CONT_INJ_RANGE: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (All::(real => bool) => bool)
+                (%d::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((op &::bool => bool => bool)
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((op =::real => real => bool) (g (f z)) z)))
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((contl::(real => real) => real => bool) f
+                                  z)))))
+                     ((Ex::(real => bool) => bool)
+                       (%e::real.
+                           (op &::bool => bool => bool)
+                            ((op <::real => real => bool) (0::real) e)
+                            ((All::(real => bool) => bool)
+                              (%y::real.
+                                  (op -->::bool => bool => bool)
+                                   ((op <=::real => real => bool)
+                                     ((abs::real => real)
+ ((op -::real => real => real) y (f x)))
+                                     e)
+                                   ((Ex::(real => bool) => bool)
+                                     (%z::real.
+   (op &::bool => bool => bool)
+    ((op <=::real => real => bool)
+      ((abs::real => real) ((op -::real => real => real) z x)) d)
+    ((op =::real => real => bool) (f z) y)))))))))))"
+  by (import lim CONT_INJ_RANGE)
+
+lemma CONT_INVERSE: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (All::(real => bool) => bool)
+                (%d::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((op &::bool => bool => bool)
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((op =::real => real => bool) (g (f z)) z)))
+                         ((All::(real => bool) => bool)
+                           (%z::real.
+                               (op -->::bool => bool => bool)
+                                ((op <=::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real) z x))
+                                  d)
+                                ((contl::(real => real) => real => bool) f
+                                  z)))))
+                     ((contl::(real => real) => real => bool) g (f x))))))"
+  by (import lim CONT_INVERSE)
+
+lemma DIFF_INVERSE: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%x::real.
+                    (All::(real => bool) => bool)
+                     (%d::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((op <::real => real => bool) (0::real) d)
+                            ((op &::bool => bool => bool)
+                              ((All::(real => bool) => bool)
+                                (%z::real.
+                                    (op -->::bool => bool => bool)
+                                     ((op <=::real => real => bool)
+ ((abs::real => real) ((op -::real => real => real) z x)) d)
+                                     ((op =::real => real => bool) (g (f z))
+ z)))
+                              ((op &::bool => bool => bool)
+                                ((All::(real => bool) => bool)
+                                  (%z::real.
+(op -->::bool => bool => bool)
+ ((op <=::real => real => bool)
+   ((abs::real => real) ((op -::real => real => real) z x)) d)
+ ((contl::(real => real) => real => bool) f z)))
+                                ((op &::bool => bool => bool)
+                                  ((diffl::(real => real)
+     => real => real => bool)
+                                    f l x)
+                                  ((Not::bool => bool)
+                                    ((op =::real => real => bool) l
+(0::real)))))))
+                          ((diffl::(real => real) => real => real => bool) g
+                            ((inverse::real => real) l) (f x)))))))"
+  by (import lim DIFF_INVERSE)
+
+lemma DIFF_INVERSE_LT: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%x::real.
+                    (All::(real => bool) => bool)
+                     (%d::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((op <::real => real => bool) (0::real) d)
+                            ((op &::bool => bool => bool)
+                              ((All::(real => bool) => bool)
+                                (%z::real.
+                                    (op -->::bool => bool => bool)
+                                     ((op <::real => real => bool)
+ ((abs::real => real) ((op -::real => real => real) z x)) d)
+                                     ((op =::real => real => bool) (g (f z))
+ z)))
+                              ((op &::bool => bool => bool)
+                                ((All::(real => bool) => bool)
+                                  (%z::real.
+(op -->::bool => bool => bool)
+ ((op <::real => real => bool)
+   ((abs::real => real) ((op -::real => real => real) z x)) d)
+ ((contl::(real => real) => real => bool) f z)))
+                                ((op &::bool => bool => bool)
+                                  ((diffl::(real => real)
+     => real => real => bool)
+                                    f l x)
+                                  ((Not::bool => bool)
+                                    ((op =::real => real => bool) l
+(0::real)))))))
+                          ((diffl::(real => real) => real => real => bool) g
+                            ((inverse::real => real) l) (f x)))))))"
+  by (import lim DIFF_INVERSE_LT)
+
+lemma INTERVAL_CLEMMA: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (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))
+                ((Ex::(real => bool) => bool)
+                  (%d::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) d)
+                       ((All::(real => bool) => bool)
+                         (%y::real.
+                             (op -->::bool => bool => bool)
+                              ((op <=::real => real => bool)
+                                ((abs::real => real)
+                                  ((op -::real => real => real) y x))
+                                d)
+                              ((op &::bool => bool => bool)
+                                ((op <::real => real => bool) a y)
+                                ((op <::real => real => bool) y b)))))))))"
+  by (import lim INTERVAL_CLEMMA)
+
+lemma DIFF_INVERSE_OPEN: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (All::(real => bool) => bool)
+           (%l::real.
+               (All::(real => bool) => bool)
+                (%a::real.
+                    (All::(real => bool) => bool)
+                     (%x::real.
+                         (All::(real => bool) => bool)
+                          (%b::real.
+                              (op -->::bool => bool => bool)
+                               ((op &::bool => bool => bool)
+                                 ((op <::real => real => bool) a x)
+                                 ((op &::bool => bool => bool)
+                                   ((op <::real => real => bool) x b)
+                                   ((op &::bool => bool => bool)
+                                     ((All::(real => bool) => bool)
+ (%z::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool) ((op <::real => real => bool) a z)
+        ((op <::real => real => bool) z b))
+      ((op &::bool => bool => bool)
+        ((op =::real => real => bool) (g (f z)) z)
+        ((contl::(real => real) => real => bool) f z))))
+                                     ((op &::bool => bool => bool)
+ ((diffl::(real => real) => real => real => bool) f l x)
+ ((Not::bool => bool) ((op =::real => real => bool) l (0::real)))))))
+                               ((diffl::(real => real)
+  => real => real => bool)
+                                 g ((inverse::real => real) l) (f x))))))))"
+  by (import lim DIFF_INVERSE_OPEN)
+
+;end_setup
+
+;setup_theory powser
+
+lemma POWDIFF_LEMMA: "ALL n x y.
+   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 x y.
+   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 x y.
+   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::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%z::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((summable::(nat => real) => bool)
+                    (%n::nat.
+                        (op *::real => real => real) (f n)
+                         ((op ^::real => nat => real) x n)))
+                  ((op <::real => real => bool) ((abs::real => real) z)
+                    ((abs::real => real) x)))
+                ((summable::(nat => real) => bool)
+                  (%n::nat.
+                      (op *::real => real => real)
+                       ((abs::real => real) (f n))
+                       ((op ^::real => nat => real) z n))))))"
+  by (import powser POWSER_INSIDEA)
+
+lemma POWSER_INSIDE: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%z::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((summable::(nat => real) => bool)
+                    (%n::nat.
+                        (op *::real => real => real) (f n)
+                         ((op ^::real => nat => real) x n)))
+                  ((op <::real => real => bool) ((abs::real => real) z)
+                    ((abs::real => real) x)))
+                ((summable::(nat => real) => bool)
+                  (%n::nat.
+                      (op *::real => real => real) (f n)
+                       ((op ^::real => nat => real) z n))))))"
+  by (import powser POWSER_INSIDE)
+
+constdefs
+  diffs :: "(nat => real) => nat => real" 
+  "diffs == %c n. real (Suc n) * c (Suc n)"
+
+lemma diffs: "ALL c. diffs c = (%n. real (Suc n) * c (Suc n))"
+  by (import powser diffs)
+
+lemma DIFFS_NEG: "ALL c. diffs (%n. - c n) = (%x. - diffs c x)"
+  by (import powser DIFFS_NEG)
+
+lemma DIFFS_LEMMA: "ALL n c x.
+   real.sum (0, n) (%n. 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 c x.
+   real.sum (0, n) (%n. real n * (c n * x ^ (n - 1))) =
+   real.sum (0, n) (%n. diffs c n * x ^ n) - real n * (c n * x ^ (n - 1))"
+  by (import powser DIFFS_LEMMA2)
+
+lemma DIFFS_EQUIV: "(All::((nat => real) => bool) => bool)
+ (%c::nat => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((summable::(nat => real) => bool)
+             (%n::nat.
+                 (op *::real => real => real)
+                  ((diffs::(nat => real) => nat => real) c n)
+                  ((op ^::real => nat => real) x n)))
+           ((sums::(nat => real) => real => bool)
+             (%n::nat.
+                 (op *::real => real => real) ((real::nat => real) n)
+                  ((op *::real => real => real) (c n)
+                    ((op ^::real => nat => real) x
+                      ((op -::nat => nat => nat) n (1::nat)))))
+             ((suminf::(nat => real) => real)
+               (%n::nat.
+                   (op *::real => real => real)
+                    ((diffs::(nat => real) => nat => real) c n)
+                    ((op ^::real => nat => real) x n))))))"
+  by (import powser DIFFS_EQUIV)
+
+lemma TERMDIFF_LEMMA1: "ALL m z h.
+   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::(real => bool) => bool)
+ (%z::real.
+     (All::(real => bool) => bool)
+      (%h::real.
+          (All::(nat => bool) => bool)
+           (%n::nat.
+               (op -->::bool => bool => bool)
+                ((Not::bool => bool)
+                  ((op =::real => real => bool) h (0::real)))
+                ((op =::real => real => bool)
+                  ((op -::real => real => real)
+                    ((op /::real => real => real)
+                      ((op -::real => real => real)
+                        ((op ^::real => nat => real)
+                          ((op +::real => real => real) z h) n)
+                        ((op ^::real => nat => real) z n))
+                      h)
+                    ((op *::real => real => real) ((real::nat => real) n)
+                      ((op ^::real => nat => real) z
+                        ((op -::nat => nat => nat) n (1::nat)))))
+                  ((op *::real => real => real) h
+                    ((real.sum::nat * nat => (nat => real) => real)
+                      ((Pair::nat => nat => nat * nat) (0::nat)
+                        ((op -::nat => nat => nat) n (1::nat)))
+                      (%p::nat.
+                          (op *::real => real => real)
+                           ((op ^::real => nat => real) z p)
+                           ((real.sum::nat * nat => (nat => real) => real)
+                             ((Pair::nat => nat => nat * nat) (0::nat)
+                               ((op -::nat => nat => nat)
+                                 ((op -::nat => nat => nat) n (1::nat)) p))
+                             (%q::nat.
+                                 (op *::real => real => real)
+                                  ((op ^::real => nat => real)
+                                    ((op +::real => real => real) z h) q)
+                                  ((op ^::real => nat => real) z
+                                    ((op -::nat => nat => nat)
+((op -::nat => nat => nat)
+  ((op -::nat => nat => nat) n
+    ((number_of::bin => nat)
+      ((op BIT::bin => bool => bin)
+        ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+        (False::bool))))
+  p)
+q)))))))))))"
+  by (import powser TERMDIFF_LEMMA2)
+
+lemma TERMDIFF_LEMMA3: "(All::(real => bool) => bool)
+ (%z::real.
+     (All::(real => bool) => bool)
+      (%h::real.
+          (All::(nat => bool) => bool)
+           (%n::nat.
+               (All::(real => bool) => bool)
+                (%k'::real.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((Not::bool => bool)
+                         ((op =::real => real => bool) h (0::real)))
+                       ((op &::bool => bool => bool)
+                         ((op <=::real => real => bool)
+                           ((abs::real => real) z) k')
+                         ((op <=::real => real => bool)
+                           ((abs::real => real)
+                             ((op +::real => real => real) z h))
+                           k')))
+                     ((op <=::real => real => bool)
+                       ((abs::real => real)
+                         ((op -::real => real => real)
+                           ((op /::real => real => real)
+                             ((op -::real => real => real)
+                               ((op ^::real => nat => real)
+                                 ((op +::real => real => real) z h) n)
+                               ((op ^::real => nat => real) z n))
+                             h)
+                           ((op *::real => real => real)
+                             ((real::nat => real) n)
+                             ((op ^::real => nat => real) z
+                               ((op -::nat => nat => nat) n (1::nat))))))
+                       ((op *::real => real => real) ((real::nat => real) n)
+                         ((op *::real => real => real)
+                           ((real::nat => real)
+                             ((op -::nat => nat => nat) n (1::nat)))
+                           ((op *::real => real => real)
+                             ((op ^::real => nat => real) k'
+                               ((op -::nat => nat => nat) n
+                                 ((number_of::bin => nat)
+                                   ((op BIT::bin => bool => bin)
+                                     ((op BIT::bin => bool => bin)
+ (bin.Pls::bin) (True::bool))
+                                     (False::bool)))))
+                             ((abs::real => real) h)))))))))"
+  by (import powser TERMDIFF_LEMMA3)
+
+lemma TERMDIFF_LEMMA4: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::(real => bool) => bool)
+      (%k'::real.
+          (All::(real => bool) => bool)
+           (%k::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <::real => real => bool) (0::real) k)
+                  ((All::(real => bool) => bool)
+                    (%h::real.
+                        (op -->::bool => bool => bool)
+                         ((op &::bool => bool => bool)
+                           ((op <::real => real => bool) (0::real)
+                             ((abs::real => real) h))
+                           ((op <::real => real => bool)
+                             ((abs::real => real) h) k))
+                         ((op <=::real => real => bool)
+                           ((abs::real => real) (f h))
+                           ((op *::real => real => real) k'
+                             ((abs::real => real) h))))))
+                ((tends_real_real::(real => real) => real => real => bool) f
+                  (0::real) (0::real)))))"
+  by (import powser TERMDIFF_LEMMA4)
+
+lemma TERMDIFF_LEMMA5: "(All::((nat => real) => bool) => bool)
+ (%f::nat => real.
+     (All::((real => nat => real) => bool) => bool)
+      (%g::real => nat => real.
+          (All::(real => bool) => bool)
+           (%k::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <::real => real => bool) (0::real) k)
+                  ((op &::bool => bool => bool)
+                    ((summable::(nat => real) => bool) f)
+                    ((All::(real => bool) => bool)
+                      (%h::real.
+                          (op -->::bool => bool => bool)
+                           ((op &::bool => bool => bool)
+                             ((op <::real => real => bool) (0::real)
+                               ((abs::real => real) h))
+                             ((op <::real => real => bool)
+                               ((abs::real => real) h) k))
+                           ((All::(nat => bool) => bool)
+                             (%n::nat.
+                                 (op <=::real => real => bool)
+                                  ((abs::real => real) (g h n))
+                                  ((op *::real => real => real) (f n)
+                                    ((abs::real => real) h))))))))
+                ((tends_real_real::(real => real) => real => real => bool)
+                  (%h::real. (suminf::(nat => real) => real) (g h))
+                  (0::real) (0::real)))))"
+  by (import powser TERMDIFF_LEMMA5)
+
+lemma TERMDIFF: "(All::((nat => real) => bool) => bool)
+ (%c::nat => real.
+     (All::(real => bool) => bool)
+      (%k'::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((summable::(nat => real) => bool)
+                    (%n::nat.
+                        (op *::real => real => real) (c n)
+                         ((op ^::real => nat => real) k' n)))
+                  ((op &::bool => bool => bool)
+                    ((summable::(nat => real) => bool)
+                      (%n::nat.
+                          (op *::real => real => real)
+                           ((diffs::(nat => real) => nat => real) c n)
+                           ((op ^::real => nat => real) k' n)))
+                    ((op &::bool => bool => bool)
+                      ((summable::(nat => real) => bool)
+                        (%n::nat.
+                            (op *::real => real => real)
+                             ((diffs::(nat => real) => nat => real)
+                               ((diffs::(nat => real) => nat => real) c) n)
+                             ((op ^::real => nat => real) k' n)))
+                      ((op <::real => real => bool) ((abs::real => real) x)
+                        ((abs::real => real) k')))))
+                ((diffl::(real => real) => real => real => bool)
+                  (%x::real.
+                      (suminf::(nat => real) => real)
+                       (%n::nat.
+                           (op *::real => real => real) (c n)
+                            ((op ^::real => nat => real) x n)))
+                  ((suminf::(nat => real) => real)
+                    (%n::nat.
+                        (op *::real => real => real)
+                         ((diffs::(nat => real) => nat => real) c n)
+                         ((op ^::real => nat => real) x n)))
+                  x))))"
+  by (import powser TERMDIFF)
+
+;end_setup
+
+;setup_theory transc
+
+constdefs
+  exp :: "real => real" 
+  "exp == %x. suminf (%n. inverse (real (FACT n)) * x ^ n)"
+
+lemma exp: "ALL x. exp x = suminf (%n. inverse (real (FACT n)) * x ^ n)"
+  by (import transc exp)
+
+constdefs
+  cos :: "real => real" 
+  "cos ==
+%x. suminf
+     (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
+
+lemma cos: "ALL x.
+   cos x =
+   suminf
+    (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)"
+  by (import transc cos)
+
+constdefs
+  sin :: "real => real" 
+  "sin ==
+%x. suminf
+     (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+          x ^ n)"
+
+lemma sin: "ALL x.
+   sin x =
+   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. sums (%n. inverse (real (FACT n)) * x ^ n) (exp x)"
+  by (import transc EXP_CONVERGES)
+
+lemma SIN_CONVERGES: "ALL x.
+   sums
+    (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) *
+         x ^ n)
+    (sin x)"
+  by (import transc SIN_CONVERGES)
+
+lemma COS_CONVERGES: "ALL x.
+   sums
+    (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)
+    (cos x)"
+  by (import transc COS_CONVERGES)
+
+lemma EXP_FDIFF: "diffs (%n. inverse (real (FACT n))) = (%n. inverse (real (FACT n)))"
+  by (import transc EXP_FDIFF)
+
+lemma SIN_FDIFF: "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. 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.
+   - sin x =
+   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. diffl exp (exp x) x"
+  by (import transc DIFF_EXP)
+
+lemma DIFF_SIN: "ALL x. diffl sin (cos x) x"
+  by (import transc DIFF_SIN)
+
+lemma DIFF_COS: "ALL x. diffl cos (- sin x) x"
+  by (import transc DIFF_COS)
+
+lemma DIFF_COMPOSITE: "(op &::bool => bool => bool)
+ ((op -->::bool => bool => bool)
+   ((op &::bool => bool => bool)
+     ((diffl::(real => real) => real => real => bool) (f::real => real)
+       (l::real) (x::real))
+     ((Not::bool => bool) ((op =::real => real => bool) (f x) (0::real))))
+   ((diffl::(real => real) => real => real => bool)
+     (%x::real. (inverse::real => real) (f x))
+     ((uminus::real => real)
+       ((op /::real => real => real) l
+         ((op ^::real => nat => real) (f x)
+           ((number_of::bin => nat)
+             ((op BIT::bin => bool => bin)
+               ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+               (False::bool))))))
+     x))
+ ((op &::bool => bool => bool)
+   ((op -->::bool => bool => bool)
+     ((op &::bool => bool => bool)
+       ((diffl::(real => real) => real => real => bool) f l x)
+       ((op &::bool => bool => bool)
+         ((diffl::(real => real) => real => real => bool) (g::real => real)
+           (m::real) x)
+         ((Not::bool => bool)
+           ((op =::real => real => bool) (g x) (0::real)))))
+     ((diffl::(real => real) => real => real => bool)
+       (%x::real. (op /::real => real => real) (f x) (g x))
+       ((op /::real => real => real)
+         ((op -::real => real => real)
+           ((op *::real => real => real) l (g x))
+           ((op *::real => real => real) m (f x)))
+         ((op ^::real => nat => real) (g x)
+           ((number_of::bin => nat)
+             ((op BIT::bin => bool => bin)
+               ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+               (False::bool)))))
+       x))
+   ((op &::bool => bool => bool)
+     ((op -->::bool => bool => bool)
+       ((op &::bool => bool => bool)
+         ((diffl::(real => real) => real => real => bool) f l x)
+         ((diffl::(real => real) => real => real => bool) g m x))
+       ((diffl::(real => real) => real => real => bool)
+         (%x::real. (op +::real => real => real) (f x) (g x))
+         ((op +::real => real => real) l m) x))
+     ((op &::bool => bool => bool)
+       ((op -->::bool => bool => bool)
+         ((op &::bool => bool => bool)
+           ((diffl::(real => real) => real => real => bool) f l x)
+           ((diffl::(real => real) => real => real => bool) g m x))
+         ((diffl::(real => real) => real => real => bool)
+           (%x::real. (op *::real => real => real) (f x) (g x))
+           ((op +::real => real => real)
+             ((op *::real => real => real) l (g x))
+             ((op *::real => real => real) m (f x)))
+           x))
+       ((op &::bool => bool => bool)
+         ((op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((diffl::(real => real) => real => real => bool) f l x)
+             ((diffl::(real => real) => real => real => bool) g m x))
+           ((diffl::(real => real) => real => real => bool)
+             (%x::real. (op -::real => real => real) (f x) (g x))
+             ((op -::real => real => real) l m) x))
+         ((op &::bool => bool => bool)
+           ((op -->::bool => bool => bool)
+             ((diffl::(real => real) => real => real => bool) f l x)
+             ((diffl::(real => real) => real => real => bool)
+               (%x::real. (uminus::real => real) (f x))
+               ((uminus::real => real) l) x))
+           ((op &::bool => bool => bool)
+             ((op -->::bool => bool => bool)
+               ((diffl::(real => real) => real => real => bool) g m x)
+               ((diffl::(real => real) => real => real => bool)
+                 (%x::real. (op ^::real => nat => real) (g x) (n::nat))
+                 ((op *::real => real => real)
+                   ((op *::real => real => real) ((real::nat => real) n)
+                     ((op ^::real => nat => real) (g x)
+                       ((op -::nat => nat => nat) n (1::nat))))
+                   m)
+                 x))
+             ((op &::bool => bool => bool)
+               ((op -->::bool => bool => bool)
+                 ((diffl::(real => real) => real => real => bool) g m x)
+                 ((diffl::(real => real) => real => real => bool)
+                   (%x::real. (exp::real => real) (g x))
+                   ((op *::real => real => real) ((exp::real => real) (g x))
+                     m)
+                   x))
+               ((op &::bool => bool => bool)
+                 ((op -->::bool => bool => bool)
+                   ((diffl::(real => real) => real => real => bool) g m x)
+                   ((diffl::(real => real) => real => real => bool)
+                     (%x::real. (sin::real => real) (g x))
+                     ((op *::real => real => real)
+                       ((cos::real => real) (g x)) m)
+                     x))
+                 ((op -->::bool => bool => bool)
+                   ((diffl::(real => real) => real => real => bool) g m x)
+                   ((diffl::(real => real) => real => real => bool)
+                     (%x::real. (cos::real => real) (g x))
+                     ((op *::real => real => real)
+                       ((uminus::real => real) ((sin::real => real) (g x)))
+                       m)
+                     x))))))))))"
+  by (import transc DIFF_COMPOSITE)
+
+lemma EXP_0: "exp 0 = 1"
+  by (import transc EXP_0)
+
+lemma EXP_LE_X: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op <=::real => real => bool)
+        ((op +::real => real => real) (1::real) x) ((exp::real => real) x)))"
+  by (import transc EXP_LE_X)
+
+lemma EXP_LT_1: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) x)
+      ((op <::real => real => bool) (1::real) ((exp::real => real) x)))"
+  by (import transc EXP_LT_1)
+
+lemma EXP_ADD_MUL: "ALL x y. exp (x + y) * exp (- x) = exp y"
+  by (import transc EXP_ADD_MUL)
+
+lemma EXP_NEG_MUL: "ALL x. exp x * exp (- x) = 1"
+  by (import transc EXP_NEG_MUL)
+
+lemma EXP_NEG_MUL2: "ALL x. exp (- x) * exp x = 1"
+  by (import transc EXP_NEG_MUL2)
+
+lemma EXP_NEG: "ALL x. exp (- x) = inverse (exp x)"
+  by (import transc EXP_NEG)
+
+lemma EXP_ADD: "ALL x y. exp (x + y) = exp x * exp y"
+  by (import transc EXP_ADD)
+
+lemma EXP_POS_LE: "ALL x. 0 <= exp x"
+  by (import transc EXP_POS_LE)
+
+lemma EXP_NZ: "ALL x. exp x ~= 0"
+  by (import transc EXP_NZ)
+
+lemma EXP_POS_LT: "ALL x. 0 < exp x"
+  by (import transc EXP_POS_LT)
+
+lemma EXP_N: "ALL n x. exp (real n * x) = exp x ^ n"
+  by (import transc EXP_N)
+
+lemma EXP_SUB: "ALL x y. exp (x - y) = exp x / exp y"
+  by (import transc EXP_SUB)
+
+lemma EXP_MONO_IMP: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool) ((op <::real => real => bool) x y)
+           ((op <::real => real => bool) ((exp::real => real) x)
+             ((exp::real => real) y))))"
+  by (import transc EXP_MONO_IMP)
+
+lemma EXP_MONO_LT: "ALL x y. (exp x < exp y) = (x < y)"
+  by (import transc EXP_MONO_LT)
+
+lemma EXP_MONO_LE: "ALL x y. (exp x <= exp y) = (x <= y)"
+  by (import transc EXP_MONO_LE)
+
+lemma EXP_INJ: "ALL x y. (exp x = exp y) = (x = y)"
+  by (import transc EXP_INJ)
+
+lemma EXP_TOTAL_LEMMA: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (1::real) y)
+      ((Ex::(real => bool) => bool)
+        (%x::real.
+            (op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <=::real => real => bool) x
+                 ((op -::real => real => real) y (1::real)))
+               ((op =::real => real => bool) ((exp::real => real) x) y)))))"
+  by (import transc EXP_TOTAL_LEMMA)
+
+lemma EXP_TOTAL: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) y)
+      ((Ex::(real => bool) => bool)
+        (%x::real. (op =::real => real => bool) ((exp::real => real) x) y)))"
+  by (import transc EXP_TOTAL)
+
+constdefs
+  ln :: "real => real" 
+  "ln == %x. SOME u. exp u = x"
+
+lemma ln: "ALL x. ln x = (SOME u. exp u = x)"
+  by (import transc ln)
+
+lemma LN_EXP: "ALL x. ln (exp x) = x"
+  by (import transc LN_EXP)
+
+lemma EXP_LN: "ALL x. (exp (ln x) = x) = (0 < x)"
+  by (import transc EXP_LN)
+
+lemma LN_MUL: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) y))
+           ((op =::real => real => bool)
+             ((ln::real => real) ((op *::real => real => real) x y))
+             ((op +::real => real => real) ((ln::real => real) x)
+               ((ln::real => real) y)))))"
+  by (import transc LN_MUL)
+
+lemma LN_INJ: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) y))
+           ((op =::bool => bool => bool)
+             ((op =::real => real => bool) ((ln::real => real) x)
+               ((ln::real => real) y))
+             ((op =::real => real => bool) x y))))"
+  by (import transc LN_INJ)
+
+lemma LN_1: "ln 1 = 0"
+  by (import transc LN_1)
+
+lemma LN_INV: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) x)
+      ((op =::real => real => bool)
+        ((ln::real => real) ((inverse::real => real) x))
+        ((uminus::real => real) ((ln::real => real) x))))"
+  by (import transc LN_INV)
+
+lemma LN_DIV: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) y))
+           ((op =::real => real => bool)
+             ((ln::real => real) ((op /::real => real => real) x y))
+             ((op -::real => real => real) ((ln::real => real) x)
+               ((ln::real => real) y)))))"
+  by (import transc LN_DIV)
+
+lemma LN_MONO_LT: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) y))
+           ((op =::bool => bool => bool)
+             ((op <::real => real => bool) ((ln::real => real) x)
+               ((ln::real => real) y))
+             ((op <::real => real => bool) x y))))"
+  by (import transc LN_MONO_LT)
+
+lemma LN_MONO_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) y))
+           ((op =::bool => bool => bool)
+             ((op <=::real => real => bool) ((ln::real => real) x)
+               ((ln::real => real) y))
+             ((op <=::real => real => bool) x y))))"
+  by (import transc LN_MONO_LE)
+
+lemma LN_POW: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((ln::real => real) ((op ^::real => nat => real) x n))
+             ((op *::real => real => real) ((real::nat => real) n)
+               ((ln::real => real) x)))))"
+  by (import transc LN_POW)
+
+lemma LN_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op <=::real => real => bool)
+        ((ln::real => real) ((op +::real => real => real) (1::real) x)) x))"
+  by (import transc LN_LE)
+
+lemma LN_LT_X: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) x)
+      ((op <::real => real => bool) ((ln::real => real) x) x))"
+  by (import transc LN_LT_X)
+
+lemma LN_POS: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (1::real) x)
+      ((op <=::real => real => bool) (0::real) ((ln::real => real) x)))"
+  by (import transc LN_POS)
+
+constdefs
+  root :: "nat => real => real" 
+  "(op ==::(nat => real => real) => (nat => real => real) => prop)
+ (root::nat => real => real)
+ (%(n::nat) x::real.
+     (Eps::(real => bool) => real)
+      (%u::real.
+          (op &::bool => bool => bool)
+           ((op -->::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op <::real => real => bool) (0::real) u))
+           ((op =::real => real => bool) ((op ^::real => nat => real) u n)
+             x)))"
+
+lemma root: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op =::real => real => bool) ((root::nat => real => real) n x)
+           ((Eps::(real => bool) => real)
+             (%u::real.
+                 (op &::bool => bool => bool)
+                  ((op -->::bool => bool => bool)
+                    ((op <::real => real => bool) (0::real) x)
+                    ((op <::real => real => bool) (0::real) u))
+                  ((op =::real => real => bool)
+                    ((op ^::real => nat => real) u n) x)))))"
+  by (import transc root)
+
+constdefs
+  sqrt :: "real => real" 
+  "sqrt == root 2"
+
+lemma sqrt: "ALL x. sqrt x = root 2 x"
+  by (import transc sqrt)
+
+lemma ROOT_LT_LEMMA: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((op ^::real => nat => real)
+               ((exp::real => real)
+                 ((op /::real => real => real) ((ln::real => real) x)
+                   ((real::nat => real) ((Suc::nat => nat) n))))
+               ((Suc::nat => nat) n))
+             x)))"
+  by (import transc ROOT_LT_LEMMA)
+
+lemma ROOT_LN: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((root::nat => real => real) ((Suc::nat => nat) n) x)
+             ((exp::real => real)
+               ((op /::real => real => real) ((ln::real => real) x)
+                 ((real::nat => real) ((Suc::nat => nat) n)))))))"
+  by (import transc ROOT_LN)
+
+lemma ROOT_0: "ALL n. root (Suc n) 0 = 0"
+  by (import transc ROOT_0)
+
+lemma ROOT_1: "ALL n. root (Suc n) 1 = 1"
+  by (import transc ROOT_1)
+
+lemma ROOT_POS_LT: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <::real => real => bool) (0::real) x)
+           ((op <::real => real => bool) (0::real)
+             ((root::nat => real => real) ((Suc::nat => nat) n) x))))"
+  by (import transc ROOT_POS_LT)
+
+lemma ROOT_POW_POS: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <=::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((op ^::real => nat => real)
+               ((root::nat => real => real) ((Suc::nat => nat) n) x)
+               ((Suc::nat => nat) n))
+             x)))"
+  by (import transc ROOT_POW_POS)
+
+lemma POW_ROOT_POS: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <=::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((root::nat => real => real) ((Suc::nat => nat) n)
+               ((op ^::real => nat => real) x ((Suc::nat => nat) n)))
+             x)))"
+  by (import transc POW_ROOT_POS)
+
+lemma ROOT_POS: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <=::real => real => bool) (0::real) x)
+           ((op <=::real => real => bool) (0::real)
+             ((root::nat => real => real) ((Suc::nat => nat) n) x))))"
+  by (import transc ROOT_POS)
+
+lemma ROOT_POS_UNIQ: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%y::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <=::real => real => bool) (0::real) x)
+                  ((op &::bool => bool => bool)
+                    ((op <=::real => real => bool) (0::real) y)
+                    ((op =::real => real => bool)
+                      ((op ^::real => nat => real) y ((Suc::nat => nat) n))
+                      x)))
+                ((op =::real => real => bool)
+                  ((root::nat => real => real) ((Suc::nat => nat) n) x)
+                  y))))"
+  by (import transc ROOT_POS_UNIQ)
+
+lemma ROOT_MUL: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (All::(real => bool) => bool)
+           (%y::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <=::real => real => bool) (0::real) x)
+                  ((op <=::real => real => bool) (0::real) y))
+                ((op =::real => real => bool)
+                  ((root::nat => real => real) ((Suc::nat => nat) n)
+                    ((op *::real => real => real) x y))
+                  ((op *::real => real => real)
+                    ((root::nat => real => real) ((Suc::nat => nat) n) x)
+                    ((root::nat => real => real) ((Suc::nat => nat) n)
+                      y))))))"
+  by (import transc ROOT_MUL)
+
+lemma ROOT_INV: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool)
+           ((op <=::real => real => bool) (0::real) x)
+           ((op =::real => real => bool)
+             ((root::nat => real => real) ((Suc::nat => nat) n)
+               ((inverse::real => real) x))
+             ((inverse::real => real)
+               ((root::nat => real => real) ((Suc::nat => nat) n) x)))))"
+  by (import transc ROOT_INV)
+
+lemma ROOT_DIV: "(All::(nat => bool) => bool)
+ (%x::nat.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (All::(real => bool) => bool)
+           (%xb::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <=::real => real => bool) (0::real) xa)
+                  ((op <=::real => real => bool) (0::real) xb))
+                ((op =::real => real => bool)
+                  ((root::nat => real => real) ((Suc::nat => nat) x)
+                    ((op /::real => real => real) xa xb))
+                  ((op /::real => real => real)
+                    ((root::nat => real => real) ((Suc::nat => nat) x) xa)
+                    ((root::nat => real => real) ((Suc::nat => nat) x)
+                      xb))))))"
+  by (import transc ROOT_DIV)
+
+lemma ROOT_MONO_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op <=::real => real => bool) x y))
+           ((op <=::real => real => bool)
+             ((root::nat => real => real) ((Suc::nat => nat) (n::nat)) x)
+             ((root::nat => real => real) ((Suc::nat => nat) n) y))))"
+  by (import transc ROOT_MONO_LE)
+
+lemma SQRT_0: "sqrt 0 = 0"
+  by (import transc SQRT_0)
+
+lemma SQRT_1: "sqrt 1 = 1"
+  by (import transc SQRT_1)
+
+lemma SQRT_POS_LT: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) x)
+      ((op <::real => real => bool) (0::real) ((sqrt::real => real) x)))"
+  by (import transc SQRT_POS_LT)
+
+lemma SQRT_POS_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op <=::real => real => bool) (0::real) ((sqrt::real => real) x)))"
+  by (import transc SQRT_POS_LE)
+
+lemma SQRT_POW2: "ALL x. (sqrt x ^ 2 = x) = (0 <= x)"
+  by (import transc SQRT_POW2)
+
+lemma SQRT_POW_2: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op =::real => real => bool)
+        ((op ^::real => nat => real) ((sqrt::real => real) x)
+          ((number_of::bin => nat)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool))))
+        x))"
+  by (import transc SQRT_POW_2)
+
+lemma POW_2_SQRT: "(op -->::bool => bool => bool)
+ ((op <=::real => real => bool) (0::real) (x::real))
+ ((op =::real => real => bool)
+   ((sqrt::real => real)
+     ((op ^::real => nat => real) x
+       ((number_of::bin => nat)
+         ((op BIT::bin => bool => bin)
+           ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+           (False::bool)))))
+   x)"
+  by (import transc POW_2_SQRT)
+
+lemma SQRT_POS_UNIQ: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <=::real => real => bool) (0::real) xa)
+               ((op =::real => real => bool)
+                 ((op ^::real => nat => real) xa
+                   ((number_of::bin => nat)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool))))
+                 x)))
+           ((op =::real => real => bool) ((sqrt::real => real) x) xa)))"
+  by (import transc SQRT_POS_UNIQ)
+
+lemma SQRT_MUL: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op <=::real => real => bool) (0::real) xa))
+           ((op =::real => real => bool)
+             ((sqrt::real => real) ((op *::real => real => real) x xa))
+             ((op *::real => real => real) ((sqrt::real => real) x)
+               ((sqrt::real => real) xa)))))"
+  by (import transc SQRT_MUL)
+
+lemma SQRT_INV: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op =::real => real => bool)
+        ((sqrt::real => real) ((inverse::real => real) x))
+        ((inverse::real => real) ((sqrt::real => real) x))))"
+  by (import transc SQRT_INV)
+
+lemma SQRT_DIV: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op <=::real => real => bool) (0::real) xa))
+           ((op =::real => real => bool)
+             ((sqrt::real => real) ((op /::real => real => real) x xa))
+             ((op /::real => real => real) ((sqrt::real => real) x)
+               ((sqrt::real => real) xa)))))"
+  by (import transc SQRT_DIV)
+
+lemma SQRT_MONO_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op <=::real => real => bool) x xa))
+           ((op <=::real => real => bool) ((sqrt::real => real) x)
+             ((sqrt::real => real) xa))))"
+  by (import transc SQRT_MONO_LE)
+
+lemma SQRT_EVEN_POW2: "(All::(nat => bool) => bool)
+ (%n::nat.
+     (op -->::bool => bool => bool) ((EVEN::nat => bool) n)
+      ((op =::real => real => bool)
+        ((sqrt::real => real)
+          ((op ^::real => nat => real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))
+            n))
+        ((op ^::real => nat => real)
+          ((number_of::bin => real)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool)))
+          ((op div::nat => nat => nat) n
+            ((number_of::bin => nat)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))))"
+  by (import transc SQRT_EVEN_POW2)
+
+lemma REAL_DIV_SQRT: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) x)
+      ((op =::real => real => bool)
+        ((op /::real => real => real) x ((sqrt::real => real) x))
+        ((sqrt::real => real) x)))"
+  by (import transc REAL_DIV_SQRT)
+
+lemma SQRT_EQ: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op =::real => real => bool)
+               ((op ^::real => nat => real) x
+                 ((number_of::bin => nat)
+                   ((op BIT::bin => bool => bin)
+                     ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                       (True::bool))
+                     (False::bool))))
+               y)
+             ((op <=::real => real => bool) (0::real) x))
+           ((op =::real => real => bool) x ((sqrt::real => real) y))))"
+  by (import transc SQRT_EQ)
+
+lemma SIN_0: "sin 0 = 0"
+  by (import transc SIN_0)
+
+lemma COS_0: "cos 0 = 1"
+  by (import transc COS_0)
+
+lemma SIN_CIRCLE: "ALL x. sin x ^ 2 + cos x ^ 2 = 1"
+  by (import transc SIN_CIRCLE)
+
+lemma SIN_BOUND: "ALL x. abs (sin x) <= 1"
+  by (import transc SIN_BOUND)
+
+lemma SIN_BOUNDS: "ALL x. - 1 <= sin x & sin x <= 1"
+  by (import transc SIN_BOUNDS)
+
+lemma COS_BOUND: "ALL x. abs (cos x) <= 1"
+  by (import transc COS_BOUND)
+
+lemma COS_BOUNDS: "ALL x. - 1 <= cos x & cos x <= 1"
+  by (import transc COS_BOUNDS)
+
+lemma SIN_COS_ADD: "ALL x y.
+   (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"
+  by (import transc SIN_COS_ADD)
+
+lemma SIN_COS_NEG: "ALL x. (sin (- x) + sin x) ^ 2 + (cos (- x) - cos x) ^ 2 = 0"
+  by (import transc SIN_COS_NEG)
+
+lemma SIN_ADD: "ALL x y. sin (x + y) = sin x * cos y + cos x * sin y"
+  by (import transc SIN_ADD)
+
+lemma COS_ADD: "ALL x y. cos (x + y) = cos x * cos y - sin x * sin y"
+  by (import transc COS_ADD)
+
+lemma SIN_NEG: "ALL x. sin (- x) = - sin x"
+  by (import transc SIN_NEG)
+
+lemma COS_NEG: "ALL x. cos (- x) = cos x"
+  by (import transc COS_NEG)
+
+lemma SIN_DOUBLE: "ALL x. sin (2 * x) = 2 * (sin x * cos x)"
+  by (import transc SIN_DOUBLE)
+
+lemma COS_DOUBLE: "ALL x. cos (2 * x) = cos x ^ 2 - sin x ^ 2"
+  by (import transc COS_DOUBLE)
+
+lemma SIN_PAIRED: "ALL x.
+   sums (%n. (- 1) ^ n / real (FACT (2 * n + 1)) * x ^ (2 * n + 1)) (sin x)"
+  by (import transc SIN_PAIRED)
+
+lemma SIN_POS: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) x)
+        ((op <::real => real => bool) x
+          ((number_of::bin => real)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool)))))
+      ((op <::real => real => bool) (0::real) ((sin::real => real) x)))"
+  by (import transc SIN_POS)
+
+lemma COS_PAIRED: "ALL x. sums (%n. (- 1) ^ n / real (FACT (2 * n)) * x ^ (2 * n)) (cos x)"
+  by (import transc COS_PAIRED)
+
+lemma COS_2: "cos 2 < 0"
+  by (import transc COS_2)
+
+lemma COS_ISZERO: "EX! x. 0 <= x & x <= 2 & cos x = 0"
+  by (import transc COS_ISZERO)
+
+constdefs
+  pi :: "real" 
+  "pi == 2 * (SOME x. 0 <= x & x <= 2 & cos x = 0)"
+
+lemma pi: "pi = 2 * (SOME x. 0 <= x & x <= 2 & cos x = 0)"
+  by (import transc pi)
+
+lemma PI2: "pi / 2 = (SOME x. 0 <= x & x <= 2 & cos x = 0)"
+  by (import transc PI2)
+
+lemma COS_PI2: "cos (pi / 2) = 0"
+  by (import transc COS_PI2)
+
+lemma PI2_BOUNDS: "0 < pi / 2 & pi / 2 < 2"
+  by (import transc PI2_BOUNDS)
+
+lemma PI_POS: "0 < pi"
+  by (import transc PI_POS)
+
+lemma SIN_PI2: "sin (pi / 2) = 1"
+  by (import transc SIN_PI2)
+
+lemma COS_PI: "cos pi = - 1"
+  by (import transc COS_PI)
+
+lemma SIN_PI: "sin pi = 0"
+  by (import transc SIN_PI)
+
+lemma SIN_COS: "ALL x. sin x = cos (pi / 2 - x)"
+  by (import transc SIN_COS)
+
+lemma COS_SIN: "ALL x. cos x = sin (pi / 2 - x)"
+  by (import transc COS_SIN)
+
+lemma SIN_PERIODIC_PI: "ALL x. sin (x + pi) = - sin x"
+  by (import transc SIN_PERIODIC_PI)
+
+lemma COS_PERIODIC_PI: "ALL x. cos (x + pi) = - cos x"
+  by (import transc COS_PERIODIC_PI)
+
+lemma SIN_PERIODIC: "ALL x. sin (x + 2 * pi) = sin x"
+  by (import transc SIN_PERIODIC)
+
+lemma COS_PERIODIC: "ALL x. cos (x + 2 * pi) = cos x"
+  by (import transc COS_PERIODIC)
+
+lemma COS_NPI: "ALL n. cos (real n * pi) = (- 1) ^ n"
+  by (import transc COS_NPI)
+
+lemma SIN_NPI: "ALL n::nat. sin (real n * pi) = (0::real)"
+  by (import transc SIN_NPI)
+
+lemma SIN_POS_PI2: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) x)
+        ((op <::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <::real => real => bool) (0::real) ((sin::real => real) x)))"
+  by (import transc SIN_POS_PI2)
+
+lemma COS_POS_PI2: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) x)
+        ((op <::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <::real => real => bool) (0::real) ((cos::real => real) x)))"
+  by (import transc COS_POS_PI2)
+
+lemma COS_POS_PI: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          x)
+        ((op <::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <::real => real => bool) (0::real) ((cos::real => real) x)))"
+  by (import transc COS_POS_PI)
+
+lemma SIN_POS_PI: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) x)
+        ((op <::real => real => bool) x (pi::real)))
+      ((op <::real => real => bool) (0::real) ((sin::real => real) x)))"
+  by (import transc SIN_POS_PI)
+
+lemma COS_POS_PI2_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op <=::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <=::real => real => bool) (0::real) ((cos::real => real) x)))"
+  by (import transc COS_POS_PI2_LE)
+
+lemma COS_POS_PI_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          x)
+        ((op <=::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <=::real => real => bool) (0::real) ((cos::real => real) x)))"
+  by (import transc COS_POS_PI_LE)
+
+lemma SIN_POS_PI2_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op <=::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <=::real => real => bool) (0::real) ((sin::real => real) x)))"
+  by (import transc SIN_POS_PI2_LE)
+
+lemma SIN_POS_PI_LE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op <=::real => real => bool) x (pi::real)))
+      ((op <=::real => real => bool) (0::real) ((sin::real => real) x)))"
+  by (import transc SIN_POS_PI_LE)
+
+lemma COS_TOTAL: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((Ex1::(real => bool) => bool)
+        (%x::real.
+            (op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <=::real => real => bool) x (pi::real))
+               ((op =::real => real => bool) ((cos::real => real) x) y)))))"
+  by (import transc COS_TOTAL)
+
+lemma SIN_TOTAL: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((Ex1::(real => bool) => bool)
+        (%x::real.
+            (op &::bool => bool => bool)
+             ((op <=::real => real => bool)
+               ((uminus::real => real)
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))
+               x)
+             ((op &::bool => bool => bool)
+               ((op <=::real => real => bool) x
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))
+               ((op =::real => real => bool) ((sin::real => real) x) y)))))"
+  by (import transc SIN_TOTAL)
+
+lemma COS_ZERO_LEMMA: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op =::real => real => bool) ((cos::real => real) x) (0::real)))
+      ((Ex::(nat => bool) => bool)
+        (%n::nat.
+            (op &::bool => bool => bool)
+             ((Not::bool => bool) ((EVEN::nat => bool) n))
+             ((op =::real => real => bool) x
+               ((op *::real => real => real) ((real::nat => real) n)
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))))))"
+  by (import transc COS_ZERO_LEMMA)
+
+lemma SIN_ZERO_LEMMA: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op =::real => real => bool) ((sin::real => real) x) (0::real)))
+      ((Ex::(nat => bool) => bool)
+        (%n::nat.
+            (op &::bool => bool => bool) ((EVEN::nat => bool) n)
+             ((op =::real => real => bool) x
+               ((op *::real => real => real) ((real::nat => real) n)
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))))))"
+  by (import transc SIN_ZERO_LEMMA)
+
+lemma COS_ZERO: "ALL x.
+   (cos x = 0) =
+   ((EX n. ~ EVEN n & x = real n * (pi / 2)) |
+    (EX n. ~ EVEN n & x = - (real n * (pi / 2))))"
+  by (import transc COS_ZERO)
+
+lemma SIN_ZERO: "ALL x.
+   (sin x = 0) =
+   ((EX n. EVEN n & x = real n * (pi / 2)) |
+    (EX n. EVEN n & x = - (real n * (pi / 2))))"
+  by (import transc SIN_ZERO)
+
+constdefs
+  tan :: "real => real" 
+  "tan == %x. sin x / cos x"
+
+lemma tan: "ALL x. tan x = sin x / cos x"
+  by (import transc tan)
+
+lemma TAN_0: "tan 0 = 0"
+  by (import transc TAN_0)
+
+lemma TAN_PI: "tan pi = 0"
+  by (import transc TAN_PI)
+
+lemma TAN_NPI: "ALL n::nat. tan (real n * pi) = (0::real)"
+  by (import transc TAN_NPI)
+
+lemma TAN_NEG: "ALL x. tan (- x) = - tan x"
+  by (import transc TAN_NEG)
+
+lemma TAN_PERIODIC: "ALL x. tan (x + 2 * pi) = tan x"
+  by (import transc TAN_PERIODIC)
+
+lemma TAN_ADD: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%y::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((Not::bool => bool)
+               ((op =::real => real => bool) ((cos::real => real) x)
+                 (0::real)))
+             ((op &::bool => bool => bool)
+               ((Not::bool => bool)
+                 ((op =::real => real => bool) ((cos::real => real) y)
+                   (0::real)))
+               ((Not::bool => bool)
+                 ((op =::real => real => bool)
+                   ((cos::real => real) ((op +::real => real => real) x y))
+                   (0::real)))))
+           ((op =::real => real => bool)
+             ((tan::real => real) ((op +::real => real => real) x y))
+             ((op /::real => real => real)
+               ((op +::real => real => real) ((tan::real => real) x)
+                 ((tan::real => real) y))
+               ((op -::real => real => real) (1::real)
+                 ((op *::real => real => real) ((tan::real => real) x)
+                   ((tan::real => real) y)))))))"
+  by (import transc TAN_ADD)
+
+lemma TAN_DOUBLE: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((Not::bool => bool)
+          ((op =::real => real => bool) ((cos::real => real) x) (0::real)))
+        ((Not::bool => bool)
+          ((op =::real => real => bool)
+            ((cos::real => real)
+              ((op *::real => real => real)
+                ((number_of::bin => real)
+                  ((op BIT::bin => bool => bin)
+                    ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                      (True::bool))
+                    (False::bool)))
+                x))
+            (0::real))))
+      ((op =::real => real => bool)
+        ((tan::real => real)
+          ((op *::real => real => real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))
+            x))
+        ((op /::real => real => real)
+          ((op *::real => real => real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))
+            ((tan::real => real) x))
+          ((op -::real => real => real) (1::real)
+            ((op ^::real => nat => real) ((tan::real => real) x)
+              ((number_of::bin => nat)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool))))))))"
+  by (import transc TAN_DOUBLE)
+
+lemma TAN_POS_PI2: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) x)
+        ((op <::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op <::real => real => bool) (0::real) ((tan::real => real) x)))"
+  by (import transc TAN_POS_PI2)
+
+lemma DIFF_TAN: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((Not::bool => bool)
+        ((op =::real => real => bool) ((cos::real => real) x) (0::real)))
+      ((diffl::(real => real) => real => real => bool) (tan::real => real)
+        ((inverse::real => real)
+          ((op ^::real => nat => real) ((cos::real => real) x)
+            ((number_of::bin => nat)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))
+        x))"
+  by (import transc DIFF_TAN)
+
+lemma TAN_TOTAL_LEMMA: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) y)
+      ((Ex::(real => bool) => bool)
+        (%x::real.
+            (op &::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <::real => real => bool) x
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))
+               ((op <::real => real => bool) y ((tan::real => real) x))))))"
+  by (import transc TAN_TOTAL_LEMMA)
+
+lemma TAN_TOTAL_POS: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) y)
+      ((Ex::(real => bool) => bool)
+        (%x::real.
+            (op &::bool => bool => bool)
+             ((op <=::real => real => bool) (0::real) x)
+             ((op &::bool => bool => bool)
+               ((op <::real => real => bool) x
+                 ((op /::real => real => real) (pi::real)
+                   ((number_of::bin => real)
+                     ((op BIT::bin => bool => bin)
+                       ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                         (True::bool))
+                       (False::bool)))))
+               ((op =::real => real => bool) ((tan::real => real) x) y)))))"
+  by (import transc TAN_TOTAL_POS)
+
+lemma TAN_TOTAL: "ALL y. EX! x. - (pi / 2) < x & x < pi / 2 & tan x = y"
+  by (import transc TAN_TOTAL)
+
+constdefs
+  asn :: "real => real" 
+  "asn == %y. SOME x. - (pi / 2) <= x & x <= pi / 2 & sin x = y"
+
+lemma asn: "ALL y. asn y = (SOME x. - (pi / 2) <= x & x <= pi / 2 & sin x = y)"
+  by (import transc asn)
+
+constdefs
+  acs :: "real => real" 
+  "acs == %y. SOME x. 0 <= x & x <= pi & cos x = y"
+
+lemma acs: "ALL y. acs y = (SOME x. 0 <= x & x <= pi & cos x = y)"
+  by (import transc acs)
+
+constdefs
+  atn :: "real => real" 
+  "atn == %y. SOME x. - (pi / 2) < x & x < pi / 2 & tan x = y"
+
+lemma atn: "ALL y. atn y = (SOME x. - (pi / 2) < x & x < pi / 2 & tan x = y)"
+  by (import transc atn)
+
+lemma ASN: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          ((asn::real => real) y))
+        ((op &::bool => bool => bool)
+          ((op <=::real => real => bool) ((asn::real => real) y)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          ((op =::real => real => bool)
+            ((sin::real => real) ((asn::real => real) y)) y))))"
+  by (import transc ASN)
+
+lemma ASN_SIN: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op =::real => real => bool)
+        ((sin::real => real) ((asn::real => real) y)) y))"
+  by (import transc ASN_SIN)
+
+lemma ASN_BOUNDS: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          ((asn::real => real) y))
+        ((op <=::real => real => bool) ((asn::real => real) y)
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))))"
+  by (import transc ASN_BOUNDS)
+
+lemma ASN_BOUNDS_LT: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          ((asn::real => real) y))
+        ((op <::real => real => bool) ((asn::real => real) y)
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))))"
+  by (import transc ASN_BOUNDS_LT)
+
+lemma SIN_ASN: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          x)
+        ((op <=::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op =::real => real => bool)
+        ((asn::real => real) ((sin::real => real) x)) x))"
+  by (import transc SIN_ASN)
+
+lemma ACS: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) ((acs::real => real) y))
+        ((op &::bool => bool => bool)
+          ((op <=::real => real => bool) ((acs::real => real) y) (pi::real))
+          ((op =::real => real => bool)
+            ((cos::real => real) ((acs::real => real) y)) y))))"
+  by (import transc ACS)
+
+lemma ACS_COS: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op =::real => real => bool)
+        ((cos::real => real) ((acs::real => real) y)) y))"
+  by (import transc ACS_COS)
+
+lemma ACS_BOUNDS: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <=::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) ((acs::real => real) y))
+        ((op <=::real => real => bool) ((acs::real => real) y) (pi::real))))"
+  by (import transc ACS_BOUNDS)
+
+lemma ACS_BOUNDS_LT: "(All::(real => bool) => bool)
+ (%y::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) y)
+        ((op <::real => real => bool) y (1::real)))
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) (0::real) ((acs::real => real) y))
+        ((op <::real => real => bool) ((acs::real => real) y) (pi::real))))"
+  by (import transc ACS_BOUNDS_LT)
+
+lemma COS_ACS: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op <=::real => real => bool) x (pi::real)))
+      ((op =::real => real => bool)
+        ((acs::real => real) ((cos::real => real) x)) x))"
+  by (import transc COS_ACS)
+
+lemma ATN: "ALL y. - (pi / 2) < atn y & atn y < pi / 2 & tan (atn y) = y"
+  by (import transc ATN)
+
+lemma ATN_TAN: "ALL x. tan (atn x) = x"
+  by (import transc ATN_TAN)
+
+lemma ATN_BOUNDS: "ALL x. - (pi / 2) < atn x & atn x < pi / 2"
+  by (import transc ATN_BOUNDS)
+
+lemma TAN_ATN: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          x)
+        ((op <::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op =::real => real => bool)
+        ((atn::real => real) ((tan::real => real) x)) x))"
+  by (import transc TAN_ATN)
+
+lemma TAN_SEC: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((Not::bool => bool)
+        ((op =::real => real => bool) ((cos::real => real) x) (0::real)))
+      ((op =::real => real => bool)
+        ((op +::real => real => real) (1::real)
+          ((op ^::real => nat => real) ((tan::real => real) x)
+            ((number_of::bin => nat)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool)))))
+        ((op ^::real => nat => real)
+          ((inverse::real => real) ((cos::real => real) x))
+          ((number_of::bin => nat)
+            ((op BIT::bin => bool => bin)
+              ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+              (False::bool))))))"
+  by (import transc TAN_SEC)
+
+lemma SIN_COS_SQ: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool) (0::real) x)
+        ((op <=::real => real => bool) x (pi::real)))
+      ((op =::real => real => bool) ((sin::real => real) x)
+        ((sqrt::real => real)
+          ((op -::real => real => real) (1::real)
+            ((op ^::real => nat => real) ((cos::real => real) x)
+              ((number_of::bin => nat)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool))))))))"
+  by (import transc SIN_COS_SQ)
+
+lemma COS_SIN_SQ: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <=::real => real => bool)
+          ((uminus::real => real)
+            ((op /::real => real => real) (pi::real)
+              ((number_of::bin => real)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool)))))
+          x)
+        ((op <=::real => real => bool) x
+          ((op /::real => real => real) (pi::real)
+            ((number_of::bin => real)
+              ((op BIT::bin => bool => bin)
+                ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                (False::bool))))))
+      ((op =::real => real => bool) ((cos::real => real) x)
+        ((sqrt::real => real)
+          ((op -::real => real => real) (1::real)
+            ((op ^::real => nat => real) ((sin::real => real) x)
+              ((number_of::bin => nat)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool))))))))"
+  by (import transc COS_SIN_SQ)
+
+lemma COS_ATN_NZ: "ALL x. cos (atn x) ~= 0"
+  by (import transc COS_ATN_NZ)
+
+lemma COS_ASN_NZ: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((Not::bool => bool)
+        ((op =::real => real => bool)
+          ((cos::real => real) ((asn::real => real) x)) (0::real))))"
+  by (import transc COS_ASN_NZ)
+
+lemma SIN_ACS_NZ: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((Not::bool => bool)
+        ((op =::real => real => bool)
+          ((sin::real => real) ((acs::real => real) x)) (0::real))))"
+  by (import transc SIN_ACS_NZ)
+
+lemma COS_SIN_SQRT: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) ((cos::real => real) x))
+      ((op =::real => real => bool) ((cos::real => real) x)
+        ((sqrt::real => real)
+          ((op -::real => real => real) (1::real)
+            ((op ^::real => nat => real) ((sin::real => real) x)
+              ((number_of::bin => nat)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool))))))))"
+  by (import transc COS_SIN_SQRT)
+
+lemma SIN_COS_SQRT: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <=::real => real => bool) (0::real) ((sin::real => real) x))
+      ((op =::real => real => bool) ((sin::real => real) x)
+        ((sqrt::real => real)
+          ((op -::real => real => real) (1::real)
+            ((op ^::real => nat => real) ((cos::real => real) x)
+              ((number_of::bin => nat)
+                ((op BIT::bin => bool => bin)
+                  ((op BIT::bin => bool => bin) (bin.Pls::bin) (True::bool))
+                  (False::bool))))))))"
+  by (import transc SIN_COS_SQRT)
+
+lemma DIFF_LN: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op <::real => real => bool) (0::real) x)
+      ((diffl::(real => real) => real => real => bool) (ln::real => real)
+        ((inverse::real => real) x) x))"
+  by (import transc DIFF_LN)
+
+lemma DIFF_ASN_LEMMA: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((diffl::(real => real) => real => real => bool) (asn::real => real)
+        ((inverse::real => real)
+          ((cos::real => real) ((asn::real => real) x)))
+        x))"
+  by (import transc DIFF_ASN_LEMMA)
+
+lemma DIFF_ASN: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((diffl::(real => real) => real => real => bool) (asn::real => real)
+        ((inverse::real => real)
+          ((sqrt::real => real)
+            ((op -::real => real => real) (1::real)
+              ((op ^::real => nat => real) x
+                ((number_of::bin => nat)
+                  ((op BIT::bin => bool => bin)
+                    ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                      (True::bool))
+                    (False::bool)))))))
+        x))"
+  by (import transc DIFF_ASN)
+
+lemma DIFF_ACS_LEMMA: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((diffl::(real => real) => real => real => bool) (acs::real => real)
+        ((inverse::real => real)
+          ((uminus::real => real)
+            ((sin::real => real) ((acs::real => real) x))))
+        x))"
+  by (import transc DIFF_ACS_LEMMA)
+
+lemma DIFF_ACS: "(All::(real => bool) => bool)
+ (%x::real.
+     (op -->::bool => bool => bool)
+      ((op &::bool => bool => bool)
+        ((op <::real => real => bool) ((uminus::real => real) (1::real)) x)
+        ((op <::real => real => bool) x (1::real)))
+      ((diffl::(real => real) => real => real => bool) (acs::real => real)
+        ((uminus::real => real)
+          ((inverse::real => real)
+            ((sqrt::real => real)
+              ((op -::real => real => real) (1::real)
+                ((op ^::real => nat => real) x
+                  ((number_of::bin => nat)
+                    ((op BIT::bin => bool => bin)
+                      ((op BIT::bin => bool => bin) (bin.Pls::bin)
+                        (True::bool))
+                      (False::bool))))))))
+        x))"
+  by (import transc DIFF_ACS)
+
+lemma DIFF_ATN: "ALL x. diffl atn (inverse (1 + x ^ 2)) x"
+  by (import transc DIFF_ATN)
+
+constdefs
+  division :: "real * real => (nat => real) => bool" 
+  "(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)))))))))"
+  by (import transc division)
+
+constdefs
+  dsize :: "(nat => real) => nat" 
+  "(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)))))))"
+  by (import transc dsize)
+
+constdefs
+  tdiv :: "real * real => (nat => real) * (nat => real) => bool" 
+  "tdiv ==
+%(a, b) (D, p). division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n))"
+
+lemma tdiv: "ALL a b D p.
+   tdiv (a, b) (D, p) =
+   (division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n)))"
+  by (import transc tdiv)
+
+constdefs
+  gauge :: "(real => bool) => (real => real) => bool" 
+  "(op ==::((real => bool) => (real => real) => bool)
+        => ((real => bool) => (real => real) => bool) => prop)
+ (gauge::(real => bool) => (real => real) => bool)
+ (%(E::real => bool) g::real => real.
+     (All::(real => bool) => bool)
+      (%x::real.
+          (op -->::bool => bool => bool) (E x)
+           ((op <::real => real => bool) (0::real) (g x))))"
+
+lemma gauge: "(All::((real => bool) => bool) => bool)
+ (%E::real => bool.
+     (All::((real => real) => bool) => bool)
+      (%g::real => real.
+          (op =::bool => bool => bool)
+           ((gauge::(real => bool) => (real => real) => bool) E g)
+           ((All::(real => bool) => bool)
+             (%x::real.
+                 (op -->::bool => bool => bool) (E x)
+                  ((op <::real => real => bool) (0::real) (g x))))))"
+  by (import transc gauge)
+
+constdefs
+  fine :: "(real => real) => (nat => real) * (nat => real) => bool" 
+  "fine == %g (D, p). ALL n<dsize D. D (Suc n) - D n < g (p n)"
+
+lemma fine: "ALL g D p. fine g (D, p) = (ALL n<dsize D. D (Suc n) - D n < g (p n))"
+  by (import transc fine)
+
+constdefs
+  rsum :: "(nat => real) * (nat => real) => (real => real) => real" 
+  "rsum == %(D, p) f. real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))"
+
+lemma rsum: "ALL D p f.
+   rsum (D, p) f = real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))"
+  by (import transc rsum)
+
+constdefs
+  Dint :: "real * real => (real => real) => real => bool" 
+  "(op ==::(real * real => (real => real) => real => bool)
+        => (real * real => (real => real) => real => bool) => prop)
+ (Dint::real * real => (real => real) => real => bool)
+ ((split::(real => real => (real => real) => real => bool)
+          => real * real => (real => real) => real => bool)
+   (%(a::real) (b::real) (f::real => real) k::real.
+       (All::(real => bool) => bool)
+        (%e::real.
+            (op -->::bool => bool => bool)
+             ((op <::real => real => bool) (0::real) e)
+             ((Ex::((real => real) => bool) => bool)
+               (%g::real => real.
+                   (op &::bool => bool => bool)
+                    ((gauge::(real => bool) => (real => real) => bool)
+                      (%x::real.
+                          (op &::bool => bool => bool)
+                           ((op <=::real => real => bool) a x)
+                           ((op <=::real => real => bool) x b))
+                      g)
+                    ((All::((nat => real) => bool) => bool)
+                      (%D::nat => real.
+                          (All::((nat => real) => bool) => bool)
+                           (%p::nat => real.
+                               (op -->::bool => bool => bool)
+                                ((op &::bool => bool => bool)
+                                  ((tdiv::real * real
+    => (nat => real) * (nat => real) => bool)
+                                    ((Pair::real => real => real * real) a
+b)
+                                    ((Pair::(nat => real)
+      => (nat => real) => (nat => real) * (nat => real))
+D p))
+                                  ((fine::(real => real)
+    => (nat => real) * (nat => real) => bool)
+                                    g ((Pair::(nat => real)
+        => (nat => real) => (nat => real) * (nat => real))
+  D p)))
+                                ((op <::real => real => bool)
+                                  ((abs::real => real)
+                                    ((op -::real => real => real)
+((rsum::(nat => real) * (nat => real) => (real => real) => real)
+  ((Pair::(nat => real) => (nat => real) => (nat => real) * (nat => real)) D
+    p)
+  f)
+k))
+                                  e)))))))))"
+
+lemma Dint: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (All::((real => real) => bool) => bool)
+           (%f::real => real.
+               (All::(real => bool) => bool)
+                (%k::real.
+                    (op =::bool => bool => bool)
+                     ((Dint::real * real => (real => real) => real => bool)
+                       ((Pair::real => real => real * real) a b) f k)
+                     ((All::(real => bool) => bool)
+                       (%e::real.
+                           (op -->::bool => bool => bool)
+                            ((op <::real => real => bool) (0::real) e)
+                            ((Ex::((real => real) => bool) => bool)
+                              (%g::real => real.
+                                  (op &::bool => bool => bool)
+                                   ((gauge::(real => bool)
+      => (real => real) => bool)
+                                     (%x::real.
+   (op &::bool => bool => bool) ((op <=::real => real => bool) a x)
+    ((op <=::real => real => bool) x b))
+                                     g)
+                                   ((All::((nat => real) => bool) => bool)
+                                     (%D::nat => real.
+   (All::((nat => real) => bool) => bool)
+    (%p::nat => real.
+        (op -->::bool => bool => bool)
+         ((op &::bool => bool => bool)
+           ((tdiv::real * real => (nat => real) * (nat => real) => bool)
+             ((Pair::real => real => real * real) a b)
+             ((Pair::(nat => real)
+                     => (nat => real) => (nat => real) * (nat => real))
+               D p))
+           ((fine::(real => real) => (nat => real) * (nat => real) => bool)
+             g ((Pair::(nat => real)
+                       => (nat => real) => (nat => real) * (nat => real))
+                 D p)))
+         ((op <::real => real => bool)
+           ((abs::real => real)
+             ((op -::real => real => real)
+               ((rsum::(nat => real) * (nat => real)
+                       => (real => real) => real)
+                 ((Pair::(nat => real)
+                         => (nat => real) => (nat => real) * (nat => real))
+                   D p)
+                 f)
+               k))
+           e))))))))))))"
+  by (import transc Dint)
+
+lemma DIVISION_0: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (op -->::bool => bool => bool) ((op =::real => real => bool) a b)
+           ((op =::nat => nat => bool)
+             ((dsize::(nat => real) => nat)
+               (%n::nat.
+                   (If::bool => real => real => real)
+                    ((op =::nat => nat => bool) n (0::nat)) a b))
+             (0::nat))))"
+  by (import transc DIVISION_0)
+
+lemma DIVISION_1: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (op -->::bool => bool => bool) ((op <::real => real => bool) a b)
+           ((op =::nat => nat => bool)
+             ((dsize::(nat => real) => nat)
+               (%n::nat.
+                   (If::bool => real => real => real)
+                    ((op =::nat => nat => bool) n (0::nat)) a b))
+             (1::nat))))"
+  by (import transc DIVISION_1)
+
+lemma DIVISION_SINGLE: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (op -->::bool => bool => bool) ((op <=::real => real => bool) a b)
+           ((division::real * real => (nat => real) => bool)
+             ((Pair::real => real => real * real) a b)
+             (%n::nat.
+                 (If::bool => real => real => real)
+                  ((op =::nat => nat => bool) n (0::nat)) a b))))"
+  by (import transc DIVISION_SINGLE)
+
+lemma DIVISION_LHS: "(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 =::real => real => bool) (D (0::nat)) 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))))))))"
+  by (import transc DIVISION_THM)
+
+lemma DIVISION_RHS: "(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 =::real => real => bool)
+                  (D ((dsize::(nat => real) => nat) D)) b))))"
+  by (import transc DIVISION_RHS)
+
+lemma DIVISION_LT_GEN: "(All::((nat => real) => bool) => bool)
+ (%D::nat => real.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(real => bool) => bool)
+           (%b::real.
+               (All::(nat => bool) => bool)
+                (%m::nat.
+                    (All::(nat => bool) => bool)
+                     (%n::nat.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((division::real * real
+  => (nat => real) => bool)
+                              ((Pair::real => real => real * real) a b) D)
+                            ((op &::bool => bool => bool)
+                              ((op <::nat => nat => bool) m n)
+                              ((op <=::nat => nat => bool) n
+                                ((dsize::(nat => real) => nat) D))))
+                          ((op <::real => real => bool) (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))))))))"
+  by (import transc DIVISION_LT)
+
+lemma DIVISION_LE: "(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 <=::real => real => bool) 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))))))))"
+  by (import transc DIVISION_GT)
+
+lemma DIVISION_EQ: "(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) a b)
+                  ((op =::nat => nat => bool)
+                    ((dsize::(nat => real) => nat) D) (0::nat))))))"
+  by (import transc DIVISION_EQ)
+
+lemma DIVISION_LBOUND: "(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)
+                  (%r::nat. (op <=::real => real => bool) a (D r))))))"
+  by (import transc DIVISION_LBOUND)
+
+lemma DIVISION_LBOUND_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)
+                ((op &::bool => bool => bool)
+                  ((division::real * real => (nat => real) => bool)
+                    ((Pair::real => real => real * real) a b) D)
+                  ((Not::bool => bool)
+                    ((op =::nat => nat => bool)
+                      ((dsize::(nat => real) => nat) D) (0::nat))))
+                ((All::(nat => bool) => bool)
+                  (%n::nat.
+                      (op <::real => real => bool) a
+                       (D ((Suc::nat => nat) n)))))))"
+  by (import transc DIVISION_LBOUND_LT)
+
+lemma DIVISION_UBOUND: "(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)
+                  (%r::nat. (op <=::real => real => bool) (D r) b)))))"
+  by (import transc DIVISION_UBOUND)
+
+lemma DIVISION_UBOUND_LT: "(All::((nat => real) => bool) => bool)
+ (%D::nat => real.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(real => bool) => bool)
+           (%b::real.
+               (All::(nat => bool) => bool)
+                (%n::nat.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((division::real * real => (nat => real) => bool)
+                         ((Pair::real => real => real * real) a b) D)
+                       ((op <::nat => nat => bool) n
+                         ((dsize::(nat => real) => nat) D)))
+                     ((op <::real => real => bool) (D n) b)))))"
+  by (import transc DIVISION_UBOUND_LT)
+
+lemma DIVISION_APPEND: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (All::(real => bool) => bool)
+           (%c::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((Ex::((nat => real) => bool) => bool)
+                    (%D1::nat => real.
+                        (Ex::((nat => real) => bool) => bool)
+                         (%p1::nat => real.
+                             (op &::bool => bool => bool)
+                              ((tdiv::real * real
+=> (nat => real) * (nat => real) => bool)
+                                ((Pair::real => real => real * real) a b)
+                                ((Pair::(nat => real)
+  => (nat => real) => (nat => real) * (nat => real))
+                                  D1 p1))
+                              ((fine::(real => real)
+=> (nat => real) * (nat => real) => bool)
+                                (g::real => real)
+                                ((Pair::(nat => real)
+  => (nat => real) => (nat => real) * (nat => real))
+                                  D1 p1)))))
+                  ((Ex::((nat => real) => bool) => bool)
+                    (%D2::nat => real.
+                        (Ex::((nat => real) => bool) => bool)
+                         (%p2::nat => real.
+                             (op &::bool => bool => bool)
+                              ((tdiv::real * real
+=> (nat => real) * (nat => real) => bool)
+                                ((Pair::real => real => real * real) b c)
+                                ((Pair::(nat => real)
+  => (nat => real) => (nat => real) * (nat => real))
+                                  D2 p2))
+                              ((fine::(real => real)
+=> (nat => real) * (nat => real) => bool)
+                                g ((Pair::(nat => real)
+    => (nat => real) => (nat => real) * (nat => real))
+                                    D2 p2))))))
+                ((Ex::((nat => real) => bool) => bool)
+                  (%x::nat => real.
+                      (Ex::((nat => real) => bool) => bool)
+                       (%p::nat => real.
+                           (op &::bool => bool => bool)
+                            ((tdiv::real * real
+                                    => (nat => real) * (nat => real)
+ => bool)
+                              ((Pair::real => real => real * real) a c)
+                              ((Pair::(nat => real)
+=> (nat => real) => (nat => real) * (nat => real))
+                                x p))
+                            ((fine::(real => real)
+                                    => (nat => real) * (nat => real)
+ => bool)
+                              g ((Pair::(nat => real)
+  => (nat => real) => (nat => real) * (nat => real))
+                                  x p))))))))"
+  by (import transc DIVISION_APPEND)
+
+lemma DIVISION_EXISTS: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (All::((real => real) => bool) => bool)
+           (%g::real => real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <=::real => real => bool) a b)
+                  ((gauge::(real => bool) => (real => real) => bool)
+                    (%x::real.
+                        (op &::bool => bool => bool)
+                         ((op <=::real => real => bool) a x)
+                         ((op <=::real => real => bool) x b))
+                    g))
+                ((Ex::((nat => real) => bool) => bool)
+                  (%D::nat => real.
+                      (Ex::((nat => real) => bool) => bool)
+                       (%p::nat => real.
+                           (op &::bool => bool => bool)
+                            ((tdiv::real * real
+                                    => (nat => real) * (nat => real)
+ => bool)
+                              ((Pair::real => real => real * real) a b)
+                              ((Pair::(nat => real)
+=> (nat => real) => (nat => real) * (nat => real))
+                                D p))
+                            ((fine::(real => real)
+                                    => (nat => real) * (nat => real)
+ => bool)
+                              g ((Pair::(nat => real)
+  => (nat => real) => (nat => real) * (nat => real))
+                                  D p))))))))"
+  by (import transc DIVISION_EXISTS)
+
+lemma GAUGE_MIN: "(All::((real => bool) => bool) => bool)
+ (%E::real => bool.
+     (All::((real => real) => bool) => bool)
+      (%g1::real => real.
+          (All::((real => real) => bool) => bool)
+           (%g2::real => real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((gauge::(real => bool) => (real => real) => bool) E g1)
+                  ((gauge::(real => bool) => (real => real) => bool) E g2))
+                ((gauge::(real => bool) => (real => real) => bool) E
+                  (%x::real.
+                      (If::bool => real => real => real)
+                       ((op <::real => real => bool) (g1 x) (g2 x)) (g1 x)
+                       (g2 x))))))"
+  by (import transc GAUGE_MIN)
+
+lemma FINE_MIN: "(All::((real => real) => bool) => bool)
+ (%g1::real => real.
+     (All::((real => real) => bool) => bool)
+      (%g2::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)
+                       (%x::real.
+                           (If::bool => real => real => real)
+                            ((op <::real => real => bool) (g1 x) (g2 x))
+                            (g1 x) (g2 x))
+                       ((Pair::(nat => real)
+                               => (nat => real)
+                                  => (nat => real) * (nat => real))
+                         D p))
+                     ((op &::bool => bool => bool)
+                       ((fine::(real => real)
+                               => (nat => real) * (nat => real) => bool)
+                         g1 ((Pair::(nat => real)
+                                    => (nat => real)
+ => (nat => real) * (nat => real))
+                              D p))
+                       ((fine::(real => real)
+                               => (nat => real) * (nat => real) => bool)
+                         g2 ((Pair::(nat => real)
+                                    => (nat => real)
+ => (nat => real) * (nat => real))
+                              D p)))))))"
+  by (import transc FINE_MIN)
+
+lemma DINT_UNIQ: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real => bool) => bool)
+      (%b::real.
+          (All::((real => real) => bool) => bool)
+           (%f::real => real.
+               (All::(real => bool) => bool)
+                (%k1::real.
+                    (All::(real => bool) => bool)
+                     (%k2::real.
+                         (op -->::bool => bool => bool)
+                          ((op &::bool => bool => bool)
+                            ((op <=::real => real => bool) a b)
+                            ((op &::bool => bool => bool)
+                              ((Dint::real * real
+=> (real => real) => real => bool)
+                                ((Pair::real => real => real * real) a b) f
+                                k1)
+                              ((Dint::real * real
+=> (real => real) => real => bool)
+                                ((Pair::real => real => real * real) a b) f
+                                k2)))
+                          ((op =::real => real => bool) k1 k2))))))"
+  by (import transc DINT_UNIQ)
+
+lemma INTEGRAL_NULL: "ALL f a. Dint (a, a) f 0"
+  by (import transc INTEGRAL_NULL)
+
+lemma FTC1: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (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))
+                              ((diffl::(real => real)
+ => real => real => bool)
+                                f (f' x) x))))
+                     ((Dint::real * real => (real => real) => real => bool)
+                       ((Pair::real => real => real * real) a b) f'
+                       ((op -::real => real => real) (f b) (f a)))))))"
+  by (import transc FTC1)
+
+lemma MCLAURIN: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((nat => real => real) => bool) => bool)
+      (%diff::nat => real => real.
+          (All::(real => bool) => bool)
+           (%h::real.
+               (All::(nat => bool) => bool)
+                (%n::nat.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) (0::real) h)
+                       ((op &::bool => bool => bool)
+                         ((op <::nat => nat => bool) (0::nat) n)
+                         ((op &::bool => bool => bool)
+                           ((op =::(real => real) => (real => real) => bool)
+                             (diff (0::nat)) f)
+                           ((All::(nat => bool) => bool)
+                             (%m::nat.
+                                 (All::(real => bool) => bool)
+                                  (%t::real.
+(op -->::bool => bool => bool)
+ ((op &::bool => bool => bool) ((op <::nat => nat => bool) m n)
+   ((op &::bool => bool => bool) ((op <=::real => real => bool) (0::real) t)
+     ((op <=::real => real => bool) t h)))
+ ((diffl::(real => real) => real => real => bool) (diff m)
+   (diff ((Suc::nat => nat) m) t) t)))))))
+                     ((Ex::(real => bool) => bool)
+                       (%t::real.
+                           (op &::bool => bool => bool)
+                            ((op <::real => real => bool) (0::real) t)
+                            ((op &::bool => bool => bool)
+                              ((op <::real => real => bool) t h)
+                              ((op =::real => real => bool) (f h)
+                                ((op +::real => real => real)
+                                  ((real.sum::nat * nat
+        => (nat => real) => real)
+                                    ((Pair::nat => nat => nat * nat)
+(0::nat) n)
+                                    (%m::nat.
+  (op *::real => real => real)
+   ((op /::real => real => real) (diff m (0::real))
+     ((real::nat => real) ((FACT::nat => nat) m)))
+   ((op ^::real => nat => real) h m)))
+                                  ((op *::real => real => real)
+                                    ((op /::real => real => real) (diff n t)
+((real::nat => real) ((FACT::nat => nat) n)))
+                                    ((op ^::real => nat => real) h
+n)))))))))))"
+  by (import transc MCLAURIN)
+
+lemma MCLAURIN_NEG: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((nat => real => real) => bool) => bool)
+      (%diff::nat => real => real.
+          (All::(real => bool) => bool)
+           (%h::real.
+               (All::(nat => bool) => bool)
+                (%n::nat.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((op <::real => real => bool) h (0::real))
+                       ((op &::bool => bool => bool)
+                         ((op <::nat => nat => bool) (0::nat) n)
+                         ((op &::bool => bool => bool)
+                           ((op =::(real => real) => (real => real) => bool)
+                             (diff (0::nat)) f)
+                           ((All::(nat => bool) => bool)
+                             (%m::nat.
+                                 (All::(real => bool) => bool)
+                                  (%t::real.
+(op -->::bool => bool => bool)
+ ((op &::bool => bool => bool) ((op <::nat => nat => bool) m n)
+   ((op &::bool => bool => bool) ((op <=::real => real => bool) h t)
+     ((op <=::real => real => bool) t (0::real))))
+ ((diffl::(real => real) => real => real => bool) (diff m)
+   (diff ((Suc::nat => nat) m) t) t)))))))
+                     ((Ex::(real => bool) => bool)
+                       (%t::real.
+                           (op &::bool => bool => bool)
+                            ((op <::real => real => bool) h t)
+                            ((op &::bool => bool => bool)
+                              ((op <::real => real => bool) t (0::real))
+                              ((op =::real => real => bool) (f h)
+                                ((op +::real => real => real)
+                                  ((real.sum::nat * nat
+        => (nat => real) => real)
+                                    ((Pair::nat => nat => nat * nat)
+(0::nat) n)
+                                    (%m::nat.
+  (op *::real => real => real)
+   ((op /::real => real => real) (diff m (0::real))
+     ((real::nat => real) ((FACT::nat => nat) m)))
+   ((op ^::real => nat => real) h m)))
+                                  ((op *::real => real => real)
+                                    ((op /::real => real => real) (diff n t)
+((real::nat => real) ((FACT::nat => nat) n)))
+                                    ((op ^::real => nat => real) h
+n)))))))))))"
+  by (import transc MCLAURIN_NEG)
+
+lemma MCLAURIN_ALL_LT: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((nat => real => real) => bool) => bool)
+      (%diff::nat => real => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op =::(real => real) => (real => real) => bool)
+               (diff (0::nat)) f)
+             ((All::(nat => bool) => bool)
+               (%m::nat.
+                   (All::(real => bool) => bool)
+                    (%x::real.
+                        (diffl::(real => real) => real => real => bool)
+                         (diff m) (diff ((Suc::nat => nat) m) x) x))))
+           ((All::(real => bool) => bool)
+             (%x::real.
+                 (All::(nat => bool) => bool)
+                  (%n::nat.
+                      (op -->::bool => bool => bool)
+                       ((op &::bool => bool => bool)
+                         ((Not::bool => bool)
+                           ((op =::real => real => bool) x (0::real)))
+                         ((op <::nat => nat => bool) (0::nat) n))
+                       ((Ex::(real => bool) => bool)
+                         (%t::real.
+                             (op &::bool => bool => bool)
+                              ((op <::real => real => bool) (0::real)
+                                ((abs::real => real) t))
+                              ((op &::bool => bool => bool)
+                                ((op <::real => real => bool)
+                                  ((abs::real => real) t)
+                                  ((abs::real => real) x))
+                                ((op =::real => real => bool) (f x)
+                                  ((op +::real => real => real)
+                                    ((real.sum::nat * nat
+          => (nat => real) => real)
+((Pair::nat => nat => nat * nat) (0::nat) n)
+(%m::nat.
+    (op *::real => real => real)
+     ((op /::real => real => real) (diff m (0::real))
+       ((real::nat => real) ((FACT::nat => nat) m)))
+     ((op ^::real => nat => real) x m)))
+                                    ((op *::real => real => real)
+((op /::real => real => real) (diff n t)
+  ((real::nat => real) ((FACT::nat => nat) n)))
+((op ^::real => nat => real) x n))))))))))))"
+  by (import transc MCLAURIN_ALL_LT)
+
+lemma MCLAURIN_ZERO: "(All::((nat => real => real) => bool) => bool)
+ (%diff::nat => real => real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op =::real => real => bool) x (0::real))
+                  ((op <::nat => nat => bool) (0::nat) n))
+                ((op =::real => real => bool)
+                  ((real.sum::nat * nat => (nat => real) => real)
+                    ((Pair::nat => nat => nat * nat) (0::nat) n)
+                    (%m::nat.
+                        (op *::real => real => real)
+                         ((op /::real => real => real) (diff m (0::real))
+                           ((real::nat => real) ((FACT::nat => nat) m)))
+                         ((op ^::real => nat => real) x m)))
+                  (diff (0::nat) (0::real))))))"
+  by (import transc MCLAURIN_ZERO)
+
+lemma MCLAURIN_ALL_LE: "(All::((real => real) => bool) => bool)
+ (%f::real => real.
+     (All::((nat => real => real) => bool) => bool)
+      (%diff::nat => real => real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((op =::(real => real) => (real => real) => bool)
+               (diff (0::nat)) f)
+             ((All::(nat => bool) => bool)
+               (%m::nat.
+                   (All::(real => bool) => bool)
+                    (%x::real.
+                        (diffl::(real => real) => real => real => bool)
+                         (diff m) (diff ((Suc::nat => nat) m) x) x))))
+           ((All::(real => bool) => bool)
+             (%x::real.
+                 (All::(nat => bool) => bool)
+                  (%n::nat.
+                      (Ex::(real => bool) => bool)
+                       (%t::real.
+                           (op &::bool => bool => bool)
+                            ((op <=::real => real => bool)
+                              ((abs::real => real) t)
+                              ((abs::real => real) x))
+                            ((op =::real => real => bool) (f x)
+                              ((op +::real => real => real)
+                                ((real.sum::nat * nat
+      => (nat => real) => real)
+                                  ((Pair::nat => nat => nat * nat) (0::nat)
+                                    n)
+                                  (%m::nat.
+(op *::real => real => real)
+ ((op /::real => real => real) (diff m (0::real))
+   ((real::nat => real) ((FACT::nat => nat) m)))
+ ((op ^::real => nat => real) x m)))
+                                ((op *::real => real => real)
+                                  ((op /::real => real => real) (diff n t)
+                                    ((real::nat => real)
+((FACT::nat => nat) n)))
+                                  ((op ^::real => nat => real) x n))))))))))"
+  by (import transc MCLAURIN_ALL_LE)
+
+lemma MCLAURIN_EXP_LT: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(nat => bool) => bool)
+      (%n::nat.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((Not::bool => bool)
+               ((op =::real => real => bool) x (0::real)))
+             ((op <::nat => nat => bool) (0::nat) n))
+           ((Ex::(real => bool) => bool)
+             (%xa::real.
+                 (op &::bool => bool => bool)
+                  ((op <::real => real => bool) (0::real)
+                    ((abs::real => real) xa))
+                  ((op &::bool => bool => bool)
+                    ((op <::real => real => bool) ((abs::real => real) xa)
+                      ((abs::real => real) x))
+                    ((op =::real => real => bool) ((exp::real => real) x)
+                      ((op +::real => real => real)
+                        ((real.sum::nat * nat => (nat => real) => real)
+                          ((Pair::nat => nat => nat * nat) (0::nat) n)
+                          (%m::nat.
+                              (op /::real => real => real)
+                               ((op ^::real => nat => real) x m)
+                               ((real::nat => real)
+                                 ((FACT::nat => nat) m))))
+                        ((op *::real => real => real)
+                          ((op /::real => real => real)
+                            ((exp::real => real) xa)
+                            ((real::nat => real) ((FACT::nat => nat) n)))
+                          ((op ^::real => nat => real) x n)))))))))"
+  by (import transc MCLAURIN_EXP_LT)
+
+lemma MCLAURIN_EXP_LE: "ALL x n.
+   EX xa.
+      abs xa <= abs x &
+      exp x =
+      real.sum (0, n) (%m. x ^ m / real (FACT m)) +
+      exp xa / real (FACT n) * x ^ n"
+  by (import transc MCLAURIN_EXP_LE)
+
+lemma DIFF_LN_COMPOSITE: "(All::((real => real) => bool) => bool)
+ (%g::real => real.
+     (All::(real => bool) => bool)
+      (%m::real.
+          (All::(real => bool) => bool)
+           (%x::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((diffl::(real => real) => real => real => bool) g m x)
+                  ((op <::real => real => bool) (0::real) (g x)))
+                ((diffl::(real => real) => real => real => bool)
+                  (%x::real. (ln::real => real) (g x))
+                  ((op *::real => real => real)
+                    ((inverse::real => real) (g x)) m)
+                  x))))"
+  by (import transc DIFF_LN_COMPOSITE)
+
+;end_setup
+
+;setup_theory poly
+
+consts
+  poly :: "real list => real => real" 
+
+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. 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)
+
+consts
+  "##" :: "real => real list => real list" ("##")
+
+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
+  poly_neg :: "real list => real list" 
+
+defs
+  poly_neg_primdef: "poly_neg == ## (- 1)"
+
+lemma poly_neg_def: "poly_neg = ## (- 1)"
+  by (import poly poly_neg_def)
+
+consts
+  poly_mul :: "real list => real list => 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)
+
+consts
+  poly_exp :: "real list => nat => real list" 
+
+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. 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)
+
+constdefs
+  diff :: "real list => real list" 
+  "diff == %l. if l = [] then [] else poly_diff_aux 1 (tl l)"
+
+lemma poly_diff_def: "ALL l. 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 = 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 [] = [] & ## c (h # t) = c * h # ## c t"
+  by (import poly POLY_CMUL_CLAUSES)
+
+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 = [] &
+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] = [] & diff (h # t) = poly_diff_aux 1 t"
+  by (import poly POLY_DIFF_CLAUSES)
+
+lemma POLY_ADD: "ALL t p2 x. poly (poly_add t p2) x = poly t x + poly p2 x"
+  by (import poly POLY_ADD)
+
+lemma POLY_CMUL: "ALL t c x. poly (## c t) x = c * poly t x"
+  by (import poly POLY_CMUL)
+
+lemma POLY_NEG: "ALL x xa. poly (poly_neg x) xa = - poly x xa"
+  by (import poly POLY_NEG)
+
+lemma POLY_MUL: "ALL x t p2. poly (poly_mul t p2) x = poly t x * poly p2 x"
+  by (import poly POLY_MUL)
+
+lemma POLY_EXP: "ALL p n x. poly (poly_exp p n) x = poly p x ^ n"
+  by (import poly POLY_EXP)
+
+lemma POLY_DIFF_LEMMA: "ALL t n x.
+   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 x. diffl (poly t) (poly (diff t) x) x"
+  by (import poly POLY_DIFF)
+
+lemma POLY_DIFFERENTIABLE: "ALL l. All (differentiable (poly l))"
+  by (import poly POLY_DIFFERENTIABLE)
+
+lemma POLY_CONT: "ALL l. All (contl (poly l))"
+  by (import poly POLY_CONT)
+
+lemma POLY_IVT_POS: "(All::(real list => bool) => bool)
+ (%x::real list.
+     (All::(real => bool) => bool)
+      (%xa::real.
+          (All::(real => bool) => bool)
+           (%xb::real.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op <::real => real => bool) xa xb)
+                  ((op &::bool => bool => bool)
+                    ((op <::real => real => bool)
+                      ((poly::real list => real => real) x xa) (0::real))
+                    ((op <::real => real => bool) (0::real)
+                      ((poly::real list => real => real) x xb))))
+                ((Ex::(real => bool) => bool)
+                  (%xc::real.
+                      (op &::bool => bool => bool)
+                       ((op <::real => real => bool) xa xc)
+                       ((op &::bool => bool => bool)
+                         ((op <::real => real => bool) xc xb)
+                         ((op =::real => real => bool)
+                           ((poly::real list => real => real) x xc)
+                           (0::real))))))))"
+  by (import poly POLY_IVT_POS)
+
+lemma POLY_IVT_NEG: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (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)
+                  ((op &::bool => bool => bool)
+                    ((op <::real => real => bool) (0::real)
+                      ((poly::real list => real => real) p a))
+                    ((op <::real => real => bool)
+                      ((poly::real list => real => real) p b) (0::real))))
+                ((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)
+                           ((poly::real list => real => real) p x)
+                           (0::real))))))))"
+  by (import poly POLY_IVT_NEG)
+
+lemma POLY_MVT: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(real => bool) => bool)
+           (%b::real.
+               (op -->::bool => bool => bool)
+                ((op <::real => real => bool) a b)
+                ((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)
+                           ((op -::real => real => real)
+                             ((poly::real list => real => real) p b)
+                             ((poly::real list => real => real) p a))
+                           ((op *::real => real => real)
+                             ((op -::real => real => real) b a)
+                             ((poly::real list => real => real)
+                               ((diff::real list => real list) p) x)))))))))"
+  by (import poly POLY_MVT)
+
+lemma POLY_ADD_RZERO: "ALL x. poly (poly_add x []) = poly x"
+  by (import poly POLY_ADD_RZERO)
+
+lemma POLY_MUL_ASSOC: "ALL x xa xb.
+   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 xa xb.
+   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 p2 n.
+   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 c n. 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 xa.
+   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 n.
+   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 p2. 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 c. poly (diff (## c t)) = poly (## c (diff t))"
+  by (import poly POLY_DIFF_CMUL)
+
+lemma POLY_DIFF_NEG: "ALL x. poly (diff (poly_neg x)) = poly (poly_neg (diff x))"
+  by (import poly POLY_DIFF_NEG)
+
+lemma POLY_DIFF_MUL_LEMMA: "ALL x xa. poly (diff (xa # x)) = poly (poly_add (0 # diff x) x)"
+  by (import poly POLY_DIFF_MUL_LEMMA)
+
+lemma POLY_DIFF_MUL: "ALL t p2.
+   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 n.
+   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 a.
+   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 h. 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 t. (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. 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)))))))))"
+  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))))))))"
+  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)))))))))"
+  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)))))))"
+  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)))))))))"
+  by (import poly POLY_ROOTS_FINITE)
+
+lemma POLY_ENTIRE_LEMMA: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::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))))
+             ((Not::bool => bool)
+               ((op =::(real => real) => (real => real) => bool)
+                 ((poly::real list => real => real) q)
+                 ((poly::real list => real => real) ([]::real list)))))
+           ((Not::bool => bool)
+             ((op =::(real => real) => (real => real) => bool)
+               ((poly::real list => real => real)
+                 ((poly_mul::real list => real list => real list) p q))
+               ((poly::real list => real => real) ([]::real list))))))"
+  by (import poly POLY_ENTIRE_LEMMA)
+
+lemma POLY_ENTIRE: "ALL p q.
+   (poly (poly_mul p q) = poly []) = (poly p = poly [] | poly q = poly [])"
+  by (import poly POLY_ENTIRE)
+
+lemma POLY_MUL_LCANCEL: "ALL x xa xb.
+   (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 n. (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. poly [a, 1] ~= poly []"
+  by (import poly POLY_PRIME_EQ_0)
+
+lemma POLY_EXP_PRIME_EQ_0: "ALL a n. poly (poly_exp [a, 1] n) ~= poly []"
+  by (import poly POLY_EXP_PRIME_EQ_0)
+
+lemma POLY_ZERO_LEMMA: "(All::(real => bool) => bool)
+ (%h::real.
+     (All::(real list => bool) => bool)
+      (%t::real list.
+          (op -->::bool => bool => bool)
+           ((op =::(real => real) => (real => real) => bool)
+             ((poly::real list => real => real)
+               ((op #::real => real list => real list) h t))
+             ((poly::real list => real => real) ([]::real list)))
+           ((op &::bool => bool => bool)
+             ((op =::real => real => bool) h (0::real))
+             ((op =::(real => real) => (real => real) => bool)
+               ((poly::real list => real => real) t)
+               ((poly::real list => real => real) ([]::real list))))))"
+  by (import poly POLY_ZERO_LEMMA)
+
+lemma POLY_ZERO: "ALL t. (poly t = poly []) = list_all (%c. c = 0) t"
+  by (import poly POLY_ZERO)
+
+lemma POLY_DIFF_AUX_ISZERO: "ALL t n.
+   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::(real list => bool) => bool)
+ (%x::real list.
+     (op -->::bool => bool => bool)
+      ((op =::(real => real) => (real => real) => bool)
+        ((poly::real list => real => real)
+          ((diff::real list => real list) x))
+        ((poly::real list => real => real) ([]::real list)))
+      ((Ex::(real => bool) => bool)
+        (%h::real.
+            (op =::(real => real) => (real => real) => bool)
+             ((poly::real list => real => real) x)
+             ((poly::real list => real => real)
+               ((op #::real => real list => real list) h
+                 ([]::real list))))))"
+  by (import poly POLY_DIFF_ISZERO)
+
+lemma POLY_DIFF_ZERO: "(All::(real list => bool) => bool)
+ (%x::real list.
+     (op -->::bool => bool => bool)
+      ((op =::(real => real) => (real => real) => bool)
+        ((poly::real list => real => real) x)
+        ((poly::real list => real => real) ([]::real list)))
+      ((op =::(real => real) => (real => real) => bool)
+        ((poly::real list => real => real)
+          ((diff::real list => real list) x))
+        ((poly::real list => real => real) ([]::real list))))"
+  by (import poly POLY_DIFF_ZERO)
+
+lemma POLY_DIFF_WELLDEF: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (op -->::bool => bool => bool)
+           ((op =::(real => real) => (real => real) => bool)
+             ((poly::real list => real => real) p)
+             ((poly::real list => real => real) q))
+           ((op =::(real => real) => (real => real) => bool)
+             ((poly::real list => real => real)
+               ((diff::real list => real list) p))
+             ((poly::real list => real => real)
+               ((diff::real list => real list) q)))))"
+  by (import poly POLY_DIFF_WELLDEF)
+
+constdefs
+  poly_divides :: "real list => real list => bool" 
+  "poly_divides == %p1 p2. EX q. poly p2 = poly (poly_mul p1 q)"
+
+lemma poly_divides: "ALL p1 p2. poly_divides p1 p2 = (EX q. poly p2 = poly (poly_mul p1 q))"
+  by (import poly poly_divides)
+
+lemma POLY_PRIMES: "ALL a p q.
+   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. poly_divides p p"
+  by (import poly POLY_DIVIDES_REFL)
+
+lemma POLY_DIVIDES_TRANS: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%r::real list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((poly_divides::real list => real list => bool) p q)
+                  ((poly_divides::real list => real list => bool) q r))
+                ((poly_divides::real list => real list => bool) p r))))"
+  by (import poly POLY_DIVIDES_TRANS)
+
+lemma POLY_DIVIDES_EXP: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(nat => bool) => bool)
+      (%m::nat.
+          (All::(nat => bool) => bool)
+           (%n::nat.
+               (op -->::bool => bool => bool)
+                ((op <=::nat => nat => bool) m n)
+                ((poly_divides::real list => real list => bool)
+                  ((poly_exp::real list => nat => real list) p m)
+                  ((poly_exp::real list => nat => real list) p n)))))"
+  by (import poly POLY_DIVIDES_EXP)
+
+lemma POLY_EXP_DIVIDES: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(nat => bool) => bool)
+           (%m::nat.
+               (All::(nat => bool) => bool)
+                (%n::nat.
+                    (op -->::bool => bool => bool)
+                     ((op &::bool => bool => bool)
+                       ((poly_divides::real list => real list => bool)
+                         ((poly_exp::real list => nat => real list) p n) q)
+                       ((op <=::nat => nat => bool) m n))
+                     ((poly_divides::real list => real list => bool)
+                       ((poly_exp::real list => nat => real list) p m)
+                       q)))))"
+  by (import poly POLY_EXP_DIVIDES)
+
+lemma POLY_DIVIDES_ADD: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%r::real list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((poly_divides::real list => real list => bool) p q)
+                  ((poly_divides::real list => real list => bool) p r))
+                ((poly_divides::real list => real list => bool) p
+                  ((poly_add::real list => real list => real list) q r)))))"
+  by (import poly POLY_DIVIDES_ADD)
+
+lemma POLY_DIVIDES_SUB: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%r::real list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((poly_divides::real list => real list => bool) p q)
+                  ((poly_divides::real list => real list => bool) p
+                    ((poly_add::real list => real list => real list) q r)))
+                ((poly_divides::real list => real list => bool) p r))))"
+  by (import poly POLY_DIVIDES_SUB)
+
+lemma POLY_DIVIDES_SUB2: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%r::real list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((poly_divides::real list => real list => bool) p r)
+                  ((poly_divides::real list => real list => bool) p
+                    ((poly_add::real list => real list => real list) q r)))
+                ((poly_divides::real list => real list => bool) p q))))"
+  by (import poly POLY_DIVIDES_SUB2)
+
+lemma POLY_DIVIDES_ZERO: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (op -->::bool => bool => bool)
+           ((op =::(real => real) => (real => real) => bool)
+             ((poly::real list => real => real) p)
+             ((poly::real list => real => real) ([]::real list)))
+           ((poly_divides::real list => real list => bool) q p)))"
+  by (import poly POLY_DIVIDES_ZERO)
+
+lemma POLY_ORDER_EXISTS: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(nat => bool) => bool)
+      (%d::nat.
+          (All::(real list => bool) => bool)
+           (%p::real list.
+               (op -->::bool => bool => bool)
+                ((op &::bool => bool => bool)
+                  ((op =::nat => nat => bool) ((size::real list => nat) p)
+                    d)
+                  ((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)
+                  (%x::nat.
+                      (op &::bool => bool => bool)
+                       ((poly_divides::real list => real list => bool)
+                         ((poly_exp::real list => nat => real list)
+                           ((op #::real => real list => real list)
+                             ((uminus::real => real) a)
+                             ((op #::real => real list => real list)
+                               (1::real) ([]::real list)))
+                           x)
+                         p)
+                       ((Not::bool => bool)
+                         ((poly_divides::real list => real list => bool)
+                           ((poly_exp::real list => nat => real list)
+                             ((op #::real => real list => real list)
+                               ((uminus::real => real) a)
+                               ((op #::real => real list => real list)
+                                 (1::real) ([]::real list)))
+                             ((Suc::nat => nat) x))
+                           p)))))))"
+  by (import poly POLY_ORDER_EXISTS)
+
+lemma POLY_ORDER: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (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))))
+           ((Ex1::(nat => bool) => bool)
+             (%n::nat.
+                 (op &::bool => bool => bool)
+                  ((poly_divides::real list => real list => bool)
+                    ((poly_exp::real list => nat => real list)
+                      ((op #::real => real list => real list)
+                        ((uminus::real => real) a)
+                        ((op #::real => real list => real list) (1::real)
+                          ([]::real list)))
+                      n)
+                    p)
+                  ((Not::bool => bool)
+                    ((poly_divides::real list => real list => bool)
+                      ((poly_exp::real list => nat => real list)
+                        ((op #::real => real list => real list)
+                          ((uminus::real => real) a)
+                          ((op #::real => real list => real list) (1::real)
+                            ([]::real list)))
+                        ((Suc::nat => nat) n))
+                      p))))))"
+  by (import poly POLY_ORDER)
+
+constdefs
+  poly_order :: "real => real list => nat" 
+  "poly_order ==
+%a p. SOME n.
+         poly_divides (poly_exp [- a, 1] n) p &
+         ~ poly_divides (poly_exp [- a, 1] (Suc n)) p"
+
+lemma poly_order: "ALL a p.
+   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 a n.
+   (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::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (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 &::bool => bool => bool)
+             ((poly_divides::real list => real list => bool)
+               ((poly_exp::real list => nat => real list)
+                 ((op #::real => real list => real list)
+                   ((uminus::real => real) a)
+                   ((op #::real => real list => real list) (1::real)
+                     ([]::real list)))
+                 ((poly_order::real => real list => nat) a p))
+               p)
+             ((Not::bool => bool)
+               ((poly_divides::real list => real list => bool)
+                 ((poly_exp::real list => nat => real list)
+                   ((op #::real => real list => real list)
+                     ((uminus::real => real) a)
+                     ((op #::real => real list => real list) (1::real)
+                       ([]::real list)))
+                   ((Suc::nat => nat)
+                     ((poly_order::real => real list => nat) a p)))
+                 p)))))"
+  by (import poly ORDER_THM)
+
+lemma ORDER_UNIQUE: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (All::(nat => bool) => bool)
+           (%n::nat.
+               (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 &::bool => bool => bool)
+                    ((poly_divides::real list => real list => bool)
+                      ((poly_exp::real list => nat => real list)
+                        ((op #::real => real list => real list)
+                          ((uminus::real => real) a)
+                          ((op #::real => real list => real list) (1::real)
+                            ([]::real list)))
+                        n)
+                      p)
+                    ((Not::bool => bool)
+                      ((poly_divides::real list => real list => bool)
+                        ((poly_exp::real list => nat => real list)
+                          ((op #::real => real list => real list)
+                            ((uminus::real => real) a)
+                            ((op #::real => real list => real list)
+                              (1::real) ([]::real list)))
+                          ((Suc::nat => nat) n))
+                        p))))
+                ((op =::nat => nat => bool) n
+                  ((poly_order::real => real list => nat) a p)))))"
+  by (import poly ORDER_UNIQUE)
+
+lemma ORDER_POLY: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real => bool) => bool)
+           (%a::real.
+               (op -->::bool => bool => bool)
+                ((op =::(real => real) => (real => real) => bool)
+                  ((poly::real list => real => real) p)
+                  ((poly::real list => real => real) q))
+                ((op =::nat => nat => bool)
+                  ((poly_order::real => real list => nat) a p)
+                  ((poly_order::real => real list => nat) a q)))))"
+  by (import poly ORDER_POLY)
+
+lemma ORDER_ROOT: "ALL p a. (poly p a = 0) = (poly p = poly [] | poly_order a p ~= 0)"
+  by (import poly ORDER_ROOT)
+
+lemma ORDER_DIVIDES: "ALL p a n.
+   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::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (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::(real list => bool) => bool)
+             (%x::real list.
+                 (op &::bool => bool => bool)
+                  ((op =::(real => real) => (real => real) => bool)
+                    ((poly::real list => real => real) p)
+                    ((poly::real list => real => real)
+                      ((poly_mul::real list => real list => real list)
+                        ((poly_exp::real list => nat => real list)
+                          ((op #::real => real list => real list)
+                            ((uminus::real => real) a)
+                            ((op #::real => real list => real list)
+                              (1::real) ([]::real list)))
+                          ((poly_order::real => real list => nat) a p))
+                        x)))
+                  ((Not::bool => bool)
+                    ((poly_divides::real list => real list => bool)
+                      ((op #::real => real list => real list)
+                        ((uminus::real => real) a)
+                        ((op #::real => real list => real list) (1::real)
+                          ([]::real list)))
+                      x))))))"
+  by (import poly ORDER_DECOMP)
+
+lemma ORDER_MUL: "(All::(real => bool) => bool)
+ (%a::real.
+     (All::(real list => bool) => bool)
+      (%p::real list.
+          (All::(real list => bool) => bool)
+           (%q::real list.
+               (op -->::bool => bool => bool)
+                ((Not::bool => bool)
+                  ((op =::(real => real) => (real => real) => bool)
+                    ((poly::real list => real => real)
+                      ((poly_mul::real list => real list => real list) p q))
+                    ((poly::real list => real => real) ([]::real list))))
+                ((op =::nat => nat => bool)
+                  ((poly_order::real => real list => nat) a
+                    ((poly_mul::real list => real list => real list) p q))
+                  ((op +::nat => nat => nat)
+                    ((poly_order::real => real list => nat) a p)
+                    ((poly_order::real => real list => nat) a q))))))"
+  by (import poly ORDER_MUL)
+
+lemma ORDER_DIFF: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((Not::bool => bool)
+               ((op =::(real => real) => (real => real) => bool)
+                 ((poly::real list => real => real)
+                   ((diff::real list => real list) p))
+                 ((poly::real list => real => real) ([]::real list))))
+             ((Not::bool => bool)
+               ((op =::nat => nat => bool)
+                 ((poly_order::real => real list => nat) a p) (0::nat))))
+           ((op =::nat => nat => bool)
+             ((poly_order::real => real list => nat) a p)
+             ((Suc::nat => nat)
+               ((poly_order::real => real list => nat) a
+                 ((diff::real list => real list) p))))))"
+  by (import poly ORDER_DIFF)
+
+lemma POLY_SQUAREFREE_DECOMP_ORDER: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%d::real list.
+               (All::(real list => bool) => bool)
+                (%e::real list.
+                    (All::(real list => bool) => bool)
+                     (%r::real list.
+                         (All::(real list => bool) => bool)
+                          (%s::real list.
+                              (op -->::bool => bool => bool)
+                               ((op &::bool => bool => bool)
+                                 ((Not::bool => bool)
+                                   ((op =::(real => real)
+     => (real => real) => bool)
+                                     ((poly::real list => real => real)
+ ((diff::real list => real list) p))
+                                     ((poly::real list => real => real)
+ ([]::real list))))
+                                 ((op &::bool => bool => bool)
+                                   ((op =::(real => real)
+     => (real => real) => bool)
+                                     ((poly::real list => real => real) p)
+                                     ((poly::real list => real => real)
+ ((poly_mul::real list => real list => real list) q d)))
+                                   ((op &::bool => bool => bool)
+                                     ((op =::(real => real)
+       => (real => real) => bool)
+ ((poly::real list => real => real) ((diff::real list => real list) p))
+ ((poly::real list => real => real)
+   ((poly_mul::real list => real list => real list) e d)))
+                                     ((op =::(real => real)
+       => (real => real) => bool)
+ ((poly::real list => real => real) d)
+ ((poly::real list => real => real)
+   ((poly_add::real list => real list => real list)
+     ((poly_mul::real list => real list => real list) r p)
+     ((poly_mul::real list => real list => real list) s
+       ((diff::real list => real list) p))))))))
+                               ((All::(real => bool) => bool)
+                                 (%a::real.
+                                     (op =::nat => nat => bool)
+((poly_order::real => real list => nat) a q)
+((If::bool => nat => nat => nat)
+  ((op =::nat => nat => bool) ((poly_order::real => real list => nat) a p)
+    (0::nat))
+  (0::nat) (1::nat))))))))))"
+  by (import poly POLY_SQUAREFREE_DECOMP_ORDER)
+
+constdefs
+  rsquarefree :: "real list => bool" 
+  "rsquarefree ==
+%p. poly p ~= poly [] & (ALL a. poly_order a p = 0 | poly_order a p = 1)"
+
+lemma rsquarefree: "ALL p.
+   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. rsquarefree p = (ALL a. ~ (poly p a = 0 & poly (diff p) a = 0))"
+  by (import poly RSQUAREFREE_ROOTS)
+
+lemma RSQUAREFREE_DECOMP: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real => bool) => bool)
+      (%a::real.
+          (op -->::bool => bool => bool)
+           ((op &::bool => bool => bool)
+             ((rsquarefree::real list => bool) p)
+             ((op =::real => real => bool)
+               ((poly::real list => real => real) p a) (0::real)))
+           ((Ex::(real list => bool) => bool)
+             (%q::real list.
+                 (op &::bool => bool => bool)
+                  ((op =::(real => real) => (real => real) => bool)
+                    ((poly::real list => real => real) p)
+                    ((poly::real list => real => real)
+                      ((poly_mul::real list => real list => real list)
+                        ((op #::real => real list => real list)
+                          ((uminus::real => real) a)
+                          ((op #::real => real list => real list) (1::real)
+                            ([]::real list)))
+                        q)))
+                  ((Not::bool => bool)
+                    ((op =::real => real => bool)
+                      ((poly::real list => real => real) q a)
+                      (0::real)))))))"
+  by (import poly RSQUAREFREE_DECOMP)
+
+lemma POLY_SQUAREFREE_DECOMP: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (All::(real list => bool) => bool)
+      (%q::real list.
+          (All::(real list => bool) => bool)
+           (%d::real list.
+               (All::(real list => bool) => bool)
+                (%e::real list.
+                    (All::(real list => bool) => bool)
+                     (%r::real list.
+                         (All::(real list => bool) => bool)
+                          (%s::real list.
+                              (op -->::bool => bool => bool)
+                               ((op &::bool => bool => bool)
+                                 ((Not::bool => bool)
+                                   ((op =::(real => real)
+     => (real => real) => bool)
+                                     ((poly::real list => real => real)
+ ((diff::real list => real list) p))
+                                     ((poly::real list => real => real)
+ ([]::real list))))
+                                 ((op &::bool => bool => bool)
+                                   ((op =::(real => real)
+     => (real => real) => bool)
+                                     ((poly::real list => real => real) p)
+                                     ((poly::real list => real => real)
+ ((poly_mul::real list => real list => real list) q d)))
+                                   ((op &::bool => bool => bool)
+                                     ((op =::(real => real)
+       => (real => real) => bool)
+ ((poly::real list => real => real) ((diff::real list => real list) p))
+ ((poly::real list => real => real)
+   ((poly_mul::real list => real list => real list) e d)))
+                                     ((op =::(real => real)
+       => (real => real) => bool)
+ ((poly::real list => real => real) d)
+ ((poly::real list => real => real)
+   ((poly_add::real list => real list => real list)
+     ((poly_mul::real list => real list => real list) r p)
+     ((poly_mul::real list => real list => real list) s
+       ((diff::real list => real list) p))))))))
+                               ((op &::bool => bool => bool)
+                                 ((rsquarefree::real list => bool) q)
+                                 ((All::(real => bool) => bool)
+                                   (%x::real.
+ (op =::bool => bool => bool)
+  ((op =::real => real => bool) ((poly::real list => real => real) q x)
+    (0::real))
+  ((op =::real => real => bool) ((poly::real list => real => real) p x)
+    (0::real)))))))))))"
+  by (import poly POLY_SQUAREFREE_DECOMP)
+
+consts
+  normalize :: "real list => real list" 
+
+specification (normalize) normalize: "normalize [] = [] &
+(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. poly (normalize t) = poly t"
+  by (import poly POLY_NORMALIZE)
+
+constdefs
+  degree :: "real list => nat" 
+  "degree == %p. PRE (length (normalize p))"
+
+lemma degree: "ALL p. degree p = PRE (length (normalize p))"
+  by (import poly degree)
+
+lemma DEGREE_ZERO: "(All::(real list => bool) => bool)
+ (%p::real list.
+     (op -->::bool => 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) ((degree::real list => nat) p) (0::nat)))"
+  by (import poly DEGREE_ZERO)
+
+lemma POLY_ROOTS_FINITE_SET: "(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))))
+      ((FINITE::(real => bool) => bool)
+        ((GSPEC::(real => real * bool) => real => bool)
+          (%x::real.
+              (Pair::real => bool => real * bool) x
+               ((op =::real => real => bool)
+                 ((poly::real list => real => real) p x) (0::real))))))"
+  by (import poly POLY_ROOTS_FINITE_SET)
+
+lemma POLY_MONO: "(All::(real => bool) => bool)
+ (%x::real.
+     (All::(real => bool) => bool)
+      (%k::real.
+          (All::(real list => bool) => bool)
+           (%xa::real list.
+               (op -->::bool => bool => bool)
+                ((op <=::real => real => bool) ((abs::real => real) x) k)
+                ((op <=::real => real => bool)
+                  ((abs::real => real)
+                    ((poly::real list => real => real) xa x))
+                  ((poly::real list => real => real)
+                    ((map::(real => real) => real list => real list)
+                      (abs::real => real) xa)
+                    k)))))"
+  by (import poly POLY_MONO)
+
+;end_setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4Vec.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,1212 @@
+theory HOL4Vec = HOL4Base:
+
+;setup_theory res_quan
+
+lemma RES_FORALL_CONJ_DIST: "ALL P Q R. 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 Q R. 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 xa. RES_FORALL (op = xa) x = x xa"
+  by (import res_quan RES_FORALL_UNIQUE)
+
+lemma RES_FORALL_FORALL: "ALL (P::'a => bool) (R::'a => 'b => bool) x::'b.
+   (ALL x::'b. RES_FORALL P (%i::'a. R i x)) =
+   RES_FORALL P (%i::'a. All (R i))"
+  by (import res_quan RES_FORALL_FORALL)
+
+lemma RES_FORALL_REORDER: "ALL P Q R.
+   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)"
+  by (import res_quan RES_FORALL_EMPTY)
+
+lemma RES_FORALL_UNIV: "ALL p. RES_FORALL pred_set.UNIV p = All p"
+  by (import res_quan RES_FORALL_UNIV)
+
+lemma RES_FORALL_NULL: "ALL p m. RES_FORALL p (%x. m) = (p = EMPTY | m)"
+  by (import res_quan RES_FORALL_NULL)
+
+lemma RES_EXISTS_DISJ_DIST: "(All::(('a => bool) => bool) => bool)
+ (%P::'a => bool.
+     (All::(('a => bool) => bool) => bool)
+      (%Q::'a => bool.
+          (All::(('a => bool) => bool) => bool)
+           (%R::'a => bool.
+               (op =::bool => bool => bool)
+                ((RES_EXISTS::('a => bool) => ('a => bool) => bool) P
+                  (%i::'a. (op |::bool => bool => bool) (Q i) (R i)))
+                ((op |::bool => bool => bool)
+                  ((RES_EXISTS::('a => bool) => ('a => bool) => bool) P Q)
+                  ((RES_EXISTS::('a => bool) => ('a => bool) => bool) P
+                    R)))))"
+  by (import res_quan RES_EXISTS_DISJ_DIST)
+
+lemma RES_DISJ_EXISTS_DIST: "ALL P Q R. 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 xa. RES_EXISTS (op = xa) x = x xa"
+  by (import res_quan RES_EXISTS_EQUAL)
+
+lemma RES_EXISTS_REORDER: "ALL P Q R.
+   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. ~ RES_EXISTS EMPTY p"
+  by (import res_quan RES_EXISTS_EMPTY)
+
+lemma RES_EXISTS_UNIV: "ALL p. RES_EXISTS pred_set.UNIV p = Ex p"
+  by (import res_quan RES_EXISTS_UNIV)
+
+lemma RES_EXISTS_NULL: "ALL p m. RES_EXISTS p (%x. m) = (p ~= EMPTY & m)"
+  by (import res_quan RES_EXISTS_NULL)
+
+lemma RES_EXISTS_ALT: "ALL p m. 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. ~ RES_EXISTS_UNIQUE EMPTY p"
+  by (import res_quan RES_EXISTS_UNIQUE_EMPTY)
+
+lemma RES_EXISTS_UNIQUE_UNIV: "ALL p. RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p"
+  by (import res_quan RES_EXISTS_UNIQUE_UNIV)
+
+lemma RES_EXISTS_UNIQUE_NULL: "ALL p m. 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 m.
+   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. RES_SELECT EMPTY p = (SOME x. False)"
+  by (import res_quan RES_SELECT_EMPTY)
+
+lemma RES_SELECT_UNIV: "ALL p. RES_SELECT pred_set.UNIV p = Eps p"
+  by (import res_quan RES_SELECT_UNIV)
+
+lemma RES_ABSTRACT: "ALL p m x. IN x p --> RES_ABSTRACT p m x = m x"
+  by (import res_quan RES_ABSTRACT)
+
+lemma RES_ABSTRACT_EQUAL: "ALL p m1 m2.
+   (ALL 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 m. 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 m1 m2.
+   (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 list recspace => bool) => 'a list recspace set)
+ (%x::'a list recspace.
+     (All::(('a list recspace => bool) => bool) => bool)
+      (%word::'a list recspace => bool.
+          (op -->::bool => bool => bool)
+           ((All::('a list recspace => bool) => bool)
+             (%a0::'a list recspace.
+                 (op -->::bool => bool => bool)
+                  ((Ex::('a list => bool) => bool)
+                    (%a::'a list.
+                        (op =::'a list recspace => 'a list recspace => bool)
+                         a0 ((CONSTR::nat
+=> 'a list => (nat => 'a list recspace) => 'a list recspace)
+                              (0::nat) a
+                              (%n::nat. BOTTOM::'a list recspace))))
+                  (word a0)))
+           (word x)))" 
+  by (rule typedef_helper,import word_base word_TY_DEF)
+
+lemmas word_TY_DEF = typedef_hol2hol4 [OF type_definition_word]
+
+consts
+  mk_word :: "'a list recspace => 'a word" 
+  dest_word :: "'a word => 'a list recspace" 
+
+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))"
+  by (import word_base word_repfns)
+
+consts
+  word_base0 :: "'a list => 'a word" 
+
+defs
+  word_base0_primdef: "word_base0 == %a. mk_word (CONSTR 0 a (%n. BOTTOM))"
+
+lemma word_base0_def: "word_base0 = (%a. mk_word (CONSTR 0 a (%n. BOTTOM)))"
+  by (import word_base word_base0_def)
+
+constdefs
+  WORD :: "'a list => 'a word" 
+  "WORD == word_base0"
+
+lemma WORD: "WORD = word_base0"
+  by (import word_base WORD)
+
+consts
+  word_case :: "('a list => 'b) => 'a word => 'b" 
+
+specification (word_case_primdef: word_case) word_case_def: "ALL f a. 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. word_size f (WORD a) = 1 + list_size f a"
+  by (import word_base word_size_def)
+
+lemma word_11: "ALL a a'. (WORD a = WORD a') = (a = a')"
+  by (import word_base word_11)
+
+lemma word_case_cong: "ALL M M' f.
+   M = M' & (ALL a. M' = WORD a --> f a = f' a) -->
+   word_case f M = word_case f' M'"
+  by (import word_base word_case_cong)
+
+lemma word_nchotomy: "ALL x. EX l. x = WORD l"
+  by (import word_base word_nchotomy)
+
+lemma word_Axiom: "ALL f. EX fn. ALL a. fn (WORD a) = f a"
+  by (import word_base word_Axiom)
+
+lemma word_induction: "ALL P. (ALL a. P (WORD a)) --> All P"
+  by (import word_base word_induction)
+
+lemma word_Ax: "ALL f. EX fn. ALL a. fn (WORD a) = f a"
+  by (import word_base word_Ax)
+
+lemma WORD_11: "ALL x xa. (WORD x = WORD xa) = (x = xa)"
+  by (import word_base WORD_11)
+
+lemma word_induct: "ALL x. (ALL l. x (WORD l)) --> All x"
+  by (import word_base word_induct)
+
+lemma word_cases: "ALL x. EX l. x = WORD l"
+  by (import word_base word_cases)
+
+consts
+  WORDLEN :: "'a word => nat" 
+
+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. GSPEC (%w. (w, WORDLEN w = n))"
+
+lemma PWORDLEN_def: "ALL n. PWORDLEN n = GSPEC (%w. (w, WORDLEN w = n))"
+  by (import word_base PWORDLEN_def)
+
+lemma IN_PWORDLEN: "ALL n l. IN (WORD l) (PWORDLEN n) = (length l = n)"
+  by (import word_base IN_PWORDLEN)
+
+lemma PWORDLEN: "ALL n w. IN w (PWORDLEN n) = (WORDLEN w = n)"
+  by (import word_base PWORDLEN)
+
+lemma PWORDLEN0: "ALL w. IN w (PWORDLEN 0) --> w = WORD []"
+  by (import word_base PWORDLEN0)
+
+lemma PWORDLEN1: "ALL x. IN (WORD [x]) (PWORDLEN 1)"
+  by (import word_base PWORDLEN1)
+
+consts
+  WSEG :: "nat => nat => 'a word => 'a word" 
+
+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 w. WSEG 0 k w = WORD []"
+  by (import word_base WSEG0)
+
+lemma WSEG_PWORDLEN: "ALL n.
+   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.
+   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. 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 l. bit k (WORD l) = ELL k l"
+  by (import word_base BIT_DEF)
+
+lemma BIT0: "ALL x. bit 0 (WORD [x]) = x"
+  by (import word_base BIT0)
+
+lemma WSEG_BIT: "ALL n. 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.
+   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. MSB (WORD l) = hd l"
+  by (import word_base MSB_DEF)
+
+lemma MSB: "ALL n. 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. LSB (WORD l) = last l"
+  by (import word_base LSB_DEF)
+
+lemma LSB: "ALL n. 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 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 l2. WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)"
+  by (import word_base WCAT_DEF)
+
+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 w2 w3. WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)"
+  by (import word_base WCAT_ASSOC)
+
+lemma WCAT0: "ALL w. WCAT (WORD [], w) = w & WCAT (w, WORD []) = w"
+  by (import word_base WCAT0)
+
+lemma WCAT_11: "ALL m n.
+   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 n.
+   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.
+   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 w.
+   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.
+   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 n.
+   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 n.
+   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 n. 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.
+   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 xa.
+   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)) (%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)) (%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.
+   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.
+   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 n2.
+   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 n2.
+   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.
+   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 l. WORD (x # l) = WCAT (WORD [x], WORD l)"
+  by (import word_base WORD_CONS_WCAT)
+
+lemma WORD_SNOC_WCAT: "ALL l x. WORD (SNOC x l) = WCAT (WORD l, WORD [x])"
+  by (import word_base WORD_SNOC_WCAT)
+
+lemma BIT_WCAT_FST: "ALL n1 n2.
+   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 n1 n2.
+   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. 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 n2.
+   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 n2.
+   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 n2.
+   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 n.
+   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
+
+constdefs
+  LVAL :: "('a => nat) => nat => 'a list => nat" 
+  "LVAL == %f b. foldl (%e x. b * e + f x) 0"
+
+lemma LVAL_DEF: "ALL f b l. 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 b l. NVAL f b (WORD l) = LVAL f b l"
+  by (import word_num NVAL_DEF)
+
+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 h f b. LVAL f b (SNOC h l) = LVAL f b l * b + f h"
+  by (import word_num LVAL_SNOC)
+
+lemma LVAL_MAX: "ALL l f b. (ALL x. f x < b) --> LVAL f b l < b ^ length l"
+  by (import word_num LVAL_MAX)
+
+lemma NVAL_MAX: "ALL f b.
+   (ALL x. f x < b) -->
+   (ALL n. RES_FORALL (PWORDLEN n) (%w. NVAL f b w < b ^ n))"
+  by (import word_num NVAL_MAX)
+
+lemma NVAL0: "ALL x xa. NVAL x xa (WORD []) = 0"
+  by (import word_num NVAL0)
+
+lemma NVAL1: "ALL x xa xb. NVAL x xa (WORD [xb]) = x xb"
+  by (import word_num NVAL1)
+
+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 f b x. 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.
+   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 m.
+   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) (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)
+
+constdefs
+  NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" 
+  "NWORD == %n frep b m. WORD (NLIST n frep b m)"
+
+lemma NWORD_DEF: "ALL n frep b m. NWORD n frep b m = WORD (NLIST n frep b m)"
+  by (import word_num NWORD_DEF)
+
+lemma NWORD_LENGTH: "ALL x xa xb xc. WORDLEN (NWORD x xa xb xc) = x"
+  by (import word_num NWORD_LENGTH)
+
+lemma NWORD_PWORDLEN: "ALL x xa xb xc. IN (NWORD x xa xb xc) (PWORDLEN x)"
+  by (import word_num NWORD_PWORDLEN)
+
+;end_setup
+
+;setup_theory word_bitop
+
+consts
+  PBITOP :: "('a word => 'b word) => bool" 
+
+defs
+  PBITOP_primdef: "PBITOP ==
+GSPEC
+ (%oper.
+     (oper,
+      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)))))"
+
+lemma PBITOP_def: "PBITOP =
+GSPEC
+ (%oper.
+     (oper,
+      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 PBITOP_def)
+
+lemma IN_PBITOP: "ALL oper.
+   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. 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.
+     ALL n.
+        RES_FORALL (PWORDLEN n)
+         (%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 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
+  PBITBOP :: "('a word => 'b word => 'c word) => bool" 
+
+defs
+  PBITBOP_primdef: "PBITBOP ==
+GSPEC
+ (%oper.
+     (oper,
+      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))))))"
+
+lemma PBITBOP_def: "PBITBOP =
+GSPEC
+ (%oper.
+     (oper,
+      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 PBITBOP_def)
+
+lemma IN_PBITBOP: "ALL oper.
+   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.
+     ALL n.
+        RES_FORALL (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.
+     ALL n.
+        RES_FORALL (PWORDLEN n)
+         (%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. 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 l. WMAP f (WORD l) = WORD (map f l)"
+  by (import word_bitop WMAP_DEF)
+
+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. WMAP x (WORD []) = WORD []"
+  by (import word_bitop WMAP_0)
+
+lemma WMAP_BIT: "ALL n.
+   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.
+   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. IN (WMAP f) PBITOP"
+  by (import word_bitop WMAP_PBITOP)
+
+lemma WMAP_WCAT: "ALL w1 w2 f. WMAP f (WCAT (w1, w2)) = WCAT (WMAP f w1, WMAP f w2)"
+  by (import word_bitop WMAP_WCAT)
+
+lemma WMAP_o: "ALL w f g. WMAP g (WMAP f w) = WMAP (g o f) w"
+  by (import word_bitop WMAP_o)
+
+consts
+  FORALLBITS :: "('a => bool) => 'a word => bool" 
+
+specification (FORALLBITS) FORALLBITS_DEF: "ALL P l. FORALLBITS P (WORD l) = list_all P l"
+  by (import word_bitop FORALLBITS_DEF)
+
+lemma FORALLBITS: "ALL n.
+   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.
+   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 w2 P.
+   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 l. EXISTSABIT P (WORD l) = list_exists P l"
+  by (import word_bitop EXISTSABIT_DEF)
+
+lemma NOT_EXISTSABIT: "ALL P w. (~ EXISTSABIT P w) = FORALLBITS (Not o P) w"
+  by (import word_bitop NOT_EXISTSABIT)
+
+lemma NOT_FORALLBITS: "ALL P w. (~ FORALLBITS P w) = EXISTSABIT (Not o P) w"
+  by (import word_bitop NOT_FORALLBITS)
+
+lemma EXISTSABIT: "ALL n.
+   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.
+   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 w2 P.
+   EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)"
+  by (import word_bitop EXISTSABIT_WCAT)
+
+constdefs
+  SHR :: "bool => 'a => 'a word => 'a word * 'a" 
+  "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: "ALL f b w.
+   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)
+
+constdefs
+  SHL :: "bool => 'a word => 'a => 'a * 'a word" 
+  "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: "ALL f w b.
+   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.
+   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.
+   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.
+   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.
+   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.
+   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.
+   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.
+   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.
+   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
+
+constdefs
+  BV :: "bool => nat" 
+  "BV == %b. if b then Suc 0 else 0"
+
+lemma BV_DEF: "ALL b. 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. BNVAL (WORD l) = LVAL BV 2 l"
+  by (import bword_num BNVAL_DEF)
+
+lemma BV_LESS_2: "ALL x. BV x < 2"
+  by (import bword_num BV_LESS_2)
+
+lemma BNVAL_NVAL: "ALL w. 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 w2. WORDLEN w1 = WORDLEN w2 --> BNVAL w1 = BNVAL w2 --> w1 = w2"
+  by (import bword_num BNVAL_11)
+
+lemma BNVAL_ONTO: "ALL w. Ex (op = (BNVAL w))"
+  by (import bword_num BNVAL_ONTO)
+
+lemma BNVAL_MAX: "ALL n. RES_FORALL (PWORDLEN n) (%w. BNVAL w < 2 ^ n)"
+  by (import bword_num BNVAL_MAX)
+
+lemma BNVAL_WCAT1: "ALL n.
+   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.
+   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 m.
+   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)
+
+constdefs
+  VB :: "nat => bool" 
+  "VB == %n. n mod 2 ~= 0"
+
+lemma VB_DEF: "ALL n. VB n = (n mod 2 ~= 0)"
+  by (import bword_num VB_DEF)
+
+constdefs
+  NBWORD :: "nat => nat => bool word" 
+  "NBWORD == %n m. WORD (NLIST n VB 2 m)"
+
+lemma NBWORD_DEF: "ALL n m. NBWORD n m = WORD (NLIST n VB 2 m)"
+  by (import bword_num NBWORD_DEF)
+
+lemma NBWORD0: "ALL x. NBWORD 0 x = WORD []"
+  by (import bword_num NBWORD0)
+
+lemma WORDLEN_NBWORD: "ALL x xa. WORDLEN (NBWORD x xa) = x"
+  by (import bword_num WORDLEN_NBWORD)
+
+lemma PWORDLEN_NBWORD: "ALL x xa. IN (NBWORD x xa) (PWORDLEN x)"
+  by (import bword_num PWORDLEN_NBWORD)
+
+lemma NBWORD_SUC: "ALL n m. 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. VB (BV x) = x"
+  by (import bword_num VB_BV)
+
+lemma BV_VB: "ALL x<2. BV (VB x) = x"
+  by (import bword_num BV_VB)
+
+lemma NBWORD_BNVAL: "ALL n. RES_FORALL (PWORDLEN n) (%w. NBWORD n (BNVAL w) = w)"
+  by (import bword_num NBWORD_BNVAL)
+
+lemma BNVAL_NBWORD: "ALL n m. m < 2 ^ n --> BNVAL (NBWORD n m) = m"
+  by (import bword_num BNVAL_NBWORD)
+
+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 n2. WCAT (NBWORD n1 0, NBWORD n2 0) = NBWORD (n1 + n2) 0"
+  by (import bword_num WCAT_NBWORD_0)
+
+lemma WSPLIT_NBWORD_0: "ALL n m. 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 n.
+   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 m. NBWORD n (m mod 2 ^ n) = NBWORD n m"
+  by (import bword_num NBWORD_MOD)
+
+lemma WSEG_NBWORD_SUC: "ALL n m. WSEG n 0 (NBWORD (Suc n) m) = NBWORD n m"
+  by (import bword_num WSEG_NBWORD_SUC)
+
+lemma NBWORD_SUC_WSEG: "ALL n. 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 < 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 m. bit n (NBWORD (Suc n) m) = VB (m div 2 ^ n mod 2)"
+  by (import bword_num MSB_NBWORD)
+
+lemma NBWORD_SPLIT: "ALL n1 n2 m.
+   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 k n.
+   m + k <= n --> (ALL l. WSEG m k (NBWORD n l) = NBWORD m (l div 2 ^ k))"
+  by (import bword_num WSEG_NBWORD)
+
+lemma NBWORD_SUC_FST: "ALL n x. 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 n. k < n --> bit k (NBWORD n 0) = False"
+  by (import bword_num BIT_NBWORD0)
+
+lemma ADD_BNVAL_LEFT: "ALL n.
+   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.
+   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 n2.
+   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
+
+;setup_theory bword_arith
+
+consts
+  ACARRY :: "nat => bool word => bool word => bool => 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)
+
+consts
+  ICARRY :: "nat => bool word => bool word => bool => 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.
+   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. RES_FORALL (PWORDLEN n) (%w. BNVAL w <= 2 ^ n - 1)"
+  by (import bword_arith BNVAL_LESS_EQ)
+
+lemma ADD_BNVAL_LESS_EQ1: "ALL n cin.
+   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 x1 x2 cin.
+   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 x1 x2 cin.
+   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 x1 x2 cin.
+   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 n.
+   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 n2.
+   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.
+   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 n2.
+   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.
+   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.
+   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.
+   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.
+   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
+
+;setup_theory bword_bitop
+
+consts
+  WNOT :: "bool word => bool word" 
+
+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. WNOT (WNOT w) = w"
+  by (import bword_bitop WNOT_WNOT)
+
+lemma WCAT_WNOT: "ALL n1 n2.
+   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 l2. WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)"
+  by (import bword_bitop WAND_DEF)
+
+lemma PBITBOP_WAND: "IN WAND PBITBOP"
+  by (import bword_bitop PBITBOP_WAND)
+
+consts
+  WOR :: "bool word => bool word => bool word" 
+
+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"
+  by (import bword_bitop PBITBOP_WOR)
+
+consts
+  WXOR :: "bool word => bool word => bool word" 
+
+specification (WXOR) WXOR_DEF: "ALL l1 l2. WXOR (WORD l1) (WORD l2) = WORD (map2 (%x y. x ~= y) l1 l2)"
+  by (import bword_bitop WXOR_DEF)
+
+lemma PBITBOP_WXOR: "IN WXOR PBITBOP"
+  by (import bword_bitop PBITBOP_WXOR)
+
+;end_setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/HOL4Word32.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,1451 @@
+theory HOL4Word32 = HOL4Base:
+
+;setup_theory bits
+
+consts
+  DIV2 :: "nat => nat" 
+
+defs
+  DIV2_primdef: "DIV2 == %n. n div 2"
+
+lemma DIV2_def: "ALL n. DIV2 n = n div 2"
+  by (import bits DIV2_def)
+
+consts
+  TIMES_2EXP :: "nat => nat => nat" 
+
+defs
+  TIMES_2EXP_primdef: "TIMES_2EXP == %x n. n * 2 ^ x"
+
+lemma TIMES_2EXP_def: "ALL x n. 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 n. n div 2 ^ x"
+
+lemma DIV_2EXP_def: "ALL x n. 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 n. n mod 2 ^ x"
+
+lemma MOD_2EXP_def: "ALL x n. 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 n. (n div 2 ^ x, n mod 2 ^ x)"
+
+lemma DIVMOD_2EXP_def: "ALL x n. 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 n. if b then 2 ^ n else 0"
+
+lemma SBIT_def: "ALL b n. 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 l n. MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
+
+lemma BITS_def: "ALL h l n. BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)"
+  by (import bits BITS_def)
+
+constdefs
+  bit :: "nat => nat => bool" 
+  "bit == %b n. BITS b b n = 1"
+
+lemma BIT_def: "ALL b n. bit b n = (BITS b b n = 1)"
+  by (import bits BIT_def)
+
+consts
+  SLICE :: "nat => nat => nat => nat" 
+
+defs
+  SLICE_primdef: "SLICE == %h l n. MOD_2EXP (Suc h) n - MOD_2EXP l n"
+
+lemma SLICE_def: "ALL h l n. SLICE h l n = MOD_2EXP (Suc h) n - MOD_2EXP l n"
+  by (import bits SLICE_def)
+
+consts
+  LSBn :: "nat => bool" 
+
+defs
+  LSBn_primdef: "LSBn == bit 0"
+
+lemma LSBn_def: "LSBn = bit 0"
+  by (import bits LSBn_def)
+
+consts
+  BITWISE :: "nat => (bool => bool => bool) => nat => nat => 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::nat) = x"
+  by (import bits DIV1)
+
+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::nat)"
+  by (import bits DIV_MULT_1)
+
+lemma ZERO_LT_TWOEXP: "ALL n::nat. (0::nat) < (2::nat) ^ n"
+  by (import bits ZERO_LT_TWOEXP)
+
+lemma MOD_2EXP_LT: "ALL (n::nat) k::nat. k mod (2::nat) ^ n < (2::nat) ^ n"
+  by (import bits MOD_2EXP_LT)
+
+lemma TWOEXP_DIVISION: "ALL (n::nat) k::nat.
+   k = k div (2::nat) ^ n * (2::nat) ^ n + k mod (2::nat) ^ n"
+  by (import bits TWOEXP_DIVISION)
+
+lemma TWOEXP_MONO: "ALL (a::nat) b::nat. a < b --> (2::nat) ^ a < (2::nat) ^ b"
+  by (import bits TWOEXP_MONO)
+
+lemma TWOEXP_MONO2: "ALL (a::nat) b::nat. a <= b --> (2::nat) ^ a <= (2::nat) ^ b"
+  by (import bits TWOEXP_MONO2)
+
+lemma EXP_SUB_LESS_EQ: "ALL (a::nat) b::nat. (2::nat) ^ (a - b) <= (2::nat) ^ a"
+  by (import bits EXP_SUB_LESS_EQ)
+
+lemma BITS_THM: "ALL x xa xb. BITS x xa xb = xb div 2 ^ xa mod 2 ^ (Suc x - xa)"
+  by (import bits BITS_THM)
+
+lemma BITSLT_THM: "ALL h l n. BITS h l n < 2 ^ (Suc h - l)"
+  by (import bits BITSLT_THM)
+
+lemma DIV_MULT_LEM: "ALL (m::nat) n::nat. (0::nat) < n --> m div n * n <= m"
+  by (import bits DIV_MULT_LEM)
+
+lemma MOD_2EXP_LEM: "ALL (n::nat) x::nat.
+   n mod (2::nat) ^ x = n - n div (2::nat) ^ x * (2::nat) ^ x"
+  by (import bits MOD_2EXP_LEM)
+
+lemma BITS2_THM: "ALL h l n. BITS h l n = n mod 2 ^ Suc h div 2 ^ l"
+  by (import bits BITS2_THM)
+
+lemma BITS_COMP_THM: "ALL h1 l1 h2 l2 n.
+   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 l x n. BITS h l x div 2 ^ n = BITS h (l + n) x"
+  by (import bits BITS_DIV_THM)
+
+lemma BITS_LT_HIGH: "ALL h l n. n < 2 ^ Suc h --> BITS h l n = n div 2 ^ l"
+  by (import bits BITS_LT_HIGH)
+
+lemma BITS_ZERO: "ALL h l n. h < l --> BITS h l n = 0"
+  by (import bits BITS_ZERO)
+
+lemma BITS_ZERO2: "ALL h l. BITS h l 0 = 0"
+  by (import bits BITS_ZERO2)
+
+lemma BITS_ZERO3: "ALL h x. BITS h 0 x = x mod 2 ^ Suc h"
+  by (import bits BITS_ZERO3)
+
+lemma BITS_COMP_THM2: "ALL h1 l1 h2 l2 n.
+   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::nat) ~= (0::nat)) = (n mod (2::nat) = (1::nat))"
+  by (import bits NOT_MOD2_LEM)
+
+lemma NOT_MOD2_LEM2: "ALL (n::nat) a::'a.
+   (n mod (2::nat) ~= (1::nat)) = (n mod (2::nat) = (0::nat))"
+  by (import bits NOT_MOD2_LEM2)
+
+lemma EVEN_MOD2_LEM: "ALL n. EVEN n = (n mod 2 = 0)"
+  by (import bits EVEN_MOD2_LEM)
+
+lemma ODD_MOD2_LEM: "ALL n. 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::nat) ^ x * (2::nat) ^ x = n - n mod (2::nat) ^ x"
+  by (import bits DIV_MULT_THM)
+
+lemma DIV_MULT_THM2: "ALL x::nat. (2::nat) * (x 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::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::nat) ^ (x + y) * (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) (x::nat) y::nat.
+   (n::nat) mod (2::nat) ^ (x + y) =
+   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::nat) ^ Suc l <= n mod (2::nat) ^ h"
+  by (import bits SLICE_LEM3)
+
+lemma SLICE_THM: "ALL n h l. SLICE h l n = BITS h l n * 2 ^ l"
+  by (import bits SLICE_THM)
+
+lemma SLICELT_THM: "ALL h l n. SLICE h l n < 2 ^ Suc h"
+  by (import bits SLICELT_THM)
+
+lemma BITS_SLICE_THM: "ALL h l n. BITS h l (SLICE h l n) = BITS h l n"
+  by (import bits BITS_SLICE_THM)
+
+lemma BITS_SLICE_THM2: "ALL h l n. 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::nat) ^ l <= n mod (2::nat) ^ Suc h"
+  by (import bits MOD_2EXP_MONO)
+
+lemma SLICE_COMP_THM: "ALL h m l n.
+   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 l n. h < l --> SLICE h l n = 0"
+  by (import bits SLICE_ZERO)
+
+lemma BIT_COMP_THM3: "ALL h m l n.
+   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 a. (~ bit n a) = (BITS n n a = 0)"
+  by (import bits NOT_BIT)
+
+lemma NOT_BITS: "ALL n a. (BITS n n a ~= 0) = (BITS n n a = 1)"
+  by (import bits NOT_BITS)
+
+lemma NOT_BITS2: "ALL n a. (BITS n n a ~= 1) = (BITS n n a = 0)"
+  by (import bits NOT_BITS2)
+
+lemma BIT_SLICE: "ALL n a b. (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 x n. SBIT (bit x n) (x + y) = SLICE x x n * 2 ^ y"
+  by (import bits BIT_SLICE_LEM)
+
+lemma BIT_SLICE_THM: "ALL x xa. SBIT (bit x xa) x = SLICE x x xa"
+  by (import bits BIT_SLICE_THM)
+
+lemma SBIT_DIV: "ALL b m n. n < m --> SBIT b (m - n) = SBIT b m div 2 ^ n"
+  by (import bits SBIT_DIV)
+
+lemma BITS_SUC: "ALL h l n.
+   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 l n.
+   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 l a b.
+   (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 oper a b. BITWISE n oper a b < 2 ^ n"
+  by (import bits BITWISE_LT_2EXP)
+
+lemma LESS_EXP_MULT2: "ALL (a::nat) b::nat.
+   a < b -->
+   (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 n oper a b.
+   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 n oper a b.
+   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 n oper a b.
+   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::nat.
+   (0::nat) < n -->
+   (ALL (j::nat) k::nat. (j + k mod n) mod n = (j + k) mod n)"
+  by (import bits MOD_PLUS_RIGHT)
+
+lemma MOD_PLUS_1: "ALL n::nat.
+   (0::nat) < n -->
+   (ALL x::nat.
+       ((x + (1::nat)) mod n = (0::nat)) = (x mod n + (1::nat) = n))"
+  by (import bits MOD_PLUS_1)
+
+lemma MOD_ADD_1: "ALL n::nat.
+   (0::nat) < n -->
+   (ALL x::nat.
+       (x + (1::nat)) mod n ~= (0::nat) -->
+       (x + (1::nat)) mod n = x mod n + (1::nat))"
+  by (import bits MOD_ADD_1)
+
+;end_setup
+
+;setup_theory word32
+
+consts
+  HB :: "nat" 
+
+defs
+  HB_primdef: "HB ==
+NUMERAL
+ (NUMERAL_BIT1
+   (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))"
+
+lemma HB_def: "HB =
+NUMERAL
+ (NUMERAL_BIT1
+   (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))"
+  by (import word32 HB_def)
+
+consts
+  WL :: "nat" 
+
+defs
+  WL_primdef: "WL == Suc HB"
+
+lemma WL_def: "WL = Suc HB"
+  by (import word32 WL_def)
+
+consts
+  MODw :: "nat => nat" 
+
+defs
+  MODw_primdef: "MODw == %n. n mod 2 ^ WL"
+
+lemma MODw_def: "ALL n. MODw n = n mod 2 ^ WL"
+  by (import word32 MODw_def)
+
+consts
+  INw :: "nat => bool" 
+
+defs
+  INw_primdef: "INw == %n. n < 2 ^ WL"
+
+lemma INw_def: "ALL n. INw n = (n < 2 ^ WL)"
+  by (import word32 INw_def)
+
+consts
+  EQUIV :: "nat => nat => bool" 
+
+defs
+  EQUIV_primdef: "EQUIV == %x y. MODw x = MODw y"
+
+lemma EQUIV_def: "ALL x y. EQUIV x y = (MODw x = MODw y)"
+  by (import word32 EQUIV_def)
+
+lemma EQUIV_QT: "ALL x y. EQUIV x y = (EQUIV x = EQUIV y)"
+  by (import word32 EQUIV_QT)
+
+lemma FUNPOW_THM: "ALL f n x. (f ^ n) (f x) = f ((f ^ n) x)"
+  by (import word32 FUNPOW_THM)
+
+lemma FUNPOW_THM2: "ALL f n x. (f ^ Suc n) x = f ((f ^ n) x)"
+  by (import word32 FUNPOW_THM2)
+
+lemma FUNPOW_COMP: "ALL f m n a. (f ^ m) ((f ^ n) a) = (f ^ (m + n)) a"
+  by (import word32 FUNPOW_COMP)
+
+lemma INw_MODw: "ALL n. INw (MODw n)"
+  by (import word32 INw_MODw)
+
+lemma TOw_IDEM: "ALL a. INw a --> MODw a = a"
+  by (import word32 TOw_IDEM)
+
+lemma MODw_IDEM2: "ALL a. MODw (MODw a) = MODw a"
+  by (import word32 MODw_IDEM2)
+
+lemma TOw_QT: "ALL a. 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 b. MODw (a + b) = MODw (MODw a + MODw b)"
+  by (import word32 MOD_ADD)
+
+lemma MODw_MULT: "ALL a b. MODw (a * b) = MODw (MODw a * MODw b)"
+  by (import word32 MODw_MULT)
+
+consts
+  AONE :: "nat" 
+
+defs
+  AONE_primdef: "AONE == 1"
+
+lemma AONE_def: "AONE = 1"
+  by (import word32 AONE_def)
+
+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. EQUIV (a + 0) a"
+  by (import word32 ADD_0_QT)
+
+lemma ADD_COMM_QT: "ALL a b. EQUIV (a + b) (b + a)"
+  by (import word32 ADD_COMM_QT)
+
+lemma ADD_ASSOC_QT: "ALL a b c. EQUIV (a + (b + c)) (a + b + c)"
+  by (import word32 ADD_ASSOC_QT)
+
+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. EQUIV (Suc m) (m + AONE)"
+  by (import word32 ADD1_QT)
+
+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 b. EQUIV (Suc a) b --> EQUIV a (b + (2 ^ WL - 1))"
+  by (import word32 SUC_EQUIV_COMP)
+
+lemma INV_SUC_EQ_QT: "ALL m n. EQUIV (Suc m) (Suc n) = EQUIV m n"
+  by (import word32 INV_SUC_EQ_QT)
+
+lemma ADD_INV_0_QT: "ALL m n. EQUIV (m + n) m --> EQUIV n 0"
+  by (import word32 ADD_INV_0_QT)
+
+lemma ADD_INV_0_EQ_QT: "ALL m n. EQUIV (m + n) m = EQUIV n 0"
+  by (import word32 ADD_INV_0_EQ_QT)
+
+lemma EQ_ADD_LCANCEL_QT: "ALL m n p. EQUIV (m + n) (m + p) = EQUIV n p"
+  by (import word32 EQ_ADD_LCANCEL_QT)
+
+lemma EQ_ADD_RCANCEL_QT: "ALL x xa xb. EQUIV (x + xb) (xa + xb) = EQUIV x xa"
+  by (import word32 EQ_ADD_RCANCEL_QT)
+
+lemma LEFT_ADD_DISTRIB_QT: "ALL m n p. EQUIV (p * (m + n)) (p * m + p * n)"
+  by (import word32 LEFT_ADD_DISTRIB_QT)
+
+lemma MULT_ASSOC_QT: "ALL m n p. EQUIV (m * (n * p)) (m * n * p)"
+  by (import word32 MULT_ASSOC_QT)
+
+lemma MULT_COMM_QT: "ALL m n. EQUIV (m * n) (n * m)"
+  by (import word32 MULT_COMM_QT)
+
+lemma MULT_CLAUSES_QT: "ALL m n.
+   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
+  MSBn :: "nat => bool" 
+
+defs
+  MSBn_primdef: "MSBn == bit HB"
+
+lemma MSBn_def: "MSBn = bit HB"
+  by (import word32 MSBn_def)
+
+consts
+  ONE_COMP :: "nat => nat" 
+
+defs
+  ONE_COMP_primdef: "ONE_COMP == %x. 2 ^ WL - 1 - MODw x"
+
+lemma ONE_COMP_def: "ALL x. 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. 2 ^ WL - MODw x"
+
+lemma TWO_COMP_def: "ALL x. TWO_COMP x = 2 ^ WL - MODw x"
+  by (import word32 TWO_COMP_def)
+
+lemma ADD_TWO_COMP_QT: "ALL a. EQUIV (MODw a + TWO_COMP a) 0"
+  by (import word32 ADD_TWO_COMP_QT)
+
+lemma TWO_COMP_ONE_COMP_QT: "ALL a. EQUIV (TWO_COMP a) (ONE_COMP a + AONE)"
+  by (import word32 TWO_COMP_ONE_COMP_QT)
+
+lemma BIT_EQUIV_THM: "ALL x xa. (ALL xb<WL. bit xb x = bit xb xa) = EQUIV x xa"
+  by (import word32 BIT_EQUIV_THM)
+
+lemma BITS_SUC2: "ALL n a. 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 b. BITWISE WL (%x y. ~ x) a b = ONE_COMP a"
+  by (import word32 BITWISE_ONE_COMP_THM)
+
+lemma ONE_COMP_THM: "ALL x xa. xa < WL --> bit xa (ONE_COMP x) = (~ bit xa x)"
+  by (import word32 ONE_COMP_THM)
+
+consts
+  OR :: "nat => nat => nat" 
+
+defs
+  OR_primdef: "OR == BITWISE WL op |"
+
+lemma OR_def: "OR = BITWISE WL op |"
+  by (import word32 OR_def)
+
+consts
+  AND :: "nat => nat => nat" 
+
+defs
+  AND_primdef: "AND == BITWISE WL op &"
+
+lemma AND_def: "AND = BITWISE WL op &"
+  by (import word32 AND_def)
+
+consts
+  EOR :: "nat => nat => nat" 
+
+defs
+  EOR_primdef: "EOR == BITWISE WL (%x y. x ~= y)"
+
+lemma EOR_def: "EOR = BITWISE WL (%x y. x ~= y)"
+  by (import word32 EOR_def)
+
+consts
+  COMP0 :: "nat" 
+
+defs
+  COMP0_primdef: "COMP0 == ONE_COMP 0"
+
+lemma COMP0_def: "COMP0 = ONE_COMP 0"
+  by (import word32 COMP0_def)
+
+lemma BITWISE_THM2: "ALL y oper a b.
+   (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 b c. EQUIV (OR a (OR b c)) (OR (OR a b) c)"
+  by (import word32 OR_ASSOC_QT)
+
+lemma OR_COMM_QT: "ALL a b. EQUIV (OR a b) (OR b a)"
+  by (import word32 OR_COMM_QT)
+
+lemma OR_ABSORB_QT: "ALL a b. EQUIV (AND a (OR a b)) a"
+  by (import word32 OR_ABSORB_QT)
+
+lemma OR_IDEM_QT: "ALL a. EQUIV (OR a a) a"
+  by (import word32 OR_IDEM_QT)
+
+lemma AND_ASSOC_QT: "ALL a b c. EQUIV (AND a (AND b c)) (AND (AND a b) c)"
+  by (import word32 AND_ASSOC_QT)
+
+lemma AND_COMM_QT: "ALL a b. EQUIV (AND a b) (AND b a)"
+  by (import word32 AND_COMM_QT)
+
+lemma AND_ABSORB_QT: "ALL a b. EQUIV (OR a (AND a b)) a"
+  by (import word32 AND_ABSORB_QT)
+
+lemma AND_IDEM_QT: "ALL a. EQUIV (AND a a) a"
+  by (import word32 AND_IDEM_QT)
+
+lemma OR_COMP_QT: "ALL a. EQUIV (OR a (ONE_COMP a)) COMP0"
+  by (import word32 OR_COMP_QT)
+
+lemma AND_COMP_QT: "ALL a. EQUIV (AND a (ONE_COMP a)) 0"
+  by (import word32 AND_COMP_QT)
+
+lemma ONE_COMP_QT: "ALL a. EQUIV (ONE_COMP (ONE_COMP a)) a"
+  by (import word32 ONE_COMP_QT)
+
+lemma RIGHT_AND_OVER_OR_QT: "ALL a b c. 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 b c. 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 b.
+   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 a b. n < WL --> EQUIV a b --> bit n a = bit n b"
+  by (import word32 BIT_EQUIV)
+
+lemma LSB_WELLDEF: "ALL a b. EQUIV a b --> LSBn a = LSBn b"
+  by (import word32 LSB_WELLDEF)
+
+lemma MSB_WELLDEF: "ALL a b. EQUIV a b --> MSBn a = MSBn b"
+  by (import word32 MSB_WELLDEF)
+
+lemma BITWISE_ISTEP: "ALL n oper a b.
+   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 oper a b.
+   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 oper a b c d.
+   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 a b c d.
+   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 b. EQUIV a b --> EQUIV (Suc a) (Suc b)"
+  by (import word32 SUC_WELLDEF)
+
+lemma ADD_WELLDEF: "ALL a b c d. EQUIV a b & EQUIV c d --> EQUIV (a + c) (b + d)"
+  by (import word32 ADD_WELLDEF)
+
+lemma MUL_WELLDEF: "ALL a b c d. EQUIV a b & EQUIV c d --> EQUIV (a * c) (b * d)"
+  by (import word32 MUL_WELLDEF)
+
+lemma ONE_COMP_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ONE_COMP a) (ONE_COMP b)"
+  by (import word32 ONE_COMP_WELLDEF)
+
+lemma TWO_COMP_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (TWO_COMP a) (TWO_COMP b)"
+  by (import word32 TWO_COMP_WELLDEF)
+
+lemma TOw_WELLDEF: "ALL a b. 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. MODw a div 2"
+
+lemma LSR_ONE_def: "ALL a. 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. LSR_ONE a + SBIT (MSBn a) HB"
+
+lemma ASR_ONE_def: "ALL a. 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. LSR_ONE a + SBIT (LSBn a) HB"
+
+lemma ROR_ONE_def: "ALL a. 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 a. LSR_ONE a + SBIT c HB"
+
+lemma RRXn_def: "ALL c a. RRXn c a = LSR_ONE a + SBIT c HB"
+  by (import word32 RRXn_def)
+
+lemma LSR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (LSR_ONE a) (LSR_ONE b)"
+  by (import word32 LSR_ONE_WELLDEF)
+
+lemma ASR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ASR_ONE a) (ASR_ONE b)"
+  by (import word32 ASR_ONE_WELLDEF)
+
+lemma ROR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ROR_ONE a) (ROR_ONE b)"
+  by (import word32 ROR_ONE_WELLDEF)
+
+lemma RRX_WELLDEF: "ALL a b c. 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. EX xa. x = EQUIV xa}" 
+  by (rule typedef_helper,import word32 word32_TY_DEF)
+
+lemmas word32_TY_DEF = typedef_hol2hol4 [OF type_definition_word32]
+
+consts
+  mk_word32 :: "(nat => bool) => word32" 
+  dest_word32 :: "word32 => nat => bool" 
+
+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
+  w_0 :: "word32" 
+
+defs
+  w_0_primdef: "w_0 == mk_word32 (EQUIV 0)"
+
+lemma w_0_def: "w_0 = mk_word32 (EQUIV 0)"
+  by (import word32 w_0_def)
+
+consts
+  w_1 :: "word32" 
+
+defs
+  w_1_primdef: "w_1 == mk_word32 (EQUIV AONE)"
+
+lemma w_1_def: "w_1 = mk_word32 (EQUIV AONE)"
+  by (import word32 w_1_def)
+
+consts
+  w_T :: "word32" 
+
+defs
+  w_T_primdef: "w_T == mk_word32 (EQUIV COMP0)"
+
+lemma w_T_def: "w_T = mk_word32 (EQUIV COMP0)"
+  by (import word32 w_T_def)
+
+constdefs
+  word_suc :: "word32 => word32" 
+  "word_suc == %T1. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
+
+lemma word_suc: "ALL T1. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))"
+  by (import word32 word_suc)
+
+constdefs
+  word_add :: "word32 => word32 => word32" 
+  "word_add ==
+%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
+
+lemma word_add: "ALL T1 T2.
+   word_add T1 T2 =
+   mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))"
+  by (import word32 word_add)
+
+constdefs
+  word_mul :: "word32 => word32 => word32" 
+  "word_mul ==
+%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
+
+lemma word_mul: "ALL T1 T2.
+   word_mul T1 T2 =
+   mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))"
+  by (import word32 word_mul)
+
+constdefs
+  word_1comp :: "word32 => word32" 
+  "word_1comp == %T1. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
+
+lemma word_1comp: "ALL T1. word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))"
+  by (import word32 word_1comp)
+
+constdefs
+  word_2comp :: "word32 => word32" 
+  "word_2comp == %T1. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
+
+lemma word_2comp: "ALL T1. word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))"
+  by (import word32 word_2comp)
+
+constdefs
+  word_lsr1 :: "word32 => word32" 
+  "word_lsr1 == %T1. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
+
+lemma word_lsr1: "ALL T1. word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))"
+  by (import word32 word_lsr1)
+
+constdefs
+  word_asr1 :: "word32 => word32" 
+  "word_asr1 == %T1. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
+
+lemma word_asr1: "ALL T1. word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))"
+  by (import word32 word_asr1)
+
+constdefs
+  word_ror1 :: "word32 => word32" 
+  "word_ror1 == %T1. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))"
+
+lemma word_ror1: "ALL T1. 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 T2. mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))"
+
+lemma RRX_def: "ALL T1 T2. 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. LSBn (Eps (dest_word32 T1))"
+
+lemma LSB_def: "ALL T1. LSB T1 = LSBn (Eps (dest_word32 T1))"
+  by (import word32 LSB_def)
+
+consts
+  MSB :: "word32 => bool" 
+
+defs
+  MSB_primdef: "MSB == %T1. MSBn (Eps (dest_word32 T1))"
+
+lemma MSB_def: "ALL T1. MSB T1 = MSBn (Eps (dest_word32 T1))"
+  by (import word32 MSB_def)
+
+constdefs
+  bitwise_or :: "word32 => word32 => word32" 
+  "bitwise_or ==
+%T1 T2. mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+
+lemma bitwise_or: "ALL T1 T2.
+   bitwise_or T1 T2 =
+   mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+  by (import word32 bitwise_or)
+
+constdefs
+  bitwise_eor :: "word32 => word32 => word32" 
+  "bitwise_eor ==
+%T1 T2.
+   mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+
+lemma bitwise_eor: "ALL T1 T2.
+   bitwise_eor T1 T2 =
+   mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+  by (import word32 bitwise_eor)
+
+constdefs
+  bitwise_and :: "word32 => word32 => word32" 
+  "bitwise_and ==
+%T1 T2.
+   mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))"
+
+lemma bitwise_and: "ALL T1 T2.
+   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. mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
+
+lemma TOw_def: "ALL T1. TOw T1 = mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))"
+  by (import word32 TOw_def)
+
+consts
+  n2w :: "nat => word32" 
+
+defs
+  n2w_primdef: "n2w == %n. mk_word32 (EQUIV n)"
+
+lemma n2w_def: "ALL n. n2w n = mk_word32 (EQUIV n)"
+  by (import word32 n2w_def)
+
+consts
+  w2n :: "word32 => nat" 
+
+defs
+  w2n_primdef: "w2n == %w. MODw (Eps (dest_word32 w))"
+
+lemma w2n_def: "ALL w. w2n w = MODw (Eps (dest_word32 w))"
+  by (import word32 w2n_def)
+
+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. word_add x w_0 = x"
+  by (import word32 ADD_0w)
+
+lemma ADD1w: "ALL x. word_suc x = word_add x w_1"
+  by (import word32 ADD1w)
+
+lemma ADD_ASSOCw: "ALL x xa xb. word_add x (word_add xa xb) = word_add (word_add x xa) xb"
+  by (import word32 ADD_ASSOCw)
+
+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 xa. word_add x xa = word_add xa x"
+  by (import word32 ADD_COMMw)
+
+lemma ADD_INV_0_EQw: "ALL x xa. (word_add x xa = x) = (xa = w_0)"
+  by (import word32 ADD_INV_0_EQw)
+
+lemma EQ_ADD_LCANCELw: "ALL x xa xb. (word_add x xa = word_add x xb) = (xa = xb)"
+  by (import word32 EQ_ADD_LCANCELw)
+
+lemma EQ_ADD_RCANCELw: "ALL x xa xb. (word_add x xb = word_add xa xb) = (x = xa)"
+  by (import word32 EQ_ADD_RCANCELw)
+
+lemma LEFT_ADD_DISTRIBw: "ALL x xa xb.
+   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 xa xb. word_mul x (word_mul xa xb) = word_mul (word_mul x xa) xb"
+  by (import word32 MULT_ASSOCw)
+
+lemma MULT_COMMw: "ALL x xa. word_mul x xa = word_mul xa x"
+  by (import word32 MULT_COMMw)
+
+lemma MULT_CLAUSESw: "ALL x xa.
+   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. word_2comp x = word_add (word_1comp x) w_1"
+  by (import word32 TWO_COMP_ONE_COMP)
+
+lemma OR_ASSOCw: "ALL x xa xb.
+   bitwise_or x (bitwise_or xa xb) = bitwise_or (bitwise_or x xa) xb"
+  by (import word32 OR_ASSOCw)
+
+lemma OR_COMMw: "ALL x xa. bitwise_or x xa = bitwise_or xa x"
+  by (import word32 OR_COMMw)
+
+lemma OR_IDEMw: "ALL x. bitwise_or x x = x"
+  by (import word32 OR_IDEMw)
+
+lemma OR_ABSORBw: "ALL x xa. bitwise_and x (bitwise_or x xa) = x"
+  by (import word32 OR_ABSORBw)
+
+lemma AND_ASSOCw: "ALL x xa xb.
+   bitwise_and x (bitwise_and xa xb) = bitwise_and (bitwise_and x xa) xb"
+  by (import word32 AND_ASSOCw)
+
+lemma AND_COMMw: "ALL x xa. bitwise_and x xa = bitwise_and xa x"
+  by (import word32 AND_COMMw)
+
+lemma AND_IDEMw: "ALL x. bitwise_and x x = x"
+  by (import word32 AND_IDEMw)
+
+lemma AND_ABSORBw: "ALL x xa. bitwise_or x (bitwise_and x xa) = x"
+  by (import word32 AND_ABSORBw)
+
+lemma ONE_COMPw: "ALL x. word_1comp (word_1comp x) = x"
+  by (import word32 ONE_COMPw)
+
+lemma RIGHT_AND_OVER_ORw: "ALL x xa xb.
+   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 xa xb.
+   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 xa.
+   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"
+  by (import word32 w_0)
+
+lemma w_1: "w_1 = n2w 1"
+  by (import word32 w_1)
+
+lemma w_T: "w_T =
+n2w (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
+            (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 w_T)
+
+lemma ADD_TWO_COMP: "ALL x. word_add x (word_2comp x) = w_0"
+  by (import word32 ADD_TWO_COMP)
+
+lemma ADD_TWO_COMP2: "ALL x. word_add (word_2comp x) x = w_0"
+  by (import word32 ADD_TWO_COMP2)
+
+constdefs
+  word_sub :: "word32 => word32 => word32" 
+  "word_sub == %a b. word_add a (word_2comp b)"
+
+lemma word_sub: "ALL a b. word_sub a b = word_add a (word_2comp b)"
+  by (import word32 word_sub)
+
+constdefs
+  word_lsl :: "word32 => nat => word32" 
+  "word_lsl == %a n. word_mul a (n2w (2 ^ n))"
+
+lemma word_lsl: "ALL a n. word_lsl a n = word_mul a (n2w (2 ^ n))"
+  by (import word32 word_lsl)
+
+constdefs
+  word_lsr :: "word32 => nat => word32" 
+  "word_lsr == %a n. (word_lsr1 ^ n) a"
+
+lemma word_lsr: "ALL a n. word_lsr a n = (word_lsr1 ^ n) a"
+  by (import word32 word_lsr)
+
+constdefs
+  word_asr :: "word32 => nat => word32" 
+  "word_asr == %a n. (word_asr1 ^ n) a"
+
+lemma word_asr: "ALL a n. word_asr a n = (word_asr1 ^ n) a"
+  by (import word32 word_asr)
+
+constdefs
+  word_ror :: "word32 => nat => word32" 
+  "word_ror == %a n. (word_ror1 ^ n) a"
+
+lemma word_ror: "ALL a n. word_ror a n = (word_ror1 ^ n) a"
+  by (import word32 word_ror)
+
+consts
+  BITw :: "nat => word32 => bool" 
+
+defs
+  BITw_primdef: "BITw == %b n. bit b (w2n n)"
+
+lemma BITw_def: "ALL b n. BITw b n = bit b (w2n n)"
+  by (import word32 BITw_def)
+
+consts
+  BITSw :: "nat => nat => word32 => nat" 
+
+defs
+  BITSw_primdef: "BITSw == %h l n. BITS h l (w2n n)"
+
+lemma BITSw_def: "ALL h l n. 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 l n. SLICE h l (w2n n)"
+
+lemma SLICEw_def: "ALL h l n. SLICEw h l n = SLICE h l (w2n n)"
+  by (import word32 SLICEw_def)
+
+lemma TWO_COMP_ADD: "ALL a b. 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. word_2comp (word_2comp a) = a"
+  by (import word32 TWO_COMP_ELIM)
+
+lemma ADD_SUB_ASSOC: "ALL a b c. 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 b c. 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. word_sub a a = w_0"
+  by (import word32 SUB_EQUALw)
+
+lemma ADD_SUBw: "ALL a b. word_sub (word_add a b) b = a"
+  by (import word32 ADD_SUBw)
+
+lemma SUB_SUBw: "ALL a b c. 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. word_1comp a = word_sub (word_2comp a) w_1"
+  by (import word32 ONE_COMP_TWO_COMP)
+
+lemma SUBw: "ALL m n. word_sub (word_suc m) n = word_suc (word_sub m n)"
+  by (import word32 SUBw)
+
+lemma ADD_EQ_SUBw: "ALL m n p. (word_add m n = p) = (m = word_sub p n)"
+  by (import word32 ADD_EQ_SUBw)
+
+lemma CANCEL_SUBw: "ALL m n p. (word_sub n p = word_sub m p) = (n = m)"
+  by (import word32 CANCEL_SUBw)
+
+lemma SUB_PLUSw: "ALL a b c. word_sub a (word_add b c) = word_sub (word_sub a b) c"
+  by (import word32 SUB_PLUSw)
+
+lemma word_nchotomy: "ALL w. EX n. w = n2w n"
+  by (import word32 word_nchotomy)
+
+lemma dest_word_mk_word_eq3: "ALL a. dest_word32 (mk_word32 (EQUIV a)) = EQUIV a"
+  by (import word32 dest_word_mk_word_eq3)
+
+lemma MODw_ELIM: "ALL n. n2w (MODw n) = n2w n"
+  by (import word32 MODw_ELIM)
+
+lemma w2n_EVAL: "ALL n. w2n (n2w n) = MODw n"
+  by (import word32 w2n_EVAL)
+
+lemma w2n_ELIM: "ALL a. n2w (w2n a) = a"
+  by (import word32 w2n_ELIM)
+
+lemma n2w_11: "ALL a b. (n2w a = n2w b) = (MODw a = MODw b)"
+  by (import word32 n2w_11)
+
+lemma ADD_EVAL: "word_add (n2w a) (n2w b) = n2w (a + b)"
+  by (import word32 ADD_EVAL)
+
+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) = n2w (ONE_COMP a)"
+  by (import word32 ONE_COMP_EVAL)
+
+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) = n2w (LSR_ONE a)"
+  by (import word32 LSR_ONE_EVAL)
+
+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) = n2w (ROR_ONE a)"
+  by (import word32 ROR_ONE_EVAL)
+
+lemma RRX_EVAL: "RRX c (n2w a) = n2w (RRXn c a)"
+  by (import word32 RRX_EVAL)
+
+lemma LSB_EVAL: "LSB (n2w a) = LSBn a"
+  by (import word32 LSB_EVAL)
+
+lemma MSB_EVAL: "MSB (n2w a) = MSBn a"
+  by (import word32 MSB_EVAL)
+
+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) (n2w b) = n2w (EOR a b)"
+  by (import word32 EOR_EVAL)
+
+lemma AND_EVAL: "bitwise_and (n2w a) (n2w b) = n2w (AND a b)"
+  by (import word32 AND_EVAL)
+
+lemma BITS_EVAL: "ALL h l a. BITSw h l (n2w a) = BITS h l (MODw a)"
+  by (import word32 BITS_EVAL)
+
+lemma BIT_EVAL: "ALL b a. BITw b (n2w a) = bit b (MODw a)"
+  by (import word32 BIT_EVAL)
+
+lemma SLICE_EVAL: "ALL h l a. SLICEw h l (n2w a) = SLICE h l (MODw a)"
+  by (import word32 SLICE_EVAL)
+
+lemma LSL_ADD: "ALL a m n. word_lsl (word_lsl a m) n = word_lsl a (m + n)"
+  by (import word32 LSL_ADD)
+
+lemma LSR_ADD: "ALL x xa xb. word_lsr (word_lsr x xa) xb = word_lsr x (xa + xb)"
+  by (import word32 LSR_ADD)
+
+lemma ASR_ADD: "ALL x xa xb. word_asr (word_asr x xa) xb = word_asr x (xa + xb)"
+  by (import word32 ASR_ADD)
+
+lemma ROR_ADD: "ALL x xa xb. word_ror (word_ror x xa) xb = word_ror x (xa + xb)"
+  by (import word32 ROR_ADD)
+
+lemma LSL_LIMIT: "ALL w n. HB < n --> word_lsl w n = w_0"
+  by (import word32 LSL_LIMIT)
+
+lemma MOD_MOD_DIV: "ALL a b. INw (MODw a div 2 ^ b)"
+  by (import word32 MOD_MOD_DIV)
+
+lemma MOD_MOD_DIV_2EXP: "ALL a n. 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. word_lsr (n2w a) n = n2w (MODw a div 2 ^ n)"
+  by (import word32 LSR_EVAL)
+
+lemma LSR_THM: "ALL x n. word_lsr (n2w n) x = n2w (BITS HB (min WL x) n)"
+  by (import word32 LSR_THM)
+
+lemma LSR_LIMIT: "ALL x w. 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::nat) ^ m -->
+   (2::nat) ^ n + a * (2::nat) ^ n <= (2::nat) ^ (m + n)"
+  by (import word32 LEFT_SHIFT_LESS)
+
+lemma ROR_THM: "ALL x n.
+   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 w. word_ror w (x * WL) = w"
+  by (import word32 ROR_CYCLE)
+
+lemma ASR_THM: "ALL x n.
+   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 w. 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. 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. 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. word_asr w_T n = w_T"
+  by (import word32 ASR_w_T)
+
+lemma ROR_w_T: "ALL n. word_ror w_T n = w_T"
+  by (import word32 ROR_w_T)
+
+lemma MODw_EVAL: "ALL x.
+   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 a. word_add (n2w a) (n2w b) = n2w (MODw (a + b))"
+  by (import word32 ADD_EVAL2)
+
+lemma MUL_EVAL2: "ALL b a. word_mul (n2w a) (n2w b) = n2w (MODw (a * b))"
+  by (import word32 MUL_EVAL2)
+
+lemma ONE_COMP_EVAL2: "ALL a.
+   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.
+   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. word_lsr1 (n2w a) = n2w (MODw a div 2)"
+  by (import word32 LSR_ONE_EVAL2)
+
+lemma ASR_ONE_EVAL2: "ALL a.
+   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: "ALL a.
+   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 a.
+   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. LSB (n2w a) = ODD a"
+  by (import word32 LSB_EVAL2)
+
+lemma MSB_EVAL2: "ALL a.
+   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 a.
+   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 a.
+   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 a.
+   bitwise_eor (n2w a) (n2w b) =
+   n2w (BITWISE
+         (NUMERAL
+           (NUMERAL_BIT2
+             (NUMERAL_BIT1
+               (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))))
+         (%x y. x ~= y) a b)"
+  by (import word32 EOR_EVAL2)
+
+lemma BITWISE_EVAL2: "ALL n oper x y.
+   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 l n. BITSw h l n < 2 ^ (Suc h - l)"
+  by (import word32 BITSwLT_THM)
+
+lemma BITSw_COMP_THM: "ALL h1 l1 h2 l2 n.
+   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 l n x. BITSw h l x div 2 ^ n = BITSw h (l + n) x"
+  by (import word32 BITSw_DIV_THM)
+
+lemma BITw_THM: "ALL b n. BITw b n = (BITSw b b n = 1)"
+  by (import word32 BITw_THM)
+
+lemma SLICEw_THM: "ALL n h l. SLICEw h l n = BITSw h l n * 2 ^ l"
+  by (import word32 SLICEw_THM)
+
+lemma BITS_SLICEw_THM: "ALL h l n. BITS h l (SLICEw h l n) = BITSw h l n"
+  by (import word32 BITS_SLICEw_THM)
+
+lemma SLICEw_ZERO_THM: "ALL n h. SLICEw h 0 n = BITSw h 0 n"
+  by (import word32 SLICEw_ZERO_THM)
+
+lemma SLICEw_COMP_THM: "ALL h m l a.
+   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 l n. h < l --> BITSw h l n = 0"
+  by (import word32 BITSw_ZERO)
+
+lemma SLICEw_ZERO: "ALL h l n. h < l --> SLICEw h l n = 0"
+  by (import word32 SLICEw_ZERO)
+
+;end_setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/README	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,3 @@
+All the files in this directory (except this README, HOL4.thy, and
+ROOT.ML) are automatically generated.  Edit the files in
+../Generate-HOL, if something needs to be changed.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/ROOT.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,4 @@
+with_path ".." use_thy "HOL4Compat";
+with_path ".." use_thy "HOL4Syntax";
+setmp quick_and_dirty true use_thy "HOL4Prob";
+setmp quick_and_dirty true use_thy "HOL4";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/arithmetic.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,271 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "nat_elim__magic" > "nat_elim__magic_def"
+  "ODD" > "ODD_def"
+  "FACT" > "FACT_def"
+  "EVEN" > "EVEN_def"
+
+const_maps
+  "num_case" > "Nat.nat.nat_case"
+  "nat_elim__magic" > "HOL4Base.arithmetic.nat_elim__magic"
+  "NUMERAL_BIT2" > "HOL4Compat.NUMERAL_BIT2"
+  "NUMERAL_BIT1" > "HOL4Compat.NUMERAL_BIT1"
+  "NUMERAL" > "HOL4Compat.NUMERAL"
+  "MOD" > "Divides.op mod" :: "nat => nat => nat"
+  "MIN" > "HOL.min" :: "nat => nat => nat"
+  "MAX" > "HOL.max" :: "nat => nat => nat"
+  "FUNPOW" > "HOL4Compat.FUNPOW"
+  "EXP" > "Nat.power" :: "nat => nat => nat"
+  "DIV" > "Divides.op div" :: "nat => nat => nat"
+  "ALT_ZERO" > "HOL4Compat.ALT_ZERO"
+  ">=" > "HOL4Compat.nat_ge"
+  ">" > "HOL4Compat.nat_gt"
+  "<=" > "op <=" :: "nat => nat => bool"
+  "-" > "op -" :: "nat => nat => nat"
+  "+" > "op +" :: "nat => nat => nat"
+  "*" > "op *" :: "nat => nat => nat"
+
+thm_maps
+  "num_case_def" > "HOL4Compat.num_case_def"
+  "num_case_cong" > "HOL4Base.arithmetic.num_case_cong"
+  "num_case_compute" > "HOL4Base.arithmetic.num_case_compute"
+  "num_CASES" > "Nat.nat.nchotomy"
+  "nat_elim__magic_def" > "HOL4Base.arithmetic.nat_elim__magic_def"
+  "nat_elim__magic" > "HOL4Base.arithmetic.nat_elim__magic"
+  "ZERO_MOD" > "HOL4Base.arithmetic.ZERO_MOD"
+  "ZERO_LESS_EXP" > "HOL4Base.arithmetic.ZERO_LESS_EXP"
+  "ZERO_LESS_EQ" > "Nat.le0"
+  "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"
+  "SUC_ONE_ADD" > "NatBin.Suc_eq_add_numeral_1_left"
+  "SUC_NOT" > "Nat.nat.simps_2"
+  "SUC_ELIM_THM" > "HOL4Base.arithmetic.SUC_ELIM_THM"
+  "SUC_ADD_SYM" > "HOL4Base.arithmetic.SUC_ADD_SYM"
+  "SUB_SUB" > "NatArith.diff_diff_right"
+  "SUB_RIGHT_SUB" > "Nat.diff_diff_left"
+  "SUB_RIGHT_LESS_EQ" > "HOL4Base.arithmetic.SUB_RIGHT_LESS_EQ"
+  "SUB_RIGHT_LESS" > "HOL4Base.arithmetic.SUB_RIGHT_LESS"
+  "SUB_RIGHT_GREATER_EQ" > "HOL4Base.arithmetic.SUB_RIGHT_GREATER_EQ"
+  "SUB_RIGHT_GREATER" > "HOL4Base.arithmetic.SUB_RIGHT_GREATER"
+  "SUB_RIGHT_EQ" > "HOL4Base.arithmetic.SUB_RIGHT_EQ"
+  "SUB_RIGHT_ADD" > "HOL4Base.arithmetic.SUB_RIGHT_ADD"
+  "SUB_PLUS" > "Nat.diff_diff_left"
+  "SUB_MONO_EQ" > "Nat.diff_Suc_Suc"
+  "SUB_LESS_OR" > "HOL4Base.arithmetic.SUB_LESS_OR"
+  "SUB_LESS_EQ_ADD" > "HOL4Base.arithmetic.SUB_LESS_EQ_ADD"
+  "SUB_LESS_EQ" > "Nat.diff_le_self"
+  "SUB_LESS_0" > "Nat.zero_less_diff"
+  "SUB_LEFT_SUC" > "HOL4Base.arithmetic.SUB_LEFT_SUC"
+  "SUB_LEFT_SUB" > "HOL4Base.arithmetic.SUB_LEFT_SUB"
+  "SUB_LEFT_LESS_EQ" > "HOL4Base.arithmetic.SUB_LEFT_LESS_EQ"
+  "SUB_LEFT_LESS" > "NatArith.less_diff_conv"
+  "SUB_LEFT_GREATER_EQ" > "NatArith.le_diff_conv"
+  "SUB_LEFT_GREATER" > "HOL4Base.arithmetic.SUB_LEFT_GREATER"
+  "SUB_LEFT_EQ" > "HOL4Base.arithmetic.SUB_LEFT_EQ"
+  "SUB_LEFT_ADD" > "HOL4Base.arithmetic.SUB_LEFT_ADD"
+  "SUB_EQ_EQ_0" > "HOL4Base.arithmetic.SUB_EQ_EQ_0"
+  "SUB_EQ_0" > "Nat.diff_is_0_eq"
+  "SUB_EQUAL_0" > "Nat.diff_self_eq_0"
+  "SUB_ELIM_THM" > "HOL4Base.arithmetic.SUB_ELIM_THM"
+  "SUB_CANCEL" > "HOL4Base.arithmetic.SUB_CANCEL"
+  "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"
+  "PRE_SUC_EQ" > "HOL4Base.arithmetic.PRE_SUC_EQ"
+  "PRE_SUB1" > "HOL4Base.arithmetic.PRE_SUB1"
+  "PRE_SUB" > "HOL4Base.arithmetic.PRE_SUB"
+  "PRE_ELIM_THM" > "HOL4Base.arithmetic.PRE_ELIM_THM"
+  "OR_LESS" > "Nat.Suc_le_lessD"
+  "ONE" > "Nat.One_nat_def"
+  "ODD_OR_EVEN" > "HOL4Base.arithmetic.ODD_OR_EVEN"
+  "ODD_MULT" > "HOL4Base.arithmetic.ODD_MULT"
+  "ODD_EXISTS" > "HOL4Base.arithmetic.ODD_EXISTS"
+  "ODD_EVEN" > "HOL4Base.arithmetic.ODD_EVEN"
+  "ODD_DOUBLE" > "HOL4Base.arithmetic.ODD_DOUBLE"
+  "ODD_ADD" > "HOL4Base.arithmetic.ODD_ADD"
+  "ODD" > "HOL4Base.arithmetic.ODD"
+  "NUMERAL_DEF" > "HOL4Compat.NUMERAL_def"
+  "NUMERAL_BIT2" > "HOL4Compat.NUMERAL_BIT2_def"
+  "NUMERAL_BIT1" > "HOL4Compat.NUMERAL_BIT1_def"
+  "NOT_ZERO_LT_ZERO" > "Nat.neq0_conv"
+  "NOT_SUC_LESS_EQ_0" > "HOL4Base.arithmetic.NOT_SUC_LESS_EQ_0"
+  "NOT_SUC_LESS_EQ" > "HOL4Base.arithmetic.NOT_SUC_LESS_EQ"
+  "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" > "Nat.not_le_iff_less"
+  "NOT_LESS" > "Nat.not_less_iff_le"
+  "NOT_LEQ" > "HOL4Base.arithmetic.NOT_LEQ"
+  "NOT_GREATER_EQ" > "HOL4Base.arithmetic.NOT_GREATER_EQ"
+  "NOT_GREATER" > "Nat.not_less_iff_le"
+  "NOT_EXP_0" > "HOL4Base.arithmetic.NOT_EXP_0"
+  "NORM_0" > "HOL4Base.arithmetic.NORM_0"
+  "MULT_SYM" > "Nat.nat_mult_commute"
+  "MULT_SUC_EQ" > "HOL4Base.arithmetic.MULT_SUC_EQ"
+  "MULT_SUC" > "Nat.mult_Suc_right"
+  "MULT_RIGHT_1" > "Nat.nat_mult_1_right"
+  "MULT_MONO_EQ" > "Nat.Suc_mult_cancel1"
+  "MULT_LESS_EQ_SUC" > "Nat.Suc_mult_le_cancel1"
+  "MULT_LEFT_1" > "Nat.nat_mult_1"
+  "MULT_INCREASES" > "HOL4Base.arithmetic.MULT_INCREASES"
+  "MULT_EXP_MONO" > "HOL4Base.arithmetic.MULT_EXP_MONO"
+  "MULT_EQ_1" > "HOL4Base.arithmetic.MULT_EQ_1"
+  "MULT_EQ_0" > "Nat.mult_is_0"
+  "MULT_DIV" > "Divides.div_mult_self_is_m"
+  "MULT_COMM" > "Nat.nat_mult_commute"
+  "MULT_CLAUSES" > "HOL4Base.arithmetic.MULT_CLAUSES"
+  "MULT_ASSOC" > "Nat.nat_mult_assoc"
+  "MULT_0" > "Nat.mult_0_right"
+  "MULT" > "HOL4Compat.MULT"
+  "MOD_UNIQUE" > "HOL4Base.arithmetic.MOD_UNIQUE"
+  "MOD_TIMES2" > "HOL4Base.arithmetic.MOD_TIMES2"
+  "MOD_TIMES" > "HOL4Base.arithmetic.MOD_TIMES"
+  "MOD_PLUS" > "HOL4Base.arithmetic.MOD_PLUS"
+  "MOD_P" > "HOL4Base.arithmetic.MOD_P"
+  "MOD_ONE" > "Divides.mod_1"
+  "MOD_MULT_MOD" > "HOL4Base.arithmetic.MOD_MULT_MOD"
+  "MOD_MULT" > "HOL4Base.arithmetic.MOD_MULT"
+  "MOD_MOD" > "HOL4Base.arithmetic.MOD_MOD"
+  "MOD_EQ_0" > "HOL4Base.arithmetic.MOD_EQ_0"
+  "MOD_COMMON_FACTOR" > "HOL4Base.arithmetic.MOD_COMMON_FACTOR"
+  "MIN_MAX_PRED" > "HOL4Base.arithmetic.MIN_MAX_PRED"
+  "MIN_MAX_LT" > "HOL4Base.arithmetic.MIN_MAX_LT"
+  "MIN_MAX_EQ" > "HOL4Base.arithmetic.MIN_MAX_EQ"
+  "MIN_LT" > "HOL4Base.arithmetic.MIN_LT"
+  "MIN_LE" > "HOL4Base.arithmetic.MIN_LE"
+  "MIN_IDEM" > "HOL.min_same"
+  "MIN_DEF" > "HOL4Compat.MIN_DEF"
+  "MIN_COMM" > "HOL.min_ac_2"
+  "MIN_ASSOC" > "HOL.min_ac_1"
+  "MIN_0" > "HOL4Base.arithmetic.MIN_0"
+  "MAX_LT" > "HOL4Base.arithmetic.MAX_LT"
+  "MAX_LE" > "HOL4Base.arithmetic.MAX_LE"
+  "MAX_IDEM" > "HOL.max_same"
+  "MAX_DEF" > "HOL4Compat.MAX_DEF"
+  "MAX_COMM" > "HOL.max_ac_2"
+  "MAX_ASSOC" > "HOL.max_ac_1"
+  "MAX_0" > "HOL4Base.arithmetic.MAX_0"
+  "LESS_TRANS" > "Nat.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"
+  "LESS_OR_EQ_ADD" > "HOL4Base.arithmetic.LESS_OR_EQ_ADD"
+  "LESS_OR_EQ" > "HOL4Compat.LESS_OR_EQ"
+  "LESS_OR" > "Nat.Suc_leI"
+  "LESS_NOT_SUC" > "HOL4Base.arithmetic.LESS_NOT_SUC"
+  "LESS_MULT_MONO" > "Nat.Suc_mult_less_cancel1"
+  "LESS_MULT2" > "Ring_and_Field.mult_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" > "Ring_and_Field.add_less_imp_less_right"
+  "LESS_MONO_ADD_EQ" > "Ring_and_Field.add_less_cancel_right"
+  "LESS_MONO_ADD" > "Nat.add_less_mono1"
+  "LESS_MOD" > "Divides.mod_less"
+  "LESS_LESS_SUC" > "HOL4Base.arithmetic.LESS_LESS_SUC"
+  "LESS_LESS_EQ_TRANS" > "Nat.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_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" > "Nat.le_refl"
+  "LESS_EQ_MONO_ADD_EQ" > "Ring_and_Field.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_IMP_LESS_SUC" > "Nat.le_imp_less_Suc"
+  "LESS_EQ_EXISTS" > "HOL4Base.arithmetic.LESS_EQ_EXISTS"
+  "LESS_EQ_CASES" > "Nat.nat_le_linear"
+  "LESS_EQ_ANTISYM" > "HOL4Base.arithmetic.LESS_EQ_ANTISYM"
+  "LESS_EQ_ADD_SUB" > "Nat.diff_add_assoc"
+  "LESS_EQ_ADD" > "Nat.le_add1"
+  "LESS_EQ_0" > "Nat.le_0_eq"
+  "LESS_EQUAL_ANTISYM" > "Nat.le_anti_sym"
+  "LESS_EQUAL_ADD" > "HOL4Base.arithmetic.LESS_EQUAL_ADD"
+  "LESS_EQ" > "Nat.le_simps_3"
+  "LESS_DIV_EQ_ZERO" > "Divides.div_less"
+  "LESS_CASES_IMP" > "HOL4Base.arithmetic.LESS_CASES_IMP"
+  "LESS_CASES" > "HOL4Base.arithmetic.LESS_CASES"
+  "LESS_ANTISYM" > "HOL4Base.arithmetic.LESS_ANTISYM"
+  "LESS_ADD_SUC" > "HOL4Base.arithmetic.LESS_ADD_SUC"
+  "LESS_ADD_NONZERO" > "HOL4Base.arithmetic.LESS_ADD_NONZERO"
+  "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"
+  "LE" > "HOL4Base.arithmetic.LE"
+  "INV_PRE_LESS_EQ" > "HOL4Base.arithmetic.INV_PRE_LESS_EQ"
+  "INV_PRE_LESS" > "HOL4Base.arithmetic.INV_PRE_LESS"
+  "INV_PRE_EQ" > "HOL4Base.arithmetic.INV_PRE_EQ"
+  "GREATER_OR_EQ" > "HOL4Compat.GREATER_OR_EQ"
+  "GREATER_EQ" > "HOL4Compat.real_ge"
+  "GREATER_DEF" > "HOL4Compat.GREATER_DEF"
+  "FUN_EQ_LEMMA" > "HOL4Base.arithmetic.FUN_EQ_LEMMA"
+  "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_EQ_0" > "HOL4Base.arithmetic.EXP_EQ_0"
+  "EXP_ALWAYS_BIG_ENOUGH" > "HOL4Base.arithmetic.EXP_ALWAYS_BIG_ENOUGH"
+  "EXP_ADD" > "Power.power_add"
+  "EXP_1" > "HOL4Base.arithmetic.EXP_1"
+  "EXP" > "HOL4Compat.EXP"
+  "EXISTS_GREATEST" > "HOL4Base.arithmetic.EXISTS_GREATEST"
+  "EVEN_OR_ODD" > "HOL4Base.arithmetic.EVEN_OR_ODD"
+  "EVEN_ODD_EXISTS" > "HOL4Base.arithmetic.EVEN_ODD_EXISTS"
+  "EVEN_ODD" > "HOL4Base.arithmetic.EVEN_ODD"
+  "EVEN_MULT" > "HOL4Base.arithmetic.EVEN_MULT"
+  "EVEN_EXISTS" > "HOL4Base.arithmetic.EVEN_EXISTS"
+  "EVEN_DOUBLE" > "HOL4Base.arithmetic.EVEN_DOUBLE"
+  "EVEN_AND_ODD" > "HOL4Base.arithmetic.EVEN_AND_ODD"
+  "EVEN_ADD" > "HOL4Base.arithmetic.EVEN_ADD"
+  "EVEN" > "HOL4Base.arithmetic.EVEN"
+  "EQ_MULT_LCANCEL" > "NatBin.nat_mult_eq_cancel_disj"
+  "EQ_MONO_ADD_EQ" > "Nat.nat_add_right_cancel"
+  "EQ_LESS_EQ" > "HOL.order_eq_iff"
+  "EQ_ADD_RCANCEL" > "Nat.nat_add_right_cancel"
+  "EQ_ADD_LCANCEL" > "Nat.nat_add_left_cancel"
+  "DIV_UNIQUE" > "HOL4Base.arithmetic.DIV_UNIQUE"
+  "DIV_P" > "HOL4Base.arithmetic.DIV_P"
+  "DIV_ONE" > "Divides.div_1"
+  "DIV_MULT" > "HOL4Base.arithmetic.DIV_MULT"
+  "DIV_LESS_EQ" > "HOL4Base.arithmetic.DIV_LESS_EQ"
+  "DIV_LESS" > "Divides.div_less_dividend"
+  "DIV_DIV_DIV_MULT" > "HOL4Base.arithmetic.DIV_DIV_DIV_MULT"
+  "DIVMOD_ID" > "HOL4Base.arithmetic.DIVMOD_ID"
+  "DIVISION" > "HOL4Compat.DIVISION"
+  "DA" > "HOL4Base.arithmetic.DA"
+  "COMPLETE_INDUCTION" > "Nat.less_induct"
+  "CANCEL_SUB" > "NatArith.eq_diff_iff"
+  "ALT_ZERO" > "HOL4Compat.ALT_ZERO_def"
+  "ADD_SYM" > "Nat.nat_add_commute"
+  "ADD_SUC" > "Nat.add_Suc_right"
+  "ADD_SUB" > "Nat.diff_add_inverse2"
+  "ADD_MONO_LESS_EQ" > "Nat.nat_add_left_cancel_le"
+  "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" > "Nat.nat_add_commute"
+  "ADD_CLAUSES" > "HOL4Base.arithmetic.ADD_CLAUSES"
+  "ADD_ASSOC" > "Nat.nat_add_assoc"
+  "ADD_0" > "Nat.add_0_right"
+  "ADD1" > "Presburger.Suc_plus1"
+  "ADD" > "HOL4Compat.ADD"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/bits.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,115 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "bit" > "bit_def"
+  "TIMES_2EXP" > "TIMES_2EXP_primdef"
+  "SLICE" > "SLICE_primdef"
+  "SBIT" > "SBIT_primdef"
+  "MOD_2EXP" > "MOD_2EXP_primdef"
+  "LSBn" > "LSBn_primdef"
+  "DIV_2EXP" > "DIV_2EXP_primdef"
+  "DIVMOD_2EXP" > "DIVMOD_2EXP_primdef"
+  "DIV2" > "DIV2_primdef"
+  "BITWISE" > "BITWISE_primdef"
+  "BITS" > "BITS_primdef"
+
+const_maps
+  "bit" > "HOL4Word32.bits.bit"
+  "TIMES_2EXP" > "HOL4Word32.bits.TIMES_2EXP"
+  "SLICE" > "HOL4Word32.bits.SLICE"
+  "SBIT" > "HOL4Word32.bits.SBIT"
+  "MOD_2EXP" > "HOL4Word32.bits.MOD_2EXP"
+  "LSBn" > "HOL4Word32.bits.LSBn"
+  "DIV_2EXP" > "HOL4Word32.bits.DIV_2EXP"
+  "DIVMOD_2EXP" > "HOL4Word32.bits.DIVMOD_2EXP"
+  "DIV2" > "HOL4Word32.bits.DIV2"
+  "BITS" > "HOL4Word32.bits.BITS"
+
+const_renames
+  "BIT" > "bit"
+
+thm_maps
+  "bit_def" > "HOL4Word32.bits.bit_def"
+  "ZERO_LT_TWOEXP" > "HOL4Word32.bits.ZERO_LT_TWOEXP"
+  "TWOEXP_MONO2" > "HOL4Word32.bits.TWOEXP_MONO2"
+  "TWOEXP_MONO" > "HOL4Word32.bits.TWOEXP_MONO"
+  "TWOEXP_DIVISION" > "HOL4Word32.bits.TWOEXP_DIVISION"
+  "TIMES_2EXP_primdef" > "HOL4Word32.bits.TIMES_2EXP_primdef"
+  "TIMES_2EXP_def" > "HOL4Word32.bits.TIMES_2EXP_def"
+  "SUC_SUB" > "HOL4Word32.bits.SUC_SUB"
+  "SLICE_primdef" > "HOL4Word32.bits.SLICE_primdef"
+  "SLICE_def" > "HOL4Word32.bits.SLICE_def"
+  "SLICE_ZERO" > "HOL4Word32.bits.SLICE_ZERO"
+  "SLICE_THM" > "HOL4Word32.bits.SLICE_THM"
+  "SLICE_LEM3" > "HOL4Word32.bits.SLICE_LEM3"
+  "SLICE_LEM2" > "HOL4Word32.bits.SLICE_LEM2"
+  "SLICE_LEM1" > "HOL4Word32.bits.SLICE_LEM1"
+  "SLICE_COMP_THM" > "HOL4Word32.bits.SLICE_COMP_THM"
+  "SLICELT_THM" > "HOL4Word32.bits.SLICELT_THM"
+  "SBIT_primdef" > "HOL4Word32.bits.SBIT_primdef"
+  "SBIT_def" > "HOL4Word32.bits.SBIT_def"
+  "SBIT_DIV" > "HOL4Word32.bits.SBIT_DIV"
+  "ODD_MOD2_LEM" > "HOL4Word32.bits.ODD_MOD2_LEM"
+  "NOT_ZERO_ADD1" > "Nat.not0_implies_Suc"
+  "NOT_MOD2_LEM2" > "HOL4Word32.bits.NOT_MOD2_LEM2"
+  "NOT_MOD2_LEM" > "HOL4Word32.bits.NOT_MOD2_LEM"
+  "NOT_BITS2" > "HOL4Word32.bits.NOT_BITS2"
+  "NOT_BITS" > "HOL4Word32.bits.NOT_BITS"
+  "NOT_BIT" > "HOL4Word32.bits.NOT_BIT"
+  "MOD_PLUS_RIGHT" > "HOL4Word32.bits.MOD_PLUS_RIGHT"
+  "MOD_PLUS_1" > "HOL4Word32.bits.MOD_PLUS_1"
+  "MOD_ADD_1" > "HOL4Word32.bits.MOD_ADD_1"
+  "MOD_2EXP_primdef" > "HOL4Word32.bits.MOD_2EXP_primdef"
+  "MOD_2EXP_def" > "HOL4Word32.bits.MOD_2EXP_def"
+  "MOD_2EXP_MONO" > "HOL4Word32.bits.MOD_2EXP_MONO"
+  "MOD_2EXP_LT" > "HOL4Word32.bits.MOD_2EXP_LT"
+  "MOD_2EXP_LEM" > "HOL4Word32.bits.MOD_2EXP_LEM"
+  "LSBn_primdef" > "HOL4Word32.bits.LSBn_primdef"
+  "LSBn_def" > "HOL4Word32.bits.LSBn_def"
+  "LSB_ODD" > "HOL4Word32.bits.LSB_ODD"
+  "LESS_EXP_MULT2" > "HOL4Word32.bits.LESS_EXP_MULT2"
+  "LESS_EQ_EXP_MULT" > "HOL4Word32.bits.LESS_EQ_EXP_MULT"
+  "EXP_SUB_LESS_EQ" > "HOL4Word32.bits.EXP_SUB_LESS_EQ"
+  "EVEN_MOD2_LEM" > "HOL4Word32.bits.EVEN_MOD2_LEM"
+  "DIV_MULT_THM2" > "HOL4Word32.bits.DIV_MULT_THM2"
+  "DIV_MULT_THM" > "HOL4Word32.bits.DIV_MULT_THM"
+  "DIV_MULT_LEM" > "HOL4Word32.bits.DIV_MULT_LEM"
+  "DIV_MULT_1" > "HOL4Word32.bits.DIV_MULT_1"
+  "DIV_2EXP_primdef" > "HOL4Word32.bits.DIV_2EXP_primdef"
+  "DIV_2EXP_def" > "HOL4Word32.bits.DIV_2EXP_def"
+  "DIVMOD_2EXP_primdef" > "HOL4Word32.bits.DIVMOD_2EXP_primdef"
+  "DIVMOD_2EXP_def" > "HOL4Word32.bits.DIVMOD_2EXP_def"
+  "DIV2_primdef" > "HOL4Word32.bits.DIV2_primdef"
+  "DIV2_def" > "HOL4Word32.bits.DIV2_def"
+  "DIV1" > "HOL4Word32.bits.DIV1"
+  "BIT_def" > "HOL4Word32.bits.BIT_def"
+  "BIT_SLICE_THM" > "HOL4Word32.bits.BIT_SLICE_THM"
+  "BIT_SLICE_LEM" > "HOL4Word32.bits.BIT_SLICE_LEM"
+  "BIT_SLICE" > "HOL4Word32.bits.BIT_SLICE"
+  "BIT_COMP_THM3" > "HOL4Word32.bits.BIT_COMP_THM3"
+  "BIT_BITS_THM" > "HOL4Word32.bits.BIT_BITS_THM"
+  "BITWISE_def" > "HOL4Word32.bits.BITWISE_def"
+  "BITWISE_THM" > "HOL4Word32.bits.BITWISE_THM"
+  "BITWISE_NOT_COR" > "HOL4Word32.bits.BITWISE_NOT_COR"
+  "BITWISE_LT_2EXP" > "HOL4Word32.bits.BITWISE_LT_2EXP"
+  "BITWISE_COR" > "HOL4Word32.bits.BITWISE_COR"
+  "BITS_primdef" > "HOL4Word32.bits.BITS_primdef"
+  "BITS_def" > "HOL4Word32.bits.BITS_def"
+  "BITS_ZERO3" > "HOL4Word32.bits.BITS_ZERO3"
+  "BITS_ZERO2" > "HOL4Word32.bits.BITS_ZERO2"
+  "BITS_ZERO" > "HOL4Word32.bits.BITS_ZERO"
+  "BITS_THM" > "HOL4Word32.bits.BITS_THM"
+  "BITS_SUC_THM" > "HOL4Word32.bits.BITS_SUC_THM"
+  "BITS_SUC" > "HOL4Word32.bits.BITS_SUC"
+  "BITS_SLICE_THM2" > "HOL4Word32.bits.BITS_SLICE_THM2"
+  "BITS_SLICE_THM" > "HOL4Word32.bits.BITS_SLICE_THM"
+  "BITS_LT_HIGH" > "HOL4Word32.bits.BITS_LT_HIGH"
+  "BITS_DIV_THM" > "HOL4Word32.bits.BITS_DIV_THM"
+  "BITS_COMP_THM2" > "HOL4Word32.bits.BITS_COMP_THM2"
+  "BITS_COMP_THM" > "HOL4Word32.bits.BITS_COMP_THM"
+  "BITSLT_THM" > "HOL4Word32.bits.BITSLT_THM"
+  "BITS2_THM" > "HOL4Word32.bits.BITS2_THM"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/bool.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,195 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "RES_SELECT" > "RES_SELECT_def"
+  "RES_FORALL" > "RES_FORALL_def"
+  "RES_EXISTS_UNIQUE" > "RES_EXISTS_UNIQUE_def"
+  "RES_EXISTS" > "RES_EXISTS_def"
+  "RES_ABSTRACT" > "RES_ABSTRACT_def"
+  "IN" > "IN_def"
+  "ARB" > "ARB_def"
+
+const_maps
+  "~" > "Not"
+  "bool_case" > "Datatype.bool.bool_case"
+  "\\/" > "op |"
+  "TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
+  "T" > "True"
+  "RES_SELECT" > "HOL4Base.bool.RES_SELECT"
+  "RES_FORALL" > "HOL4Base.bool.RES_FORALL"
+  "RES_EXISTS_UNIQUE" > "HOL4Base.bool.RES_EXISTS_UNIQUE"
+  "RES_EXISTS" > "HOL4Base.bool.RES_EXISTS"
+  "ONTO" > "HOL4Setup.ONTO"
+  "ONE_ONE" > "HOL4Setup.ONE_ONE"
+  "LET" > "HOL4Compat.LET"
+  "IN" > "HOL4Base.bool.IN"
+  "F" > "False"
+  "COND" > "If"
+  "ARB" > "HOL4Base.bool.ARB"
+  "?!" > "Ex1"
+  "?" > "Ex"
+  "/\\" > "op &"
+  "!" > "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" > "Set.bool_induct"
+  "boolAxiom" > "HOL4Base.bool.boolAxiom"
+  "UNWIND_THM2" > "HOL.simp_thms_39"
+  "UNWIND_THM1" > "HOL.simp_thms_40"
+  "UNWIND_FORALL_THM2" > "HOL.simp_thms_41"
+  "UNWIND_FORALL_THM1" > "HOL.simp_thms_42"
+  "UEXISTS_SIMP" > "HOL4Base.bool.UEXISTS_SIMP"
+  "UEXISTS_OR_THM" > "HOL4Base.bool.UEXISTS_OR_THM"
+  "T_DEF" > "HOL.True_def"
+  "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"
+  "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"
+  "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_FORALL_THM" > "HOL.all_simps_2"
+  "RES_SELECT_def" > "HOL4Base.bool.RES_SELECT_def"
+  "RES_SELECT_DEF" > "HOL4Base.bool.RES_SELECT_DEF"
+  "RES_FORALL_def" > "HOL4Base.bool.RES_FORALL_def"
+  "RES_FORALL_DEF" > "HOL4Base.bool.RES_FORALL_DEF"
+  "RES_EXISTS_def" > "HOL4Base.bool.RES_EXISTS_def"
+  "RES_EXISTS_UNIQUE_def" > "HOL4Base.bool.RES_EXISTS_UNIQUE_def"
+  "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"
+  "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_DEF" > "HOL.or_def"
+  "OR_CONG" > "HOL4Base.bool.OR_CONG"
+  "OR_CLAUSES" > "HOL4Base.bool.OR_CLAUSES"
+  "ONTO_THM" > "HOL4Setup.ONTO_DEF"
+  "ONTO_DEF" > "HOL4Setup.ONTO_DEF"
+  "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_CLAUSES" > "HOL4Base.bool.NOT_CLAUSES"
+  "NOT_AND" > "HOL4Base.bool.NOT_AND"
+  "MONO_OR" > "Inductive.basic_monos_3"
+  "MONO_NOT" > "HOL.rev_contrapos"
+  "MONO_IMP" > "Set.imp_mono"
+  "MONO_EXISTS" > "Inductive.basic_monos_5"
+  "MONO_COND" > "HOL4Base.bool.MONO_COND"
+  "MONO_AND" > "Hilbert_Choice.conj_forward"
+  "MONO_ALL" > "Inductive.basic_monos_6"
+  "LET_THM" > "HOL.Let_def"
+  "LET_RATOR" > "HOL4Base.bool.LET_RATOR"
+  "LET_RAND" > "HOL4Base.bool.LET_RAND"
+  "LET_DEF" > "HOL4Compat.LET_def"
+  "LET_CONG" > "Recdef.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_EXISTS_AND_THM" > "HOL.ex_simps_1"
+  "LEFT_AND_OVER_OR" > "HOL.conj_disj_distribL"
+  "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_CONG" > "HOL.imp_cong"
+  "IMP_CLAUSES" > "HOL4Base.bool.IMP_CLAUSES"
+  "IMP_ANTISYM_AX" > "HOL4Setup.light_imp_as"
+  "F_IMP" > "HOL4Base.bool.F_IMP"
+  "F_DEF" > "HOL.False_def"
+  "FUN_EQ_THM" > "Fun.expand_fun_eq"
+  "FORALL_THM" > "HOL4Base.bool.FORALL_THM"
+  "FORALL_SIMP" > "HOL.simp_thms_35"
+  "FORALL_DEF" > "HOL.All_def"
+  "FORALL_AND_THM" > "HOL.all_conj_distrib"
+  "FALSITY" > "HOL.FalseE"
+  "EXISTS_UNIQUE_THM" > "HOL4Compat.EXISTS_UNIQUE_DEF"
+  "EXISTS_UNIQUE_REFL" > "HOL.ex1_eq_1"
+  "EXISTS_UNIQUE_DEF" > "HOL4Compat.EXISTS_UNIQUE_DEF"
+  "EXISTS_THM" > "HOL4Base.bool.EXISTS_THM"
+  "EXISTS_SIMP" > "HOL.simp_thms_36"
+  "EXISTS_REFL" > "HOL.simp_thms_37"
+  "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"
+  "EQ_IMP_THM" > "HOL.iff_conv_conj_imp"
+  "EQ_EXT" > "HOL.meta_eq_to_obj_eq"
+  "EQ_EXPAND" > "HOL4Base.bool.EQ_EXPAND"
+  "EQ_CLAUSES" > "HOL4Base.bool.EQ_CLAUSES"
+  "DISJ_SYM" > "HOL.disj_comms_1"
+  "DISJ_IMP_THM" > "HOL.imp_disjL"
+  "DISJ_COMM" > "HOL.disj_comms_1"
+  "DISJ_ASSOC" > "Recdef.tfl_disj_assoc"
+  "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"
+  "COND_RATOR" > "HOL4Base.bool.COND_RATOR"
+  "COND_RAND" > "HOL.if_distrib"
+  "COND_ID" > "HOL.if_cancel"
+  "COND_EXPAND" > "HOL4Base.bool.COND_EXPAND"
+  "COND_DEF" > "HOL4Compat.COND_DEF"
+  "COND_CONG" > "HOL4Base.bool.COND_CONG"
+  "COND_CLAUSES" > "HOL4Base.bool.COND_CLAUSES"
+  "COND_ABS" > "HOL4Base.bool.COND_ABS"
+  "BOTH_FORALL_OR_THM" > "HOL4Base.bool.BOTH_FORALL_OR_THM"
+  "BOTH_FORALL_IMP_THM" > "HOL4Base.bool.BOTH_FORALL_IMP_THM"
+  "BOTH_EXISTS_IMP_THM" > "HOL4Base.bool.BOTH_EXISTS_IMP_THM"
+  "BOTH_EXISTS_AND_THM" > "HOL4Base.bool.BOTH_EXISTS_AND_THM"
+  "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" > "HOL.True_or_False"
+  "BETA_THM" > "Presburger.fm_modd_pinf"
+  "ARB_def" > "HOL4Base.bool.ARB_def"
+  "ARB_DEF" > "HOL4Base.bool.ARB_DEF"
+  "AND_INTRO_THM" > "HOL.conjI"
+  "AND_IMP_INTRO" > "HOL.imp_conjL"
+  "AND_DEF" > "HOL.and_def"
+  "AND_CONG" > "HOL4Base.bool.AND_CONG"
+  "AND_CLAUSES" > "HOL4Base.bool.AND_CLAUSES"
+  "AND2_THM" > "HOL.conjunct2"
+  "AND1_THM" > "HOL.conjunct1"
+  "ABS_SIMP" > "Presburger.fm_modd_pinf"
+  "ABS_REP_THM" > "HOL4Base.bool.ABS_REP_THM"
+
+ignore_thms
+  "UNBOUNDED_THM"
+  "UNBOUNDED_DEF"
+  "BOUNDED_THM"
+  "BOUNDED_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/boolean_sequence.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,39 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "STL" > "STL_primdef"
+  "STAKE" > "STAKE_primdef"
+  "SHD" > "SHD_primdef"
+  "SDROP" > "SDROP_primdef"
+  "SDEST" > "SDEST_primdef"
+  "SCONST" > "SCONST_primdef"
+  "SCONS" > "SCONS_primdef"
+
+const_maps
+  "STL" > "HOL4Prob.boolean_sequence.STL"
+  "SHD" > "HOL4Prob.boolean_sequence.SHD"
+  "SDEST" > "HOL4Prob.boolean_sequence.SDEST"
+  "SCONST" > "HOL4Prob.boolean_sequence.SCONST"
+
+thm_maps
+  "STL_primdef" > "HOL4Prob.boolean_sequence.STL_primdef"
+  "STL_def" > "HOL4Prob.boolean_sequence.STL_def"
+  "STL_SCONST" > "HOL4Prob.boolean_sequence.STL_SCONST"
+  "STL_SCONS" > "HOL4Prob.boolean_sequence.STL_SCONS"
+  "STAKE_def" > "HOL4Prob.boolean_sequence.STAKE_def"
+  "SHD_primdef" > "HOL4Prob.boolean_sequence.SHD_primdef"
+  "SHD_def" > "HOL4Prob.boolean_sequence.SHD_def"
+  "SHD_STL_ISO" > "HOL4Prob.boolean_sequence.SHD_STL_ISO"
+  "SHD_SCONST" > "HOL4Prob.boolean_sequence.SHD_SCONST"
+  "SHD_SCONS" > "HOL4Prob.boolean_sequence.SHD_SCONS"
+  "SDROP_def" > "HOL4Prob.boolean_sequence.SDROP_def"
+  "SDEST_primdef" > "HOL4Prob.boolean_sequence.SDEST_primdef"
+  "SDEST_def" > "HOL4Prob.boolean_sequence.SDEST_def"
+  "SCONS_def" > "HOL4Prob.boolean_sequence.SCONS_def"
+  "SCONS_SURJ" > "HOL4Prob.boolean_sequence.SCONS_SURJ"
+  "SCONST_primdef" > "HOL4Prob.boolean_sequence.SCONST_primdef"
+  "SCONST_def" > "HOL4Prob.boolean_sequence.SCONST_def"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/bword_arith.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,27 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "ICARRY" > "ICARRY_def"
+  "ACARRY" > "ACARRY_def"
+
+thm_maps
+  "WSEG_NBWORD_ADD" > "HOL4Vec.bword_arith.WSEG_NBWORD_ADD"
+  "ICARRY_WSEG" > "HOL4Vec.bword_arith.ICARRY_WSEG"
+  "ICARRY_DEF" > "HOL4Vec.bword_arith.ICARRY_DEF"
+  "BNVAL_LESS_EQ" > "HOL4Vec.bword_arith.BNVAL_LESS_EQ"
+  "ADD_WORD_SPLIT" > "HOL4Vec.bword_arith.ADD_WORD_SPLIT"
+  "ADD_NBWORD_EQ0_SPLIT" > "HOL4Vec.bword_arith.ADD_NBWORD_EQ0_SPLIT"
+  "ADD_BV_BNVAL_LESS_EQ1" > "HOL4Vec.bword_arith.ADD_BV_BNVAL_LESS_EQ1"
+  "ADD_BV_BNVAL_LESS_EQ" > "HOL4Vec.bword_arith.ADD_BV_BNVAL_LESS_EQ"
+  "ADD_BV_BNVAL_DIV_LESS_EQ1" > "HOL4Vec.bword_arith.ADD_BV_BNVAL_DIV_LESS_EQ1"
+  "ADD_BNVAL_LESS_EQ1" > "HOL4Vec.bword_arith.ADD_BNVAL_LESS_EQ1"
+  "ACARRY_WSEG" > "HOL4Vec.bword_arith.ACARRY_WSEG"
+  "ACARRY_MSB" > "HOL4Vec.bword_arith.ACARRY_MSB"
+  "ACARRY_EQ_ICARRY" > "HOL4Vec.bword_arith.ACARRY_EQ_ICARRY"
+  "ACARRY_EQ_ADD_DIV" > "HOL4Vec.bword_arith.ACARRY_EQ_ADD_DIV"
+  "ACARRY_DEF" > "HOL4Vec.bword_arith.ACARRY_DEF"
+  "ACARRY_ACARRY_WSEG" > "HOL4Vec.bword_arith.ACARRY_ACARRY_WSEG"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/bword_bitop.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,23 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "WXOR" > "WXOR_def"
+  "WOR" > "WOR_def"
+  "WNOT" > "WNOT_def"
+  "WAND" > "WAND_def"
+
+thm_maps
+  "WXOR_DEF" > "HOL4Vec.bword_bitop.WXOR_DEF"
+  "WOR_DEF" > "HOL4Vec.bword_bitop.WOR_DEF"
+  "WNOT_WNOT" > "HOL4Vec.bword_bitop.WNOT_WNOT"
+  "WNOT_DEF" > "HOL4Vec.bword_bitop.WNOT_DEF"
+  "WCAT_WNOT" > "HOL4Vec.bword_bitop.WCAT_WNOT"
+  "WAND_DEF" > "HOL4Vec.bword_bitop.WAND_DEF"
+  "PBITOP_WNOT" > "HOL4Vec.bword_bitop.PBITOP_WNOT"
+  "PBITBOP_WXOR" > "HOL4Vec.bword_bitop.PBITBOP_WXOR"
+  "PBITBOP_WOR" > "HOL4Vec.bword_bitop.PBITBOP_WOR"
+  "PBITBOP_WAND" > "HOL4Vec.bword_bitop.PBITBOP_WAND"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/bword_num.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,58 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "VB" > "VB_def"
+  "NBWORD" > "NBWORD_def"
+  "BV" > "BV_def"
+  "BNVAL" > "BNVAL_def"
+
+const_maps
+  "VB" > "HOL4Vec.bword_num.VB"
+  "NBWORD" > "HOL4Vec.bword_num.NBWORD"
+  "BV" > "HOL4Vec.bword_num.BV"
+
+thm_maps
+  "ZERO_WORD_VAL" > "HOL4Vec.bword_num.ZERO_WORD_VAL"
+  "WSPLIT_NBWORD_0" > "HOL4Vec.bword_num.WSPLIT_NBWORD_0"
+  "WSEG_NBWORD_SUC" > "HOL4Vec.bword_num.WSEG_NBWORD_SUC"
+  "WSEG_NBWORD" > "HOL4Vec.bword_num.WSEG_NBWORD"
+  "WORDLEN_NBWORD" > "HOL4Vec.bword_num.WORDLEN_NBWORD"
+  "WCAT_NBWORD_0" > "HOL4Vec.bword_num.WCAT_NBWORD_0"
+  "VB_def" > "HOL4Vec.bword_num.VB_def"
+  "VB_DEF" > "HOL4Vec.bword_num.VB_DEF"
+  "VB_BV" > "HOL4Vec.bword_num.VB_BV"
+  "PWORDLEN_NBWORD" > "HOL4Vec.bword_num.PWORDLEN_NBWORD"
+  "NBWORD_def" > "HOL4Vec.bword_num.NBWORD_def"
+  "NBWORD_SUC_WSEG" > "HOL4Vec.bword_num.NBWORD_SUC_WSEG"
+  "NBWORD_SUC_FST" > "HOL4Vec.bword_num.NBWORD_SUC_FST"
+  "NBWORD_SUC" > "HOL4Vec.bword_num.NBWORD_SUC"
+  "NBWORD_SPLIT" > "HOL4Vec.bword_num.NBWORD_SPLIT"
+  "NBWORD_MOD" > "HOL4Vec.bword_num.NBWORD_MOD"
+  "NBWORD_DEF" > "HOL4Vec.bword_num.NBWORD_DEF"
+  "NBWORD_BNVAL" > "HOL4Vec.bword_num.NBWORD_BNVAL"
+  "NBWORD0" > "HOL4Vec.bword_num.NBWORD0"
+  "MSB_NBWORD" > "HOL4Vec.bword_num.MSB_NBWORD"
+  "EQ_NBWORD0_SPLIT" > "HOL4Vec.bword_num.EQ_NBWORD0_SPLIT"
+  "DOUBL_EQ_SHL" > "HOL4Vec.bword_num.DOUBL_EQ_SHL"
+  "BV_def" > "HOL4Vec.bword_num.BV_def"
+  "BV_VB" > "HOL4Vec.bword_num.BV_VB"
+  "BV_LESS_2" > "HOL4Vec.bword_num.BV_LESS_2"
+  "BV_DEF" > "HOL4Vec.bword_num.BV_DEF"
+  "BNVAL_WCAT2" > "HOL4Vec.bword_num.BNVAL_WCAT2"
+  "BNVAL_WCAT1" > "HOL4Vec.bword_num.BNVAL_WCAT1"
+  "BNVAL_WCAT" > "HOL4Vec.bword_num.BNVAL_WCAT"
+  "BNVAL_ONTO" > "HOL4Vec.bword_num.BNVAL_ONTO"
+  "BNVAL_NVAL" > "HOL4Vec.bword_num.BNVAL_NVAL"
+  "BNVAL_NBWORD" > "HOL4Vec.bword_num.BNVAL_NBWORD"
+  "BNVAL_MAX" > "HOL4Vec.bword_num.BNVAL_MAX"
+  "BNVAL_DEF" > "HOL4Vec.bword_num.BNVAL_DEF"
+  "BNVAL_11" > "HOL4Vec.bword_num.BNVAL_11"
+  "BNVAL0" > "HOL4Vec.bword_num.BNVAL0"
+  "BIT_NBWORD0" > "HOL4Vec.bword_num.BIT_NBWORD0"
+  "ADD_BNVAL_SPLIT" > "HOL4Vec.bword_num.ADD_BNVAL_SPLIT"
+  "ADD_BNVAL_RIGHT" > "HOL4Vec.bword_num.ADD_BNVAL_RIGHT"
+  "ADD_BNVAL_LEFT" > "HOL4Vec.bword_num.ADD_BNVAL_LEFT"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/combin.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,41 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "W" > "W_def"
+  "S" > "S_def"
+  "K" > "K_def"
+  "I" > "I_def"
+  "C" > "C_def"
+
+const_maps
+  "o" > "Fun.comp"
+  "W" > "HOL4Base.combin.W"
+  "S" > "HOL4Base.combin.S"
+  "K" > "HOL4Base.combin.K"
+  "I" > "HOL4Base.combin.I"
+  "C" > "HOL4Base.combin.C"
+
+thm_maps
+  "o_THM" > "Fun.o_apply"
+  "o_DEF" > "Fun.o_apply"
+  "o_ASSOC" > "Fun.o_assoc"
+  "W_def" > "HOL4Base.combin.W_def"
+  "W_THM" > "HOL4Base.combin.W_def"
+  "W_DEF" > "HOL4Base.combin.W_DEF"
+  "S_def" > "HOL4Base.combin.S_def"
+  "S_THM" > "HOL4Base.combin.S_def"
+  "S_DEF" > "HOL4Base.combin.S_DEF"
+  "K_def" > "HOL4Base.combin.K_def"
+  "K_THM" > "HOL4Base.combin.K_def"
+  "K_DEF" > "HOL4Base.combin.K_DEF"
+  "I_o_ID" > "HOL4Base.combin.I_o_ID"
+  "I_def" > "HOL4Base.combin.I_def"
+  "I_THM" > "HOL4Base.combin.I_THM"
+  "I_DEF" > "HOL4Base.combin.I_DEF"
+  "C_def" > "HOL4Base.combin.C_def"
+  "C_THM" > "HOL4Base.combin.C_def"
+  "C_DEF" > "HOL4Base.combin.C_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/divides.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,24 @@
+import
+
+import_segment "hol4"
+
+const_maps
+  "divides" > "Divides.op dvd" :: "nat => nat => bool"
+
+thm_maps
+  "divides_def" > "HOL4Compat.divides_def"
+  "ONE_DIVIDES_ALL" > "HOL4Base.divides.ONE_DIVIDES_ALL"
+  "NOT_LT_DIV" > "HOL4Base.divides.NOT_LT_DIV"
+  "DIVIDES_TRANS" > "Divides.dvd_trans"
+  "DIVIDES_SUB" > "Divides.dvd_diff"
+  "DIVIDES_REFL" > "Divides.dvd_refl"
+  "DIVIDES_MULT_LEFT" > "HOL4Base.divides.DIVIDES_MULT_LEFT"
+  "DIVIDES_MULT" > "Divides.dvd_mult2"
+  "DIVIDES_LE" > "Divides.dvd_imp_le"
+  "DIVIDES_FACT" > "HOL4Base.divides.DIVIDES_FACT"
+  "DIVIDES_ANTISYM" > "Divides.dvd_anti_sym"
+  "DIVIDES_ADD_2" > "HOL4Base.divides.DIVIDES_ADD_2"
+  "DIVIDES_ADD_1" > "Divides.dvd_add"
+  "ALL_DIVIDES_0" > "Divides.dvd_0_right"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/hrat.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,95 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "trat_sucint" > "trat_sucint_def"
+  "trat_mul" > "trat_mul_def"
+  "trat_inv" > "trat_inv_def"
+  "trat_eq" > "trat_eq_def"
+  "trat_add" > "trat_add_def"
+  "trat_1" > "trat_1_def"
+  "mk_hrat" > "mk_hrat_def"
+  "hrat_sucint" > "hrat_sucint_def"
+  "hrat_mul" > "hrat_mul_def"
+  "hrat_inv" > "hrat_inv_def"
+  "hrat_add" > "hrat_add_def"
+  "hrat_1" > "hrat_1_def"
+  "dest_hrat" > "dest_hrat_def"
+
+type_maps
+  "hrat" > "HOL4Base.hrat.hrat"
+
+const_maps
+  "trat_mul" > "HOL4Base.hrat.trat_mul"
+  "trat_inv" > "HOL4Base.hrat.trat_inv"
+  "trat_eq" > "HOL4Base.hrat.trat_eq"
+  "trat_add" > "HOL4Base.hrat.trat_add"
+  "trat_1" > "HOL4Base.hrat.trat_1"
+  "hrat_sucint" > "HOL4Base.hrat.hrat_sucint"
+  "hrat_mul" > "HOL4Base.hrat.hrat_mul"
+  "hrat_inv" > "HOL4Base.hrat.hrat_inv"
+  "hrat_add" > "HOL4Base.hrat.hrat_add"
+  "hrat_1" > "HOL4Base.hrat.hrat_1"
+
+thm_maps
+  "trat_sucint" > "HOL4Base.hrat.trat_sucint"
+  "trat_mul_def" > "HOL4Base.hrat.trat_mul_def"
+  "trat_mul" > "HOL4Base.hrat.trat_mul"
+  "trat_inv_def" > "HOL4Base.hrat.trat_inv_def"
+  "trat_inv" > "HOL4Base.hrat.trat_inv"
+  "trat_eq_def" > "HOL4Base.hrat.trat_eq_def"
+  "trat_eq" > "HOL4Base.hrat.trat_eq"
+  "trat_add_def" > "HOL4Base.hrat.trat_add_def"
+  "trat_add" > "HOL4Base.hrat.trat_add"
+  "trat_1_def" > "HOL4Base.hrat.trat_1_def"
+  "trat_1" > "HOL4Base.hrat.trat_1"
+  "hrat_tybij" > "HOL4Base.hrat.hrat_tybij"
+  "hrat_sucint_def" > "HOL4Base.hrat.hrat_sucint_def"
+  "hrat_sucint" > "HOL4Base.hrat.hrat_sucint"
+  "hrat_mul_def" > "HOL4Base.hrat.hrat_mul_def"
+  "hrat_mul" > "HOL4Base.hrat.hrat_mul"
+  "hrat_inv_def" > "HOL4Base.hrat.hrat_inv_def"
+  "hrat_inv" > "HOL4Base.hrat.hrat_inv"
+  "hrat_add_def" > "HOL4Base.hrat.hrat_add_def"
+  "hrat_add" > "HOL4Base.hrat.hrat_add"
+  "hrat_TY_DEF" > "HOL4Base.hrat.hrat_TY_DEF"
+  "hrat_1_def" > "HOL4Base.hrat.hrat_1_def"
+  "hrat_1" > "HOL4Base.hrat.hrat_1"
+  "TRAT_SUCINT_0" > "HOL4Base.hrat.TRAT_SUCINT_0"
+  "TRAT_SUCINT" > "HOL4Base.hrat.TRAT_SUCINT"
+  "TRAT_NOZERO" > "HOL4Base.hrat.TRAT_NOZERO"
+  "TRAT_MUL_WELLDEFINED2" > "HOL4Base.hrat.TRAT_MUL_WELLDEFINED2"
+  "TRAT_MUL_WELLDEFINED" > "HOL4Base.hrat.TRAT_MUL_WELLDEFINED"
+  "TRAT_MUL_SYM_EQ" > "HOL4Base.hrat.TRAT_MUL_SYM_EQ"
+  "TRAT_MUL_SYM" > "HOL4Base.hrat.TRAT_MUL_SYM"
+  "TRAT_MUL_LINV" > "HOL4Base.hrat.TRAT_MUL_LINV"
+  "TRAT_MUL_LID" > "HOL4Base.hrat.TRAT_MUL_LID"
+  "TRAT_MUL_ASSOC" > "HOL4Base.hrat.TRAT_MUL_ASSOC"
+  "TRAT_LDISTRIB" > "HOL4Base.hrat.TRAT_LDISTRIB"
+  "TRAT_INV_WELLDEFINED" > "HOL4Base.hrat.TRAT_INV_WELLDEFINED"
+  "TRAT_EQ_TRANS" > "HOL4Base.hrat.TRAT_EQ_TRANS"
+  "TRAT_EQ_SYM" > "HOL4Base.hrat.TRAT_EQ_SYM"
+  "TRAT_EQ_REFL" > "HOL4Base.hrat.TRAT_EQ_REFL"
+  "TRAT_EQ_EQUIV" > "HOL4Base.hrat.TRAT_EQ_EQUIV"
+  "TRAT_EQ_AP" > "HOL4Base.hrat.TRAT_EQ_AP"
+  "TRAT_ARCH" > "HOL4Base.hrat.TRAT_ARCH"
+  "TRAT_ADD_WELLDEFINED2" > "HOL4Base.hrat.TRAT_ADD_WELLDEFINED2"
+  "TRAT_ADD_WELLDEFINED" > "HOL4Base.hrat.TRAT_ADD_WELLDEFINED"
+  "TRAT_ADD_TOTAL" > "HOL4Base.hrat.TRAT_ADD_TOTAL"
+  "TRAT_ADD_SYM_EQ" > "HOL4Base.hrat.TRAT_ADD_SYM_EQ"
+  "TRAT_ADD_SYM" > "HOL4Base.hrat.TRAT_ADD_SYM"
+  "TRAT_ADD_ASSOC" > "HOL4Base.hrat.TRAT_ADD_ASSOC"
+  "HRAT_SUCINT" > "HOL4Base.hrat.HRAT_SUCINT"
+  "HRAT_NOZERO" > "HOL4Base.hrat.HRAT_NOZERO"
+  "HRAT_MUL_SYM" > "HOL4Base.hrat.HRAT_MUL_SYM"
+  "HRAT_MUL_LINV" > "HOL4Base.hrat.HRAT_MUL_LINV"
+  "HRAT_MUL_LID" > "HOL4Base.hrat.HRAT_MUL_LID"
+  "HRAT_MUL_ASSOC" > "HOL4Base.hrat.HRAT_MUL_ASSOC"
+  "HRAT_LDISTRIB" > "HOL4Base.hrat.HRAT_LDISTRIB"
+  "HRAT_ARCH" > "HOL4Base.hrat.HRAT_ARCH"
+  "HRAT_ADD_TOTAL" > "HOL4Base.hrat.HRAT_ADD_TOTAL"
+  "HRAT_ADD_SYM" > "HOL4Base.hrat.HRAT_ADD_SYM"
+  "HRAT_ADD_ASSOC" > "HOL4Base.hrat.HRAT_ADD_ASSOC"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/hreal.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,118 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "isacut" > "isacut_def"
+  "hreal_sup" > "hreal_sup_def"
+  "hreal_sub" > "hreal_sub_def"
+  "hreal_mul" > "hreal_mul_def"
+  "hreal_lt" > "hreal_lt_def"
+  "hreal_inv" > "hreal_inv_def"
+  "hreal_add" > "hreal_add_def"
+  "hreal_1" > "hreal_1_def"
+  "hreal" > "hreal_def"
+  "hrat_lt" > "hrat_lt_def"
+  "cut_of_hrat" > "cut_of_hrat_def"
+  "cut" > "cut_def"
+
+type_maps
+  "hreal" > "HOL4Base.hreal.hreal"
+
+const_maps
+  "isacut" > "HOL4Base.hreal.isacut"
+  "hreal_sup" > "HOL4Base.hreal.hreal_sup"
+  "hreal_sub" > "HOL4Base.hreal.hreal_sub"
+  "hreal_mul" > "HOL4Base.hreal.hreal_mul"
+  "hreal_lt" > "HOL4Base.hreal.hreal_lt"
+  "hreal_inv" > "HOL4Base.hreal.hreal_inv"
+  "hreal_add" > "HOL4Base.hreal.hreal_add"
+  "hreal_1" > "HOL4Base.hreal.hreal_1"
+  "hrat_lt" > "HOL4Base.hreal.hrat_lt"
+  "cut_of_hrat" > "HOL4Base.hreal.cut_of_hrat"
+
+thm_maps
+  "isacut_def" > "HOL4Base.hreal.isacut_def"
+  "isacut" > "HOL4Base.hreal.isacut"
+  "hreal_tybij" > "HOL4Base.hreal.hreal_tybij"
+  "hreal_sup_def" > "HOL4Base.hreal.hreal_sup_def"
+  "hreal_sup" > "HOL4Base.hreal.hreal_sup"
+  "hreal_sub_def" > "HOL4Base.hreal.hreal_sub_def"
+  "hreal_sub" > "HOL4Base.hreal.hreal_sub"
+  "hreal_mul_def" > "HOL4Base.hreal.hreal_mul_def"
+  "hreal_mul" > "HOL4Base.hreal.hreal_mul"
+  "hreal_lt_def" > "HOL4Base.hreal.hreal_lt_def"
+  "hreal_lt" > "HOL4Base.hreal.hreal_lt"
+  "hreal_inv_def" > "HOL4Base.hreal.hreal_inv_def"
+  "hreal_inv" > "HOL4Base.hreal.hreal_inv"
+  "hreal_add_def" > "HOL4Base.hreal.hreal_add_def"
+  "hreal_add" > "HOL4Base.hreal.hreal_add"
+  "hreal_TY_DEF" > "HOL4Base.hreal.hreal_TY_DEF"
+  "hreal_1_def" > "HOL4Base.hreal.hreal_1_def"
+  "hreal_1" > "HOL4Base.hreal.hreal_1"
+  "hrat_lt_def" > "HOL4Base.hreal.hrat_lt_def"
+  "hrat_lt" > "HOL4Base.hreal.hrat_lt"
+  "cut_of_hrat_def" > "HOL4Base.hreal.cut_of_hrat_def"
+  "cut_of_hrat" > "HOL4Base.hreal.cut_of_hrat"
+  "ISACUT_HRAT" > "HOL4Base.hreal.ISACUT_HRAT"
+  "HREAL_SUP_ISACUT" > "HOL4Base.hreal.HREAL_SUP_ISACUT"
+  "HREAL_SUP" > "HOL4Base.hreal.HREAL_SUP"
+  "HREAL_SUB_ISACUT" > "HOL4Base.hreal.HREAL_SUB_ISACUT"
+  "HREAL_SUB_ADD" > "HOL4Base.hreal.HREAL_SUB_ADD"
+  "HREAL_NOZERO" > "HOL4Base.hreal.HREAL_NOZERO"
+  "HREAL_MUL_SYM" > "HOL4Base.hreal.HREAL_MUL_SYM"
+  "HREAL_MUL_LINV" > "HOL4Base.hreal.HREAL_MUL_LINV"
+  "HREAL_MUL_LID" > "HOL4Base.hreal.HREAL_MUL_LID"
+  "HREAL_MUL_ISACUT" > "HOL4Base.hreal.HREAL_MUL_ISACUT"
+  "HREAL_MUL_ASSOC" > "HOL4Base.hreal.HREAL_MUL_ASSOC"
+  "HREAL_LT_TOTAL" > "HOL4Base.hreal.HREAL_LT_TOTAL"
+  "HREAL_LT_LEMMA" > "HOL4Base.hreal.HREAL_LT_LEMMA"
+  "HREAL_LT" > "HOL4Base.hreal.HREAL_LT"
+  "HREAL_LDISTRIB" > "HOL4Base.hreal.HREAL_LDISTRIB"
+  "HREAL_INV_ISACUT" > "HOL4Base.hreal.HREAL_INV_ISACUT"
+  "HREAL_ADD_TOTAL" > "HOL4Base.hreal.HREAL_ADD_TOTAL"
+  "HREAL_ADD_SYM" > "HOL4Base.hreal.HREAL_ADD_SYM"
+  "HREAL_ADD_ISACUT" > "HOL4Base.hreal.HREAL_ADD_ISACUT"
+  "HREAL_ADD_ASSOC" > "HOL4Base.hreal.HREAL_ADD_ASSOC"
+  "HRAT_UP" > "HOL4Base.hreal.HRAT_UP"
+  "HRAT_RDISTRIB" > "HOL4Base.hreal.HRAT_RDISTRIB"
+  "HRAT_MUL_RINV" > "HOL4Base.hreal.HRAT_MUL_RINV"
+  "HRAT_MUL_RID" > "HOL4Base.hreal.HRAT_MUL_RID"
+  "HRAT_MEAN" > "HOL4Base.hreal.HRAT_MEAN"
+  "HRAT_LT_TRANS" > "HOL4Base.hreal.HRAT_LT_TRANS"
+  "HRAT_LT_TOTAL" > "HOL4Base.hreal.HRAT_LT_TOTAL"
+  "HRAT_LT_RMUL1" > "HOL4Base.hreal.HRAT_LT_RMUL1"
+  "HRAT_LT_RMUL" > "HOL4Base.hreal.HRAT_LT_RMUL"
+  "HRAT_LT_REFL" > "HOL4Base.hreal.HRAT_LT_REFL"
+  "HRAT_LT_RADD" > "HOL4Base.hreal.HRAT_LT_RADD"
+  "HRAT_LT_R1" > "HOL4Base.hreal.HRAT_LT_R1"
+  "HRAT_LT_NE" > "HOL4Base.hreal.HRAT_LT_NE"
+  "HRAT_LT_MUL2" > "HOL4Base.hreal.HRAT_LT_MUL2"
+  "HRAT_LT_LMUL1" > "HOL4Base.hreal.HRAT_LT_LMUL1"
+  "HRAT_LT_LMUL" > "HOL4Base.hreal.HRAT_LT_LMUL"
+  "HRAT_LT_LADD" > "HOL4Base.hreal.HRAT_LT_LADD"
+  "HRAT_LT_L1" > "HOL4Base.hreal.HRAT_LT_L1"
+  "HRAT_LT_GT" > "HOL4Base.hreal.HRAT_LT_GT"
+  "HRAT_LT_ANTISYM" > "HOL4Base.hreal.HRAT_LT_ANTISYM"
+  "HRAT_LT_ADDR" > "HOL4Base.hreal.HRAT_LT_ADDR"
+  "HRAT_LT_ADDL" > "HOL4Base.hreal.HRAT_LT_ADDL"
+  "HRAT_LT_ADD2" > "HOL4Base.hreal.HRAT_LT_ADD2"
+  "HRAT_INV_MUL" > "HOL4Base.hreal.HRAT_INV_MUL"
+  "HRAT_GT_LMUL1" > "HOL4Base.hreal.HRAT_GT_LMUL1"
+  "HRAT_GT_L1" > "HOL4Base.hreal.HRAT_GT_L1"
+  "HRAT_EQ_LMUL" > "HOL4Base.hreal.HRAT_EQ_LMUL"
+  "HRAT_EQ_LADD" > "HOL4Base.hreal.HRAT_EQ_LADD"
+  "HRAT_DOWN2" > "HOL4Base.hreal.HRAT_DOWN2"
+  "HRAT_DOWN" > "HOL4Base.hreal.HRAT_DOWN"
+  "EQUAL_CUTS" > "HOL4Base.hreal.EQUAL_CUTS"
+  "CUT_UP" > "HOL4Base.hreal.CUT_UP"
+  "CUT_UBOUND" > "HOL4Base.hreal.CUT_UBOUND"
+  "CUT_STRADDLE" > "HOL4Base.hreal.CUT_STRADDLE"
+  "CUT_NONEMPTY" > "HOL4Base.hreal.CUT_NONEMPTY"
+  "CUT_NEARTOP_MUL" > "HOL4Base.hreal.CUT_NEARTOP_MUL"
+  "CUT_NEARTOP_ADD" > "HOL4Base.hreal.CUT_NEARTOP_ADD"
+  "CUT_ISACUT" > "HOL4Base.hreal.CUT_ISACUT"
+  "CUT_DOWN" > "HOL4Base.hreal.CUT_DOWN"
+  "CUT_BOUNDED" > "HOL4Base.hreal.CUT_BOUNDED"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/ind_type.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,99 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "mk_rec" > "mk_rec_def"
+  "dest_rec" > "dest_rec_def"
+  "ZRECSPACE" > "ZRECSPACE_def"
+  "ZCONSTR" > "ZCONSTR_def"
+  "ZBOT" > "ZBOT_def"
+  "NUMSUM" > "NUMSUM_def"
+  "NUMSND" > "NUMSND_def"
+  "NUMRIGHT" > "NUMRIGHT_def"
+  "NUMPAIR" > "NUMPAIR_def"
+  "NUMLEFT" > "NUMLEFT_def"
+  "NUMFST" > "NUMFST_def"
+  "ISO" > "ISO_def"
+  "INJP" > "INJP_def"
+  "INJN" > "INJN_def"
+  "INJF" > "INJF_def"
+  "INJA" > "INJA_def"
+  "FNIL" > "FNIL_def"
+  "FCONS" > "FCONS_def"
+  "CONSTR" > "CONSTR_def"
+  "BOTTOM" > "BOTTOM_def"
+
+type_maps
+  "recspace" > "HOL4Base.ind_type.recspace"
+
+const_maps
+  "ZRECSPACE" > "HOL4Base.ind_type.ZRECSPACE"
+  "ZCONSTR" > "HOL4Base.ind_type.ZCONSTR"
+  "ZBOT" > "HOL4Base.ind_type.ZBOT"
+  "NUMSUM" > "HOL4Base.ind_type.NUMSUM"
+  "NUMPAIR" > "HOL4Base.ind_type.NUMPAIR"
+  "ISO" > "HOL4Base.ind_type.ISO"
+  "INJP" > "HOL4Base.ind_type.INJP"
+  "INJN" > "HOL4Base.ind_type.INJN"
+  "INJF" > "HOL4Base.ind_type.INJF"
+  "INJA" > "HOL4Base.ind_type.INJA"
+  "FNIL" > "HOL4Base.ind_type.FNIL"
+  "CONSTR" > "HOL4Base.ind_type.CONSTR"
+  "BOTTOM" > "HOL4Base.ind_type.BOTTOM"
+
+thm_maps
+  "recspace_repfns" > "HOL4Base.ind_type.recspace_repfns"
+  "recspace_TY_DEF" > "HOL4Base.ind_type.recspace_TY_DEF"
+  "ZRECSPACE_rules" > "HOL4Base.ind_type.ZRECSPACE_rules"
+  "ZRECSPACE_ind" > "HOL4Base.ind_type.ZRECSPACE_ind"
+  "ZRECSPACE_def" > "HOL4Base.ind_type.ZRECSPACE_def"
+  "ZRECSPACE_cases" > "HOL4Base.ind_type.ZRECSPACE_cases"
+  "ZRECSPACE" > "HOL4Base.ind_type.ZRECSPACE"
+  "ZCONSTR_def" > "HOL4Base.ind_type.ZCONSTR_def"
+  "ZCONSTR_ZBOT" > "HOL4Base.ind_type.ZCONSTR_ZBOT"
+  "ZCONSTR" > "HOL4Base.ind_type.ZCONSTR"
+  "ZBOT_def" > "HOL4Base.ind_type.ZBOT_def"
+  "ZBOT" > "HOL4Base.ind_type.ZBOT"
+  "NUMSUM_def" > "HOL4Base.ind_type.NUMSUM_def"
+  "NUMSUM_INJ" > "HOL4Base.ind_type.NUMSUM_INJ"
+  "NUMSUM_DEST" > "HOL4Base.ind_type.NUMSUM_DEST"
+  "NUMSUM" > "HOL4Base.ind_type.NUMSUM"
+  "NUMPAIR_def" > "HOL4Base.ind_type.NUMPAIR_def"
+  "NUMPAIR_INJ_LEMMA" > "HOL4Base.ind_type.NUMPAIR_INJ_LEMMA"
+  "NUMPAIR_INJ" > "HOL4Base.ind_type.NUMPAIR_INJ"
+  "NUMPAIR_DEST" > "HOL4Base.ind_type.NUMPAIR_DEST"
+  "NUMPAIR" > "HOL4Base.ind_type.NUMPAIR"
+  "MK_REC_INJ" > "HOL4Base.ind_type.MK_REC_INJ"
+  "ISO_def" > "HOL4Base.ind_type.ISO_def"
+  "ISO_USAGE" > "HOL4Base.ind_type.ISO_USAGE"
+  "ISO_REFL" > "HOL4Base.ind_type.ISO_REFL"
+  "ISO_FUN" > "HOL4Base.ind_type.ISO_FUN"
+  "ISO" > "HOL4Base.ind_type.ISO"
+  "INJ_INVERSE2" > "HOL4Base.ind_type.INJ_INVERSE2"
+  "INJP_def" > "HOL4Base.ind_type.INJP_def"
+  "INJP_INJ" > "HOL4Base.ind_type.INJP_INJ"
+  "INJP" > "HOL4Base.ind_type.INJP"
+  "INJN_def" > "HOL4Base.ind_type.INJN_def"
+  "INJN_INJ" > "HOL4Base.ind_type.INJN_INJ"
+  "INJN" > "HOL4Base.ind_type.INJN"
+  "INJF_def" > "HOL4Base.ind_type.INJF_def"
+  "INJF_INJ" > "HOL4Base.ind_type.INJF_INJ"
+  "INJF" > "HOL4Base.ind_type.INJF"
+  "INJA_def" > "HOL4Base.ind_type.INJA_def"
+  "INJA_INJ" > "HOL4Base.ind_type.INJA_INJ"
+  "INJA" > "HOL4Base.ind_type.INJA"
+  "FNIL_def" > "HOL4Base.ind_type.FNIL_def"
+  "FNIL" > "HOL4Base.ind_type.FNIL"
+  "FCONS" > "HOL4Base.ind_type.FCONS"
+  "DEST_REC_INJ" > "HOL4Base.ind_type.DEST_REC_INJ"
+  "CONSTR_def" > "HOL4Base.ind_type.CONSTR_def"
+  "CONSTR_REC" > "HOL4Base.ind_type.CONSTR_REC"
+  "CONSTR_INJ" > "HOL4Base.ind_type.CONSTR_INJ"
+  "CONSTR_IND" > "HOL4Base.ind_type.CONSTR_IND"
+  "CONSTR_BOT" > "HOL4Base.ind_type.CONSTR_BOT"
+  "CONSTR" > "HOL4Base.ind_type.CONSTR"
+  "BOTTOM_def" > "HOL4Base.ind_type.BOTTOM_def"
+  "BOTTOM" > "HOL4Base.ind_type.BOTTOM"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/lim.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,93 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "tends_real_real" > "tends_real_real_def"
+  "diffl" > "diffl_def"
+  "differentiable" > "differentiable_def"
+  "contl" > "contl_def"
+
+const_maps
+  "tends_real_real" > "HOL4Real.lim.tends_real_real"
+  "diffl" > "HOL4Real.lim.diffl"
+  "differentiable" > "HOL4Real.lim.differentiable"
+  "contl" > "HOL4Real.lim.contl"
+
+thm_maps
+  "tends_real_real_def" > "HOL4Real.lim.tends_real_real_def"
+  "tends_real_real" > "HOL4Real.lim.tends_real_real"
+  "diffl_def" > "HOL4Real.lim.diffl_def"
+  "diffl" > "HOL4Real.lim.diffl"
+  "differentiable_def" > "HOL4Real.lim.differentiable_def"
+  "differentiable" > "HOL4Real.lim.differentiable"
+  "contl_def" > "HOL4Real.lim.contl_def"
+  "contl" > "HOL4Real.lim.contl"
+  "ROLLE" > "HOL4Real.lim.ROLLE"
+  "MVT_LEMMA" > "HOL4Real.lim.MVT_LEMMA"
+  "MVT" > "HOL4Real.lim.MVT"
+  "LIM_X" > "HOL4Real.lim.LIM_X"
+  "LIM_UNIQ" > "HOL4Real.lim.LIM_UNIQ"
+  "LIM_TRANSFORM" > "HOL4Real.lim.LIM_TRANSFORM"
+  "LIM_SUB" > "HOL4Real.lim.LIM_SUB"
+  "LIM_NULL" > "HOL4Real.lim.LIM_NULL"
+  "LIM_NEG" > "HOL4Real.lim.LIM_NEG"
+  "LIM_MUL" > "HOL4Real.lim.LIM_MUL"
+  "LIM_INV" > "HOL4Real.lim.LIM_INV"
+  "LIM_EQUAL" > "HOL4Real.lim.LIM_EQUAL"
+  "LIM_DIV" > "HOL4Real.lim.LIM_DIV"
+  "LIM_CONST" > "HOL4Real.lim.LIM_CONST"
+  "LIM_ADD" > "HOL4Real.lim.LIM_ADD"
+  "LIM" > "HOL4Real.lim.LIM"
+  "IVT2" > "HOL4Real.lim.IVT2"
+  "IVT" > "HOL4Real.lim.IVT"
+  "INTERVAL_LEMMA" > "HOL4Real.lim.INTERVAL_LEMMA"
+  "INTERVAL_CLEMMA" > "HOL4Real.lim.INTERVAL_CLEMMA"
+  "INTERVAL_ABS" > "HOL4Real.lim.INTERVAL_ABS"
+  "DIFF_XM1" > "HOL4Real.lim.DIFF_XM1"
+  "DIFF_X" > "HOL4Real.lim.DIFF_X"
+  "DIFF_UNIQ" > "HOL4Real.lim.DIFF_UNIQ"
+  "DIFF_SUM" > "HOL4Real.lim.DIFF_SUM"
+  "DIFF_SUB" > "HOL4Real.lim.DIFF_SUB"
+  "DIFF_POW" > "HOL4Real.lim.DIFF_POW"
+  "DIFF_NEG" > "HOL4Real.lim.DIFF_NEG"
+  "DIFF_MUL" > "HOL4Real.lim.DIFF_MUL"
+  "DIFF_LMIN" > "HOL4Real.lim.DIFF_LMIN"
+  "DIFF_LMAX" > "HOL4Real.lim.DIFF_LMAX"
+  "DIFF_LINC" > "HOL4Real.lim.DIFF_LINC"
+  "DIFF_LDEC" > "HOL4Real.lim.DIFF_LDEC"
+  "DIFF_LCONST" > "HOL4Real.lim.DIFF_LCONST"
+  "DIFF_ISCONST_END" > "HOL4Real.lim.DIFF_ISCONST_END"
+  "DIFF_ISCONST_ALL" > "HOL4Real.lim.DIFF_ISCONST_ALL"
+  "DIFF_ISCONST" > "HOL4Real.lim.DIFF_ISCONST"
+  "DIFF_INVERSE_OPEN" > "HOL4Real.lim.DIFF_INVERSE_OPEN"
+  "DIFF_INVERSE_LT" > "HOL4Real.lim.DIFF_INVERSE_LT"
+  "DIFF_INVERSE" > "HOL4Real.lim.DIFF_INVERSE"
+  "DIFF_INV" > "HOL4Real.lim.DIFF_INV"
+  "DIFF_DIV" > "HOL4Real.lim.DIFF_DIV"
+  "DIFF_CONT" > "HOL4Real.lim.DIFF_CONT"
+  "DIFF_CONST" > "HOL4Real.lim.DIFF_CONST"
+  "DIFF_CMUL" > "HOL4Real.lim.DIFF_CMUL"
+  "DIFF_CHAIN" > "HOL4Real.lim.DIFF_CHAIN"
+  "DIFF_CARAT" > "HOL4Real.lim.DIFF_CARAT"
+  "DIFF_ADD" > "HOL4Real.lim.DIFF_ADD"
+  "CONT_SUB" > "HOL4Real.lim.CONT_SUB"
+  "CONT_NEG" > "HOL4Real.lim.CONT_NEG"
+  "CONT_MUL" > "HOL4Real.lim.CONT_MUL"
+  "CONT_INVERSE" > "HOL4Real.lim.CONT_INVERSE"
+  "CONT_INV" > "HOL4Real.lim.CONT_INV"
+  "CONT_INJ_RANGE" > "HOL4Real.lim.CONT_INJ_RANGE"
+  "CONT_INJ_LEMMA2" > "HOL4Real.lim.CONT_INJ_LEMMA2"
+  "CONT_INJ_LEMMA" > "HOL4Real.lim.CONT_INJ_LEMMA"
+  "CONT_HASSUP" > "HOL4Real.lim.CONT_HASSUP"
+  "CONT_DIV" > "HOL4Real.lim.CONT_DIV"
+  "CONT_CONST" > "HOL4Real.lim.CONT_CONST"
+  "CONT_COMPOSE" > "HOL4Real.lim.CONT_COMPOSE"
+  "CONT_BOUNDED" > "HOL4Real.lim.CONT_BOUNDED"
+  "CONT_ATTAINS_ALL" > "HOL4Real.lim.CONT_ATTAINS_ALL"
+  "CONT_ATTAINS2" > "HOL4Real.lim.CONT_ATTAINS2"
+  "CONT_ATTAINS" > "HOL4Real.lim.CONT_ATTAINS"
+  "CONT_ADD" > "HOL4Real.lim.CONT_ADD"
+  "CONTL_LIM" > "HOL4Real.lim.CONTL_LIM"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/list.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,135 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "EL" > "EL_def"
+
+type_maps
+  "list" > "List.list"
+
+const_maps
+  "list_size" > "HOL4Compat.list_size"
+  "list_case" > "List.list.list_case"
+  "ZIP" > "HOL4Compat.ZIP"
+  "UNZIP" > "HOL4Compat.unzip"
+  "TL" > "List.tl"
+  "SUM" > "HOL4Compat.sum"
+  "REVERSE" > "List.rev"
+  "REPLICATE" > "List.replicate"
+  "NULL" > "List.null"
+  "NIL" > "List.list.Nil"
+  "MEM" > "List.op mem"
+  "MAP2" > "HOL4Compat.map2"
+  "MAP" > "List.map"
+  "LENGTH" > "Nat.size"
+  "LAST" > "List.last"
+  "HD" > "List.hd"
+  "FRONT" > "List.butlast"
+  "FOLDR" > "HOL4Compat.FOLDR"
+  "FOLDL" > "List.foldl"
+  "FLAT" > "List.concat"
+  "FILTER" > "List.filter"
+  "EXISTS" > "HOL4Compat.list_exists"
+  "EVERY" > "List.list_all"
+  "CONS" > "List.list.Cons"
+  "APPEND" > "List.op @"
+
+thm_maps
+  "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_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_CASES" > "HOL4Compat.list_CASES"
+  "list_Axiom_old" > "HOL4Compat.list_Axiom_old"
+  "list_Axiom" > "HOL4Compat.list_Axiom"
+  "list_11" > "List.list.simps_3"
+  "ZIP_UNZIP" > "HOL4Base.list.ZIP_UNZIP"
+  "ZIP_MAP" > "HOL4Base.list.ZIP_MAP"
+  "ZIP" > "HOL4Compat.ZIP"
+  "WF_LIST_PRED" > "HOL4Base.list.WF_LIST_PRED"
+  "UNZIP_ZIP" > "HOL4Base.list.UNZIP_ZIP"
+  "UNZIP" > "HOL4Compat.UNZIP"
+  "TL" > "List.tl.simps_2"
+  "SUM" > "HOL4Compat.SUM"
+  "REVERSE_REVERSE" > "List.rev_rev_ident"
+  "REVERSE_DEF" > "HOL4Compat.REVERSE"
+  "REVERSE_APPEND" > "List.rev_append"
+  "NULL_DEF" > "HOL4Compat.NULL_DEF"
+  "NULL" > "HOL4Base.list.NULL"
+  "NOT_NIL_CONS" > "List.list.simps_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"
+  "MEM_ZIP" > "HOL4Base.list.MEM_ZIP"
+  "MEM_MAP" > "HOL4Base.list.MEM_MAP"
+  "MEM_EL" > "HOL4Base.list.MEM_EL"
+  "MEM_APPEND" > "HOL4Base.list.MEM_APPEND"
+  "MEM" > "HOL4Compat.MEM"
+  "MAP_EQ_NIL" > "HOL4Base.list.MAP_EQ_NIL"
+  "MAP_CONG" > "HOL4Base.list.MAP_CONG"
+  "MAP_APPEND" > "List.map_append"
+  "MAP2_ZIP" > "HOL4Base.list.MAP2_ZIP"
+  "MAP2" > "HOL4Compat.MAP2"
+  "MAP" > "HOL4Compat.MAP"
+  "LIST_NOT_EQ" > "HOL4Base.list.LIST_NOT_EQ"
+  "LENGTH_ZIP" > "HOL4Base.list.LENGTH_ZIP"
+  "LENGTH_UNZIP" > "HOL4Base.list.LENGTH_UNZIP"
+  "LENGTH_NIL" > "List.length_0_conv"
+  "LENGTH_MAP" > "List.length_map"
+  "LENGTH_EQ_NIL" > "HOL4Base.list.LENGTH_EQ_NIL"
+  "LENGTH_EQ_CONS" > "HOL4Base.list.LENGTH_EQ_CONS"
+  "LENGTH_CONS" > "HOL4Base.list.LENGTH_CONS"
+  "LENGTH_APPEND" > "List.length_append"
+  "LENGTH" > "HOL4Compat.LENGTH"
+  "LAST_DEF" > "List.last.simps"
+  "LAST_CONS" > "HOL4Base.list.LAST_CONS"
+  "HD" > "List.hd.simps"
+  "FRONT_DEF" > "List.butlast.simps_2"
+  "FRONT_CONS" > "HOL4Base.list.FRONT_CONS"
+  "FOLDR_CONG" > "HOL4Base.list.FOLDR_CONG"
+  "FOLDR" > "HOL4Compat.FOLDR"
+  "FOLDL_CONG" > "HOL4Base.list.FOLDL_CONG"
+  "FOLDL" > "HOL4Compat.FOLDL"
+  "FLAT" > "HOL4Compat.FLAT"
+  "FILTER" > "HOL4Compat.FILTER"
+  "EXISTS_MEM" > "HOL4Base.list.EXISTS_MEM"
+  "EXISTS_DEF" > "HOL4Compat.list_exists_DEF"
+  "EXISTS_CONG" > "HOL4Base.list.EXISTS_CONG"
+  "EXISTS_APPEND" > "HOL4Base.list.EXISTS_APPEND"
+  "EVERY_MONOTONIC" > "HOL4Base.list.EVERY_MONOTONIC"
+  "EVERY_MEM" > "HOL4Base.list.EVERY_MEM"
+  "EVERY_EL" > "HOL4Base.list.EVERY_EL"
+  "EVERY_DEF" > "HOL4Compat.EVERY_DEF"
+  "EVERY_CONJ" > "HOL4Base.list.EVERY_CONJ"
+  "EVERY_CONG" > "HOL4Base.list.EVERY_CONG"
+  "EVERY_APPEND" > "List.list_all_append"
+  "EQ_LIST" > "HOL4Base.list.EQ_LIST"
+  "EL_compute" > "HOL4Base.list.EL_compute"
+  "EL_ZIP" > "HOL4Base.list.EL_ZIP"
+  "EL" > "HOL4Base.list.EL"
+  "CONS_ACYCLIC" > "HOL4Base.list.CONS_ACYCLIC"
+  "CONS_11" > "List.list.simps_3"
+  "CONS" > "HOL4Base.list.CONS"
+  "APPEND_eq_NIL" > "HOL4Base.list.APPEND_eq_NIL"
+  "APPEND_NIL" > "List.append_Nil2"
+  "APPEND_FRONT_LAST" > "List.append_butlast_last_id"
+  "APPEND_ASSOC" > "List.append_assoc"
+  "APPEND_11" > "HOL4Base.list.APPEND_11"
+  "APPEND" > "HOL4Compat.APPEND"
+
+ignore_thms
+  "list_repfns"
+  "list_TY_DEF"
+  "list1_def"
+  "list0_def"
+  "NIL"
+  "CONS_def"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/marker.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,19 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "stmarker" > "stmarker_primdef"
+
+const_maps
+  "stmarker" > "HOL4Base.marker.stmarker"
+
+thm_maps
+  "stmarker_primdef" > "HOL4Base.marker.stmarker_primdef"
+  "stmarker_def" > "HOL4Base.marker.stmarker_def"
+  "move_right_disj" > "HOL4Base.marker.move_right_disj"
+  "move_right_conj" > "HOL4Base.marker.move_right_conj"
+  "move_left_disj" > "HOL4Base.marker.move_left_disj"
+  "move_left_conj" > "HOL4Base.marker.move_left_conj"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/nets.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,51 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "tendsto" > "tendsto_def"
+  "tends" > "tends_def"
+  "dorder" > "dorder_def"
+  "bounded" > "bounded_def"
+
+const_maps
+  "tendsto" > "HOL4Real.nets.tendsto"
+  "tends" > "HOL4Real.nets.tends"
+  "dorder" > "HOL4Real.nets.dorder"
+  "bounded" > "HOL4Real.nets.bounded"
+
+thm_maps
+  "tendsto_def" > "HOL4Real.nets.tendsto_def"
+  "tendsto" > "HOL4Real.nets.tendsto"
+  "tends_def" > "HOL4Real.nets.tends_def"
+  "tends" > "HOL4Real.nets.tends"
+  "dorder_def" > "HOL4Real.nets.dorder_def"
+  "dorder" > "HOL4Real.nets.dorder"
+  "bounded_def" > "HOL4Real.nets.bounded_def"
+  "bounded" > "HOL4Real.nets.bounded"
+  "SEQ_TENDS" > "HOL4Real.nets.SEQ_TENDS"
+  "NET_SUB" > "HOL4Real.nets.NET_SUB"
+  "NET_NULL_MUL" > "HOL4Real.nets.NET_NULL_MUL"
+  "NET_NULL_CMUL" > "HOL4Real.nets.NET_NULL_CMUL"
+  "NET_NULL_ADD" > "HOL4Real.nets.NET_NULL_ADD"
+  "NET_NULL" > "HOL4Real.nets.NET_NULL"
+  "NET_NEG" > "HOL4Real.nets.NET_NEG"
+  "NET_MUL" > "HOL4Real.nets.NET_MUL"
+  "NET_LE" > "HOL4Real.nets.NET_LE"
+  "NET_INV" > "HOL4Real.nets.NET_INV"
+  "NET_DIV" > "HOL4Real.nets.NET_DIV"
+  "NET_CONV_NZ" > "HOL4Real.nets.NET_CONV_NZ"
+  "NET_CONV_IBOUNDED" > "HOL4Real.nets.NET_CONV_IBOUNDED"
+  "NET_CONV_BOUNDED" > "HOL4Real.nets.NET_CONV_BOUNDED"
+  "NET_ADD" > "HOL4Real.nets.NET_ADD"
+  "NET_ABS" > "HOL4Real.nets.NET_ABS"
+  "MTOP_TENDS_UNIQ" > "HOL4Real.nets.MTOP_TENDS_UNIQ"
+  "MTOP_TENDS" > "HOL4Real.nets.MTOP_TENDS"
+  "MR1_BOUNDED" > "HOL4Real.nets.MR1_BOUNDED"
+  "LIM_TENDS2" > "HOL4Real.nets.LIM_TENDS2"
+  "LIM_TENDS" > "HOL4Real.nets.LIM_TENDS"
+  "DORDER_TENDSTO" > "HOL4Real.nets.DORDER_TENDSTO"
+  "DORDER_NGE" > "HOL4Real.nets.DORDER_NGE"
+  "DORDER_LEMMA" > "HOL4Real.nets.DORDER_LEMMA"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/num.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,26 @@
+import
+
+import_segment "hol4"
+
+type_maps
+  "num" > "nat"
+
+const_maps
+  "SUC" > "Suc"
+  "0" > "0" :: "nat"
+
+thm_maps
+  "NOT_SUC" > "Nat.nat.simps_1"
+  "INV_SUC" > "Nat.Suc_inject"
+  "INDUCTION" > "List.lexn.induct"
+
+ignore_thms
+  "num_TY_DEF"
+  "num_ISO_DEF"
+  "ZERO_REP_DEF"
+  "ZERO_DEF"
+  "SUC_REP_DEF"
+  "SUC_DEF"
+  "IS_NUM_REP"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/numeral.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,49 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "iiSUC" > "iiSUC_def"
+  "iZ" > "iZ_def"
+  "iSUB" > "iSUB_def"
+  "iSQR" > "iSQR_def"
+  "iDUB" > "iDUB_def"
+  "iBIT_cases" > "iBIT_cases_def"
+
+const_maps
+  "iiSUC" > "HOL4Base.numeral.iiSUC"
+  "iZ" > "HOL4Base.numeral.iZ"
+  "iSQR" > "HOL4Base.numeral.iSQR"
+  "iDUB" > "HOL4Base.numeral.iDUB"
+
+thm_maps
+  "numeral_suc" > "HOL4Base.numeral.numeral_suc"
+  "numeral_sub" > "HOL4Base.numeral.numeral_sub"
+  "numeral_pre" > "HOL4Base.numeral.numeral_pre"
+  "numeral_mult" > "HOL4Base.numeral.numeral_mult"
+  "numeral_lte" > "HOL4Base.numeral.numeral_lte"
+  "numeral_lt" > "HOL4Base.numeral.numeral_lt"
+  "numeral_iisuc" > "HOL4Base.numeral.numeral_iisuc"
+  "numeral_funpow" > "HOL4Base.numeral.numeral_funpow"
+  "numeral_fact" > "HOL4Base.numeral.numeral_fact"
+  "numeral_exp" > "HOL4Base.numeral.numeral_exp"
+  "numeral_evenodd" > "HOL4Base.numeral.numeral_evenodd"
+  "numeral_eq" > "HOL4Base.numeral.numeral_eq"
+  "numeral_distrib" > "HOL4Base.numeral.numeral_distrib"
+  "numeral_add" > "HOL4Base.numeral.numeral_add"
+  "iiSUC_def" > "HOL4Base.numeral.iiSUC_def"
+  "iiSUC" > "HOL4Base.numeral.iiSUC"
+  "iZ_def" > "HOL4Base.numeral.iZ_def"
+  "iZ" > "HOL4Base.numeral.iZ"
+  "iSUB_THM" > "HOL4Base.numeral.iSUB_THM"
+  "iSUB_DEF" > "HOL4Base.numeral.iSUB_DEF"
+  "iSQR_def" > "HOL4Base.numeral.iSQR_def"
+  "iSQR" > "HOL4Base.numeral.iSQR"
+  "iDUB_removal" > "HOL4Base.numeral.iDUB_removal"
+  "iDUB_def" > "HOL4Base.numeral.iDUB_def"
+  "iDUB" > "HOL4Base.numeral.iDUB"
+  "iBIT_cases" > "HOL4Base.numeral.iBIT_cases"
+  "bit_initiality" > "HOL4Base.numeral.bit_initiality"
+  "bit_induction" > "HOL4Base.numeral.bit_induction"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/one.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,20 @@
+import
+
+import_segment "hol4"
+
+type_maps
+  "one" > "Product_Type.unit"
+
+const_maps
+  "one" > "Product_Type.Unity"
+
+thm_maps
+  "one" > "HOL4Compat.one"
+
+ignore_thms
+  "one_axiom"
+  "one_TY_DEF"
+  "one_DEF"
+  "one_Axiom"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/operator.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,40 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "RIGHT_ID" > "RIGHT_ID_def"
+  "MONOID" > "MONOID_def"
+  "LEFT_ID" > "LEFT_ID_def"
+  "FCOMM" > "FCOMM_def"
+  "COMM" > "COMM_def"
+  "ASSOC" > "ASSOC_def"
+
+const_maps
+  "RIGHT_ID" > "HOL4Base.operator.RIGHT_ID"
+  "MONOID" > "HOL4Base.operator.MONOID"
+  "LEFT_ID" > "HOL4Base.operator.LEFT_ID"
+  "FCOMM" > "HOL4Base.operator.FCOMM"
+  "COMM" > "HOL4Base.operator.COMM"
+  "ASSOC" > "HOL4Base.operator.ASSOC"
+
+thm_maps
+  "RIGHT_ID_def" > "HOL4Base.operator.RIGHT_ID_def"
+  "RIGHT_ID_DEF" > "HOL4Base.operator.RIGHT_ID_DEF"
+  "MONOID_def" > "HOL4Base.operator.MONOID_def"
+  "MONOID_DISJ_F" > "HOL4Base.operator.MONOID_DISJ_F"
+  "MONOID_DEF" > "HOL4Base.operator.MONOID_DEF"
+  "MONOID_CONJ_T" > "HOL4Base.operator.MONOID_CONJ_T"
+  "LEFT_ID_def" > "HOL4Base.operator.LEFT_ID_def"
+  "LEFT_ID_DEF" > "HOL4Base.operator.LEFT_ID_DEF"
+  "FCOMM_def" > "HOL4Base.operator.FCOMM_def"
+  "FCOMM_DEF" > "HOL4Base.operator.FCOMM_DEF"
+  "FCOMM_ASSOC" > "HOL4Base.operator.FCOMM_ASSOC"
+  "COMM_def" > "HOL4Base.operator.COMM_def"
+  "COMM_DEF" > "HOL4Base.operator.COMM_DEF"
+  "ASSOC_def" > "HOL4Base.operator.ASSOC_def"
+  "ASSOC_DISJ" > "HOL4Base.operator.ASSOC_DISJ"
+  "ASSOC_DEF" > "HOL4Base.operator.ASSOC_DEF"
+  "ASSOC_CONJ" > "HOL4Base.operator.ASSOC_CONJ"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/option.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,45 @@
+import
+
+import_segment "hol4"
+
+type_maps
+  "option" > "Datatype.option"
+
+const_maps
+  "option_case" > "Datatype.option.option_case"
+  "THE" > "Datatype.the"
+  "SOME" > "Datatype.option.Some"
+  "OPTION_MAP" > "Datatype.option_map"
+  "OPTION_JOIN" > "HOL4Compat.OPTION_JOIN"
+  "NONE" > "Datatype.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_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"
+  "OPTION_MAP_EQ_SOME" > "HOL4Base.option.OPTION_MAP_EQ_SOME"
+  "OPTION_MAP_EQ_NONE" > "Datatype.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"
+  "IS_SOME_DEF" > "HOL4Compat.IS_SOME_DEF"
+  "IS_NONE_DEF" > "HOL4Compat.IS_NONE_DEF"
+
+ignore_thms
+  "option_axiom"
+  "option_TY_DEF"
+  "option_REP_ABS_DEF"
+  "option_Axiom"
+  "SOME_DEF"
+  "NONE_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/pair.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,70 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "RPROD" > "RPROD_def"
+  "LEX" > "LEX_def"
+
+type_maps
+  "prod" > "*"
+
+const_maps
+  "pair_case" > "split"
+  "UNCURRY" > "split"
+  "SND" > "snd"
+  "RPROD" > "HOL4Base.pair.RPROD"
+  "LEX" > "HOL4Base.pair.LEX"
+  "FST" > "fst"
+  "CURRY" > "curry"
+  "," > "Pair"
+  "##" > "prod_fun"
+
+thm_maps
+  "pair_induction" > "Product_Type.prod_induct"
+  "pair_case_thm" > "Product_Type.split"
+  "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_ONE_ONE_THM" > "HOL4Base.pair.UNCURRY_ONE_ONE_THM"
+  "UNCURRY_DEF" > "Product_Type.split"
+  "UNCURRY_CURRY_THM" > "Product_Type.split_curry"
+  "UNCURRY_CONG" > "HOL4Base.pair.UNCURRY_CONG"
+  "UNCURRY" > "Product_Type.split_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.prod_fun"
+  "PAIR_MAP" > "HOL4Compat.PAIR_MAP"
+  "PAIR_EQ" > "Product_Type.Pair_eq"
+  "PAIR" > "HOL4Compat.PAIR"
+  "LEX_def" > "HOL4Base.pair.LEX_def"
+  "LEX_DEF" > "HOL4Base.pair.LEX_DEF"
+  "LET2_RATOR" > "HOL4Base.pair.LET2_RATOR"
+  "LET2_RAND" > "HOL4Base.pair.LET2_RAND"
+  "LAMBDA_PROD" > "Product_Type.split_eta"
+  "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_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" > "Product_Type.Pair_eq"
+  "ABS_PAIR_THM" > "Product_Type.surj_pair"
+
+ignore_thms
+  "prod_TY_DEF"
+  "MK_PAIR_DEF"
+  "IS_PAIR_DEF"
+  "COMMA_DEF"
+  "ABS_REP_prod"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/poly.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,128 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "rsquarefree" > "rsquarefree_def"
+  "poly_order" > "poly_order_def"
+  "poly_neg" > "poly_neg_primdef"
+  "poly_mul" > "poly_mul_primdef"
+  "poly_exp" > "poly_exp_primdef"
+  "poly_divides" > "poly_divides_def"
+  "poly_diff_aux" > "poly_diff_aux_primdef"
+  "poly_add" > "poly_add_primdef"
+  "poly" > "poly_primdef"
+  "normalize" > "normalize_def"
+  "diff" > "diff_def"
+  "degree" > "degree_def"
+  "##" > "##_def"
+
+const_maps
+  "rsquarefree" > "HOL4Real.poly.rsquarefree"
+  "poly_order" > "HOL4Real.poly.poly_order"
+  "poly_neg" > "HOL4Real.poly.poly_neg"
+  "poly_divides" > "HOL4Real.poly.poly_divides"
+  "diff" > "HOL4Real.poly.diff"
+  "degree" > "HOL4Real.poly.degree"
+
+thm_maps
+  "rsquarefree_def" > "HOL4Real.poly.rsquarefree_def"
+  "rsquarefree" > "HOL4Real.poly.rsquarefree"
+  "poly_order_def" > "HOL4Real.poly.poly_order_def"
+  "poly_order" > "HOL4Real.poly.poly_order"
+  "poly_neg_primdef" > "HOL4Real.poly.poly_neg_primdef"
+  "poly_neg_def" > "HOL4Real.poly.poly_neg_def"
+  "poly_mul_def" > "HOL4Real.poly.poly_mul_def"
+  "poly_exp_def" > "HOL4Real.poly.poly_exp_def"
+  "poly_divides_def" > "HOL4Real.poly.poly_divides_def"
+  "poly_divides" > "HOL4Real.poly.poly_divides"
+  "poly_diff_def" > "HOL4Real.poly.poly_diff_def"
+  "poly_diff_aux_def" > "HOL4Real.poly.poly_diff_aux_def"
+  "poly_def" > "HOL4Real.poly.poly_def"
+  "poly_cmul_def" > "HOL4Real.poly.poly_cmul_def"
+  "poly_add_def" > "HOL4Real.poly.poly_add_def"
+  "normalize" > "HOL4Real.poly.normalize"
+  "diff_def" > "HOL4Real.poly.diff_def"
+  "degree_def" > "HOL4Real.poly.degree_def"
+  "degree" > "HOL4Real.poly.degree"
+  "RSQUAREFREE_ROOTS" > "HOL4Real.poly.RSQUAREFREE_ROOTS"
+  "RSQUAREFREE_DECOMP" > "HOL4Real.poly.RSQUAREFREE_DECOMP"
+  "POLY_ZERO_LEMMA" > "HOL4Real.poly.POLY_ZERO_LEMMA"
+  "POLY_ZERO" > "HOL4Real.poly.POLY_ZERO"
+  "POLY_SQUAREFREE_DECOMP_ORDER" > "HOL4Real.poly.POLY_SQUAREFREE_DECOMP_ORDER"
+  "POLY_SQUAREFREE_DECOMP" > "HOL4Real.poly.POLY_SQUAREFREE_DECOMP"
+  "POLY_ROOTS_INDEX_LENGTH" > "HOL4Real.poly.POLY_ROOTS_INDEX_LENGTH"
+  "POLY_ROOTS_INDEX_LEMMA" > "HOL4Real.poly.POLY_ROOTS_INDEX_LEMMA"
+  "POLY_ROOTS_FINITE_SET" > "HOL4Real.poly.POLY_ROOTS_FINITE_SET"
+  "POLY_ROOTS_FINITE_LEMMA" > "HOL4Real.poly.POLY_ROOTS_FINITE_LEMMA"
+  "POLY_ROOTS_FINITE" > "HOL4Real.poly.POLY_ROOTS_FINITE"
+  "POLY_PRIME_EQ_0" > "HOL4Real.poly.POLY_PRIME_EQ_0"
+  "POLY_PRIMES" > "HOL4Real.poly.POLY_PRIMES"
+  "POLY_ORDER_EXISTS" > "HOL4Real.poly.POLY_ORDER_EXISTS"
+  "POLY_ORDER" > "HOL4Real.poly.POLY_ORDER"
+  "POLY_NORMALIZE" > "HOL4Real.poly.POLY_NORMALIZE"
+  "POLY_NEG_CLAUSES" > "HOL4Real.poly.POLY_NEG_CLAUSES"
+  "POLY_NEG" > "HOL4Real.poly.POLY_NEG"
+  "POLY_MVT" > "HOL4Real.poly.POLY_MVT"
+  "POLY_MUL_LCANCEL" > "HOL4Real.poly.POLY_MUL_LCANCEL"
+  "POLY_MUL_CLAUSES" > "HOL4Real.poly.POLY_MUL_CLAUSES"
+  "POLY_MUL_ASSOC" > "HOL4Real.poly.POLY_MUL_ASSOC"
+  "POLY_MUL" > "HOL4Real.poly.POLY_MUL"
+  "POLY_MONO" > "HOL4Real.poly.POLY_MONO"
+  "POLY_LINEAR_REM" > "HOL4Real.poly.POLY_LINEAR_REM"
+  "POLY_LINEAR_DIVIDES" > "HOL4Real.poly.POLY_LINEAR_DIVIDES"
+  "POLY_LENGTH_MUL" > "HOL4Real.poly.POLY_LENGTH_MUL"
+  "POLY_IVT_POS" > "HOL4Real.poly.POLY_IVT_POS"
+  "POLY_IVT_NEG" > "HOL4Real.poly.POLY_IVT_NEG"
+  "POLY_EXP_PRIME_EQ_0" > "HOL4Real.poly.POLY_EXP_PRIME_EQ_0"
+  "POLY_EXP_EQ_0" > "HOL4Real.poly.POLY_EXP_EQ_0"
+  "POLY_EXP_DIVIDES" > "HOL4Real.poly.POLY_EXP_DIVIDES"
+  "POLY_EXP_ADD" > "HOL4Real.poly.POLY_EXP_ADD"
+  "POLY_EXP" > "HOL4Real.poly.POLY_EXP"
+  "POLY_ENTIRE_LEMMA" > "HOL4Real.poly.POLY_ENTIRE_LEMMA"
+  "POLY_ENTIRE" > "HOL4Real.poly.POLY_ENTIRE"
+  "POLY_DIVIDES_ZERO" > "HOL4Real.poly.POLY_DIVIDES_ZERO"
+  "POLY_DIVIDES_TRANS" > "HOL4Real.poly.POLY_DIVIDES_TRANS"
+  "POLY_DIVIDES_SUB2" > "HOL4Real.poly.POLY_DIVIDES_SUB2"
+  "POLY_DIVIDES_SUB" > "HOL4Real.poly.POLY_DIVIDES_SUB"
+  "POLY_DIVIDES_REFL" > "HOL4Real.poly.POLY_DIVIDES_REFL"
+  "POLY_DIVIDES_EXP" > "HOL4Real.poly.POLY_DIVIDES_EXP"
+  "POLY_DIVIDES_ADD" > "HOL4Real.poly.POLY_DIVIDES_ADD"
+  "POLY_DIFF_ZERO" > "HOL4Real.poly.POLY_DIFF_ZERO"
+  "POLY_DIFF_WELLDEF" > "HOL4Real.poly.POLY_DIFF_WELLDEF"
+  "POLY_DIFF_NEG" > "HOL4Real.poly.POLY_DIFF_NEG"
+  "POLY_DIFF_MUL_LEMMA" > "HOL4Real.poly.POLY_DIFF_MUL_LEMMA"
+  "POLY_DIFF_MUL" > "HOL4Real.poly.POLY_DIFF_MUL"
+  "POLY_DIFF_LEMMA" > "HOL4Real.poly.POLY_DIFF_LEMMA"
+  "POLY_DIFF_ISZERO" > "HOL4Real.poly.POLY_DIFF_ISZERO"
+  "POLY_DIFF_EXP_PRIME" > "HOL4Real.poly.POLY_DIFF_EXP_PRIME"
+  "POLY_DIFF_EXP" > "HOL4Real.poly.POLY_DIFF_EXP"
+  "POLY_DIFF_CMUL" > "HOL4Real.poly.POLY_DIFF_CMUL"
+  "POLY_DIFF_CLAUSES" > "HOL4Real.poly.POLY_DIFF_CLAUSES"
+  "POLY_DIFF_AUX_NEG" > "HOL4Real.poly.POLY_DIFF_AUX_NEG"
+  "POLY_DIFF_AUX_MUL_LEMMA" > "HOL4Real.poly.POLY_DIFF_AUX_MUL_LEMMA"
+  "POLY_DIFF_AUX_ISZERO" > "HOL4Real.poly.POLY_DIFF_AUX_ISZERO"
+  "POLY_DIFF_AUX_CMUL" > "HOL4Real.poly.POLY_DIFF_AUX_CMUL"
+  "POLY_DIFF_AUX_ADD" > "HOL4Real.poly.POLY_DIFF_AUX_ADD"
+  "POLY_DIFF_ADD" > "HOL4Real.poly.POLY_DIFF_ADD"
+  "POLY_DIFFERENTIABLE" > "HOL4Real.poly.POLY_DIFFERENTIABLE"
+  "POLY_DIFF" > "HOL4Real.poly.POLY_DIFF"
+  "POLY_CONT" > "HOL4Real.poly.POLY_CONT"
+  "POLY_CMUL_CLAUSES" > "HOL4Real.poly.POLY_CMUL_CLAUSES"
+  "POLY_CMUL" > "HOL4Real.poly.POLY_CMUL"
+  "POLY_ADD_RZERO" > "HOL4Real.poly.POLY_ADD_RZERO"
+  "POLY_ADD_CLAUSES" > "HOL4Real.poly.POLY_ADD_CLAUSES"
+  "POLY_ADD" > "HOL4Real.poly.POLY_ADD"
+  "ORDER_UNIQUE" > "HOL4Real.poly.ORDER_UNIQUE"
+  "ORDER_THM" > "HOL4Real.poly.ORDER_THM"
+  "ORDER_ROOT" > "HOL4Real.poly.ORDER_ROOT"
+  "ORDER_POLY" > "HOL4Real.poly.ORDER_POLY"
+  "ORDER_MUL" > "HOL4Real.poly.ORDER_MUL"
+  "ORDER_DIVIDES" > "HOL4Real.poly.ORDER_DIVIDES"
+  "ORDER_DIFF" > "HOL4Real.poly.ORDER_DIFF"
+  "ORDER_DECOMP" > "HOL4Real.poly.ORDER_DECOMP"
+  "ORDER" > "HOL4Real.poly.ORDER"
+  "FINITE_LEMMA" > "HOL4Real.poly.FINITE_LEMMA"
+  "DEGREE_ZERO" > "HOL4Real.poly.DEGREE_ZERO"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/powser.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,30 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "diffs" > "diffs_def"
+
+const_maps
+  "diffs" > "HOL4Real.powser.diffs"
+
+thm_maps
+  "diffs_def" > "HOL4Real.powser.diffs_def"
+  "diffs" > "HOL4Real.powser.diffs"
+  "TERMDIFF_LEMMA5" > "HOL4Real.powser.TERMDIFF_LEMMA5"
+  "TERMDIFF_LEMMA4" > "HOL4Real.powser.TERMDIFF_LEMMA4"
+  "TERMDIFF_LEMMA3" > "HOL4Real.powser.TERMDIFF_LEMMA3"
+  "TERMDIFF_LEMMA2" > "HOL4Real.powser.TERMDIFF_LEMMA2"
+  "TERMDIFF_LEMMA1" > "HOL4Real.powser.TERMDIFF_LEMMA1"
+  "TERMDIFF" > "HOL4Real.powser.TERMDIFF"
+  "POWSER_INSIDEA" > "HOL4Real.powser.POWSER_INSIDEA"
+  "POWSER_INSIDE" > "HOL4Real.powser.POWSER_INSIDE"
+  "POWREV" > "HOL4Real.powser.POWREV"
+  "POWDIFF_LEMMA" > "HOL4Real.powser.POWDIFF_LEMMA"
+  "POWDIFF" > "HOL4Real.powser.POWDIFF"
+  "DIFFS_NEG" > "HOL4Real.powser.DIFFS_NEG"
+  "DIFFS_LEMMA2" > "HOL4Real.powser.DIFFS_LEMMA2"
+  "DIFFS_LEMMA" > "HOL4Real.powser.DIFFS_LEMMA"
+  "DIFFS_EQUIV" > "HOL4Real.powser.DIFFS_EQUIV"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/pred_set.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,322 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "count" > "count_primdef"
+  "UNIV" > "UNIV_def"
+  "UNION" > "UNION_def"
+  "SURJ" > "SURJ_def"
+  "SUBSET" > "SUBSET_def"
+  "SING" > "SING_def"
+  "RINV" > "RINV_def"
+  "REST" > "REST_def"
+  "PSUBSET" > "PSUBSET_def"
+  "LINV" > "LINV_def"
+  "ITSET_tupled" > "ITSET_tupled_def"
+  "ITSET" > "ITSET_def"
+  "INTER" > "INTER_def"
+  "INSERT" > "INSERT_def"
+  "INJ" > "INJ_def"
+  "INFINITE" > "INFINITE_def"
+  "IMAGE" > "IMAGE_def"
+  "GSPEC" > "GSPEC_def"
+  "FINITE" > "FINITE_def"
+  "EMPTY" > "EMPTY_def"
+  "DISJOINT" > "DISJOINT_def"
+  "DIFF" > "DIFF_def"
+  "DELETE" > "DELETE_def"
+  "CROSS" > "CROSS_def"
+  "COMPL" > "COMPL_def"
+  "CHOICE" > "CHOICE_def"
+  "CARD" > "CARD_def"
+  "BIJ" > "BIJ_def"
+  "BIGUNION" > "BIGUNION_def"
+  "BIGINTER" > "BIGINTER_def"
+
+const_maps
+  "count" > "HOL4Base.pred_set.count"
+  "UNIV" > "HOL4Base.pred_set.UNIV"
+  "UNION" > "HOL4Base.pred_set.UNION"
+  "SURJ" > "HOL4Base.pred_set.SURJ"
+  "SUBSET" > "HOL4Base.pred_set.SUBSET"
+  "SING" > "HOL4Base.pred_set.SING"
+  "REST" > "HOL4Base.pred_set.REST"
+  "PSUBSET" > "HOL4Base.pred_set.PSUBSET"
+  "ITSET_tupled" > "HOL4Base.pred_set.ITSET_tupled"
+  "ITSET" > "HOL4Base.pred_set.ITSET"
+  "INTER" > "HOL4Base.pred_set.INTER"
+  "INSERT" > "HOL4Base.pred_set.INSERT"
+  "INJ" > "HOL4Base.pred_set.INJ"
+  "INFINITE" > "HOL4Base.pred_set.INFINITE"
+  "IMAGE" > "HOL4Base.pred_set.IMAGE"
+  "FINITE" > "HOL4Base.pred_set.FINITE"
+  "EMPTY" > "HOL4Base.pred_set.EMPTY"
+  "DISJOINT" > "HOL4Base.pred_set.DISJOINT"
+  "DIFF" > "HOL4Base.pred_set.DIFF"
+  "DELETE" > "HOL4Base.pred_set.DELETE"
+  "CROSS" > "HOL4Base.pred_set.CROSS"
+  "COMPL" > "HOL4Base.pred_set.COMPL"
+  "BIJ" > "HOL4Base.pred_set.BIJ"
+  "BIGUNION" > "HOL4Base.pred_set.BIGUNION"
+  "BIGINTER" > "HOL4Base.pred_set.BIGINTER"
+
+thm_maps
+  "count_primdef" > "HOL4Base.pred_set.count_primdef"
+  "count_def" > "HOL4Base.pred_set.count_def"
+  "UNIV_def" > "HOL4Base.pred_set.UNIV_def"
+  "UNIV_SUBSET" > "HOL4Base.pred_set.UNIV_SUBSET"
+  "UNIV_NOT_EMPTY" > "HOL4Base.pred_set.UNIV_NOT_EMPTY"
+  "UNIV_DEF" > "HOL4Base.pred_set.UNIV_DEF"
+  "UNION_def" > "HOL4Base.pred_set.UNION_def"
+  "UNION_UNIV" > "HOL4Base.pred_set.UNION_UNIV"
+  "UNION_SUBSET" > "HOL4Base.pred_set.UNION_SUBSET"
+  "UNION_OVER_INTER" > "HOL4Base.pred_set.UNION_OVER_INTER"
+  "UNION_IDEMPOT" > "HOL4Base.pred_set.UNION_IDEMPOT"
+  "UNION_EMPTY" > "HOL4Base.pred_set.UNION_EMPTY"
+  "UNION_DEF" > "HOL4Base.pred_set.UNION_DEF"
+  "UNION_COMM" > "HOL4Base.pred_set.UNION_COMM"
+  "UNION_ASSOC" > "HOL4Base.pred_set.UNION_ASSOC"
+  "SURJ_def" > "HOL4Base.pred_set.SURJ_def"
+  "SURJ_ID" > "HOL4Base.pred_set.SURJ_ID"
+  "SURJ_EMPTY" > "HOL4Base.pred_set.SURJ_EMPTY"
+  "SURJ_DEF" > "HOL4Base.pred_set.SURJ_DEF"
+  "SURJ_COMPOSE" > "HOL4Base.pred_set.SURJ_COMPOSE"
+  "SUBSET_def" > "HOL4Base.pred_set.SUBSET_def"
+  "SUBSET_UNIV" > "HOL4Base.pred_set.SUBSET_UNIV"
+  "SUBSET_UNION_ABSORPTION" > "HOL4Base.pred_set.SUBSET_UNION_ABSORPTION"
+  "SUBSET_UNION" > "HOL4Base.pred_set.SUBSET_UNION"
+  "SUBSET_TRANS" > "HOL4Base.pred_set.SUBSET_TRANS"
+  "SUBSET_REFL" > "HOL4Base.pred_set.SUBSET_REFL"
+  "SUBSET_INTER_ABSORPTION" > "HOL4Base.pred_set.SUBSET_INTER_ABSORPTION"
+  "SUBSET_INTER" > "HOL4Base.pred_set.SUBSET_INTER"
+  "SUBSET_INSERT_DELETE" > "HOL4Base.pred_set.SUBSET_INSERT_DELETE"
+  "SUBSET_INSERT" > "HOL4Base.pred_set.SUBSET_INSERT"
+  "SUBSET_FINITE" > "HOL4Base.pred_set.SUBSET_FINITE"
+  "SUBSET_EMPTY" > "HOL4Base.pred_set.SUBSET_EMPTY"
+  "SUBSET_DELETE" > "HOL4Base.pred_set.SUBSET_DELETE"
+  "SUBSET_DEF" > "HOL4Base.pred_set.SUBSET_DEF"
+  "SUBSET_BIGINTER" > "HOL4Base.pred_set.SUBSET_BIGINTER"
+  "SUBSET_ANTISYM" > "HOL4Base.pred_set.SUBSET_ANTISYM"
+  "SPECIFICATION" > "HOL4Base.bool.IN_DEF"
+  "SING_def" > "HOL4Base.pred_set.SING_def"
+  "SING_IFF_EMPTY_REST" > "HOL4Base.pred_set.SING_IFF_EMPTY_REST"
+  "SING_IFF_CARD1" > "HOL4Base.pred_set.SING_IFF_CARD1"
+  "SING_FINITE" > "HOL4Base.pred_set.SING_FINITE"
+  "SING_DELETE" > "HOL4Base.pred_set.SING_DELETE"
+  "SING_DEF" > "HOL4Base.pred_set.SING_DEF"
+  "SING" > "HOL4Base.pred_set.SING"
+  "SET_MINIMUM" > "HOL4Base.pred_set.SET_MINIMUM"
+  "SET_CASES" > "HOL4Base.pred_set.SET_CASES"
+  "RINV_DEF" > "HOL4Base.pred_set.RINV_DEF"
+  "REST_def" > "HOL4Base.pred_set.REST_def"
+  "REST_SUBSET" > "HOL4Base.pred_set.REST_SUBSET"
+  "REST_SING" > "HOL4Base.pred_set.REST_SING"
+  "REST_PSUBSET" > "HOL4Base.pred_set.REST_PSUBSET"
+  "REST_DEF" > "HOL4Base.pred_set.REST_DEF"
+  "PSUBSET_def" > "HOL4Base.pred_set.PSUBSET_def"
+  "PSUBSET_UNIV" > "HOL4Base.pred_set.PSUBSET_UNIV"
+  "PSUBSET_TRANS" > "HOL4Base.pred_set.PSUBSET_TRANS"
+  "PSUBSET_MEMBER" > "HOL4Base.pred_set.PSUBSET_MEMBER"
+  "PSUBSET_IRREFL" > "HOL4Base.pred_set.PSUBSET_IRREFL"
+  "PSUBSET_INSERT_SUBSET" > "HOL4Base.pred_set.PSUBSET_INSERT_SUBSET"
+  "PSUBSET_FINITE" > "HOL4Base.pred_set.PSUBSET_FINITE"
+  "PSUBSET_DEF" > "HOL4Base.pred_set.PSUBSET_DEF"
+  "NUM_SET_WOP" > "HOL4Base.pred_set.NUM_SET_WOP"
+  "NOT_UNIV_PSUBSET" > "HOL4Base.pred_set.NOT_UNIV_PSUBSET"
+  "NOT_SING_EMPTY" > "HOL4Base.pred_set.NOT_SING_EMPTY"
+  "NOT_PSUBSET_EMPTY" > "HOL4Base.pred_set.NOT_PSUBSET_EMPTY"
+  "NOT_IN_FINITE" > "HOL4Base.pred_set.NOT_IN_FINITE"
+  "NOT_IN_EMPTY" > "HOL4Base.pred_set.NOT_IN_EMPTY"
+  "NOT_INSERT_EMPTY" > "HOL4Base.pred_set.NOT_INSERT_EMPTY"
+  "NOT_EQUAL_SETS" > "HOL4Base.pred_set.NOT_EQUAL_SETS"
+  "NOT_EMPTY_SING" > "HOL4Base.pred_set.NOT_EMPTY_SING"
+  "NOT_EMPTY_INSERT" > "HOL4Base.pred_set.NOT_EMPTY_INSERT"
+  "MEMBER_NOT_EMPTY" > "HOL4Base.pred_set.MEMBER_NOT_EMPTY"
+  "LINV_DEF" > "HOL4Base.pred_set.LINV_DEF"
+  "LESS_CARD_DIFF" > "HOL4Base.pred_set.LESS_CARD_DIFF"
+  "ITSET_tupled_primitive_def" > "HOL4Base.pred_set.ITSET_tupled_primitive_def"
+  "ITSET_tupled_def" > "HOL4Base.pred_set.ITSET_tupled_def"
+  "ITSET_def" > "HOL4Base.pred_set.ITSET_def"
+  "ITSET_curried_def" > "HOL4Base.pred_set.ITSET_curried_def"
+  "ITSET_THM" > "HOL4Base.pred_set.ITSET_THM"
+  "ITSET_IND" > "HOL4Base.pred_set.ITSET_IND"
+  "ITSET_EMPTY" > "HOL4Base.pred_set.ITSET_EMPTY"
+  "IN_UNIV" > "HOL4Base.pred_set.IN_UNIV"
+  "IN_UNION" > "HOL4Base.pred_set.IN_UNION"
+  "IN_SING" > "HOL4Base.pred_set.IN_SING"
+  "IN_INTER" > "HOL4Base.pred_set.IN_INTER"
+  "IN_INSERT" > "HOL4Base.pred_set.IN_INSERT"
+  "IN_INFINITE_NOT_FINITE" > "HOL4Base.pred_set.IN_INFINITE_NOT_FINITE"
+  "IN_IMAGE" > "HOL4Base.pred_set.IN_IMAGE"
+  "IN_DISJOINT" > "HOL4Base.pred_set.IN_DISJOINT"
+  "IN_DIFF" > "HOL4Base.pred_set.IN_DIFF"
+  "IN_DELETE_EQ" > "HOL4Base.pred_set.IN_DELETE_EQ"
+  "IN_DELETE" > "HOL4Base.pred_set.IN_DELETE"
+  "IN_CROSS" > "HOL4Base.pred_set.IN_CROSS"
+  "IN_COUNT" > "HOL4Base.pred_set.IN_COUNT"
+  "IN_COMPL" > "HOL4Base.pred_set.IN_COMPL"
+  "IN_BIGUNION" > "HOL4Base.pred_set.IN_BIGUNION"
+  "IN_BIGINTER" > "HOL4Base.pred_set.IN_BIGINTER"
+  "INTER_def" > "HOL4Base.pred_set.INTER_def"
+  "INTER_UNIV" > "HOL4Base.pred_set.INTER_UNIV"
+  "INTER_UNION_COMPL" > "HOL4Base.pred_set.INTER_UNION_COMPL"
+  "INTER_SUBSET" > "HOL4Base.pred_set.INTER_SUBSET"
+  "INTER_OVER_UNION" > "HOL4Base.pred_set.INTER_OVER_UNION"
+  "INTER_IDEMPOT" > "HOL4Base.pred_set.INTER_IDEMPOT"
+  "INTER_FINITE" > "HOL4Base.pred_set.INTER_FINITE"
+  "INTER_EMPTY" > "HOL4Base.pred_set.INTER_EMPTY"
+  "INTER_DEF" > "HOL4Base.pred_set.INTER_DEF"
+  "INTER_COMM" > "HOL4Base.pred_set.INTER_COMM"
+  "INTER_ASSOC" > "HOL4Base.pred_set.INTER_ASSOC"
+  "INSERT_def" > "HOL4Base.pred_set.INSERT_def"
+  "INSERT_UNIV" > "HOL4Base.pred_set.INSERT_UNIV"
+  "INSERT_UNION_EQ" > "HOL4Base.pred_set.INSERT_UNION_EQ"
+  "INSERT_UNION" > "HOL4Base.pred_set.INSERT_UNION"
+  "INSERT_SUBSET" > "HOL4Base.pred_set.INSERT_SUBSET"
+  "INSERT_SING_UNION" > "HOL4Base.pred_set.INSERT_SING_UNION"
+  "INSERT_INTER" > "HOL4Base.pred_set.INSERT_INTER"
+  "INSERT_INSERT" > "HOL4Base.pred_set.INSERT_INSERT"
+  "INSERT_DIFF" > "HOL4Base.pred_set.INSERT_DIFF"
+  "INSERT_DELETE" > "HOL4Base.pred_set.INSERT_DELETE"
+  "INSERT_DEF" > "HOL4Base.pred_set.INSERT_DEF"
+  "INSERT_COMM" > "HOL4Base.pred_set.INSERT_COMM"
+  "INJ_def" > "HOL4Base.pred_set.INJ_def"
+  "INJ_ID" > "HOL4Base.pred_set.INJ_ID"
+  "INJ_EMPTY" > "HOL4Base.pred_set.INJ_EMPTY"
+  "INJ_DEF" > "HOL4Base.pred_set.INJ_DEF"
+  "INJ_COMPOSE" > "HOL4Base.pred_set.INJ_COMPOSE"
+  "INFINITE_def" > "HOL4Base.pred_set.INFINITE_def"
+  "INFINITE_UNIV" > "HOL4Base.pred_set.INFINITE_UNIV"
+  "INFINITE_SUBSET" > "HOL4Base.pred_set.INFINITE_SUBSET"
+  "INFINITE_INHAB" > "HOL4Base.pred_set.INFINITE_INHAB"
+  "INFINITE_DIFF_FINITE" > "HOL4Base.pred_set.INFINITE_DIFF_FINITE"
+  "INFINITE_DEF" > "HOL4Base.pred_set.INFINITE_DEF"
+  "IMAGE_def" > "HOL4Base.pred_set.IMAGE_def"
+  "IMAGE_UNION" > "HOL4Base.pred_set.IMAGE_UNION"
+  "IMAGE_SURJ" > "HOL4Base.pred_set.IMAGE_SURJ"
+  "IMAGE_SUBSET" > "HOL4Base.pred_set.IMAGE_SUBSET"
+  "IMAGE_INTER" > "HOL4Base.pred_set.IMAGE_INTER"
+  "IMAGE_INSERT" > "HOL4Base.pred_set.IMAGE_INSERT"
+  "IMAGE_IN" > "HOL4Base.pred_set.IMAGE_IN"
+  "IMAGE_ID" > "HOL4Base.pred_set.IMAGE_ID"
+  "IMAGE_FINITE" > "HOL4Base.pred_set.IMAGE_FINITE"
+  "IMAGE_EQ_EMPTY" > "HOL4Base.pred_set.IMAGE_EQ_EMPTY"
+  "IMAGE_EMPTY" > "HOL4Base.pred_set.IMAGE_EMPTY"
+  "IMAGE_DELETE" > "HOL4Base.pred_set.IMAGE_DELETE"
+  "IMAGE_DEF" > "HOL4Base.pred_set.IMAGE_DEF"
+  "IMAGE_COMPOSE" > "HOL4Base.pred_set.IMAGE_COMPOSE"
+  "IMAGE_11_INFINITE" > "HOL4Base.pred_set.IMAGE_11_INFINITE"
+  "GSPECIFICATION" > "HOL4Base.pred_set.GSPECIFICATION"
+  "FINITE_def" > "HOL4Base.pred_set.FINITE_def"
+  "FINITE_WEAK_ENUMERATE" > "HOL4Base.pred_set.FINITE_WEAK_ENUMERATE"
+  "FINITE_UNION" > "HOL4Base.pred_set.FINITE_UNION"
+  "FINITE_SING" > "HOL4Base.pred_set.FINITE_SING"
+  "FINITE_PSUBSET_UNIV" > "HOL4Base.pred_set.FINITE_PSUBSET_UNIV"
+  "FINITE_PSUBSET_INFINITE" > "HOL4Base.pred_set.FINITE_PSUBSET_INFINITE"
+  "FINITE_ISO_NUM" > "HOL4Base.pred_set.FINITE_ISO_NUM"
+  "FINITE_INSERT" > "HOL4Base.pred_set.FINITE_INSERT"
+  "FINITE_INDUCT" > "HOL4Base.pred_set.FINITE_INDUCT"
+  "FINITE_EMPTY" > "HOL4Base.pred_set.FINITE_EMPTY"
+  "FINITE_DIFF" > "HOL4Base.pred_set.FINITE_DIFF"
+  "FINITE_DELETE" > "HOL4Base.pred_set.FINITE_DELETE"
+  "FINITE_DEF" > "HOL4Base.pred_set.FINITE_DEF"
+  "FINITE_CROSS_EQ" > "HOL4Base.pred_set.FINITE_CROSS_EQ"
+  "FINITE_CROSS" > "HOL4Base.pred_set.FINITE_CROSS"
+  "FINITE_COUNT" > "HOL4Base.pred_set.FINITE_COUNT"
+  "FINITE_COMPLETE_INDUCTION" > "HOL4Base.pred_set.FINITE_COMPLETE_INDUCTION"
+  "FINITE_BIGUNION" > "HOL4Base.pred_set.FINITE_BIGUNION"
+  "EXTENSION" > "HOL4Base.pred_set.EXTENSION"
+  "EQ_UNIV" > "HOL4Base.pred_set.EQ_UNIV"
+  "EQUAL_SING" > "HOL4Base.pred_set.EQUAL_SING"
+  "EMPTY_def" > "HOL4Base.pred_set.EMPTY_def"
+  "EMPTY_UNION" > "HOL4Base.pred_set.EMPTY_UNION"
+  "EMPTY_SUBSET" > "HOL4Base.pred_set.EMPTY_SUBSET"
+  "EMPTY_NOT_UNIV" > "HOL4Base.pred_set.EMPTY_NOT_UNIV"
+  "EMPTY_DIFF" > "HOL4Base.pred_set.EMPTY_DIFF"
+  "EMPTY_DELETE" > "HOL4Base.pred_set.EMPTY_DELETE"
+  "EMPTY_DEF" > "HOL4Base.pred_set.EMPTY_DEF"
+  "DISJOINT_def" > "HOL4Base.pred_set.DISJOINT_def"
+  "DISJOINT_UNION_BOTH" > "HOL4Base.pred_set.DISJOINT_UNION_BOTH"
+  "DISJOINT_UNION" > "HOL4Base.pred_set.DISJOINT_UNION"
+  "DISJOINT_SYM" > "HOL4Base.pred_set.DISJOINT_SYM"
+  "DISJOINT_SING_EMPTY" > "HOL4Base.pred_set.DISJOINT_SING_EMPTY"
+  "DISJOINT_INSERT" > "HOL4Base.pred_set.DISJOINT_INSERT"
+  "DISJOINT_EMPTY_REFL" > "HOL4Base.pred_set.DISJOINT_EMPTY_REFL"
+  "DISJOINT_EMPTY" > "HOL4Base.pred_set.DISJOINT_EMPTY"
+  "DISJOINT_DELETE_SYM" > "HOL4Base.pred_set.DISJOINT_DELETE_SYM"
+  "DISJOINT_DEF" > "HOL4Base.pred_set.DISJOINT_DEF"
+  "DISJOINT_BIGUNION" > "HOL4Base.pred_set.DISJOINT_BIGUNION"
+  "DISJOINT_BIGINTER" > "HOL4Base.pred_set.DISJOINT_BIGINTER"
+  "DIFF_def" > "HOL4Base.pred_set.DIFF_def"
+  "DIFF_UNIV" > "HOL4Base.pred_set.DIFF_UNIV"
+  "DIFF_INSERT" > "HOL4Base.pred_set.DIFF_INSERT"
+  "DIFF_EQ_EMPTY" > "HOL4Base.pred_set.DIFF_EQ_EMPTY"
+  "DIFF_EMPTY" > "HOL4Base.pred_set.DIFF_EMPTY"
+  "DIFF_DIFF" > "HOL4Base.pred_set.DIFF_DIFF"
+  "DIFF_DEF" > "HOL4Base.pred_set.DIFF_DEF"
+  "DELETE_def" > "HOL4Base.pred_set.DELETE_def"
+  "DELETE_SUBSET" > "HOL4Base.pred_set.DELETE_SUBSET"
+  "DELETE_NON_ELEMENT" > "HOL4Base.pred_set.DELETE_NON_ELEMENT"
+  "DELETE_INTER" > "HOL4Base.pred_set.DELETE_INTER"
+  "DELETE_INSERT" > "HOL4Base.pred_set.DELETE_INSERT"
+  "DELETE_EQ_SING" > "HOL4Base.pred_set.DELETE_EQ_SING"
+  "DELETE_DELETE" > "HOL4Base.pred_set.DELETE_DELETE"
+  "DELETE_DEF" > "HOL4Base.pred_set.DELETE_DEF"
+  "DELETE_COMM" > "HOL4Base.pred_set.DELETE_COMM"
+  "DECOMPOSITION" > "HOL4Base.pred_set.DECOMPOSITION"
+  "CROSS_def" > "HOL4Base.pred_set.CROSS_def"
+  "CROSS_SUBSET" > "HOL4Base.pred_set.CROSS_SUBSET"
+  "CROSS_SINGS" > "HOL4Base.pred_set.CROSS_SINGS"
+  "CROSS_INSERT_RIGHT" > "HOL4Base.pred_set.CROSS_INSERT_RIGHT"
+  "CROSS_INSERT_LEFT" > "HOL4Base.pred_set.CROSS_INSERT_LEFT"
+  "CROSS_EMPTY" > "HOL4Base.pred_set.CROSS_EMPTY"
+  "CROSS_DEF" > "HOL4Base.pred_set.CROSS_DEF"
+  "COUNT_ZERO" > "HOL4Base.pred_set.COUNT_ZERO"
+  "COUNT_SUC" > "HOL4Base.pred_set.COUNT_SUC"
+  "COMPONENT" > "HOL4Base.pred_set.COMPONENT"
+  "COMPL_def" > "HOL4Base.pred_set.COMPL_def"
+  "COMPL_SPLITS" > "HOL4Base.pred_set.COMPL_SPLITS"
+  "COMPL_EMPTY" > "HOL4Base.pred_set.COMPL_EMPTY"
+  "COMPL_DEF" > "HOL4Base.pred_set.COMPL_DEF"
+  "COMPL_COMPL" > "HOL4Base.pred_set.COMPL_COMPL"
+  "COMPL_CLAUSES" > "HOL4Base.pred_set.COMPL_CLAUSES"
+  "CHOICE_SING" > "HOL4Base.pred_set.CHOICE_SING"
+  "CHOICE_NOT_IN_REST" > "HOL4Base.pred_set.CHOICE_NOT_IN_REST"
+  "CHOICE_INSERT_REST" > "HOL4Base.pred_set.CHOICE_INSERT_REST"
+  "CHOICE_DEF" > "HOL4Base.pred_set.CHOICE_DEF"
+  "CARD_UNION" > "HOL4Base.pred_set.CARD_UNION"
+  "CARD_SUBSET" > "HOL4Base.pred_set.CARD_SUBSET"
+  "CARD_SING_CROSS" > "HOL4Base.pred_set.CARD_SING_CROSS"
+  "CARD_SING" > "HOL4Base.pred_set.CARD_SING"
+  "CARD_PSUBSET" > "HOL4Base.pred_set.CARD_PSUBSET"
+  "CARD_INTER_LESS_EQ" > "HOL4Base.pred_set.CARD_INTER_LESS_EQ"
+  "CARD_INSERT" > "HOL4Base.pred_set.CARD_INSERT"
+  "CARD_EQ_0" > "HOL4Base.pred_set.CARD_EQ_0"
+  "CARD_EMPTY" > "HOL4Base.pred_set.CARD_EMPTY"
+  "CARD_DIFF" > "HOL4Base.pred_set.CARD_DIFF"
+  "CARD_DELETE" > "HOL4Base.pred_set.CARD_DELETE"
+  "CARD_DEF" > "HOL4Base.pred_set.CARD_DEF"
+  "CARD_CROSS" > "HOL4Base.pred_set.CARD_CROSS"
+  "CARD_COUNT" > "HOL4Base.pred_set.CARD_COUNT"
+  "BIJ_def" > "HOL4Base.pred_set.BIJ_def"
+  "BIJ_ID" > "HOL4Base.pred_set.BIJ_ID"
+  "BIJ_EMPTY" > "HOL4Base.pred_set.BIJ_EMPTY"
+  "BIJ_DEF" > "HOL4Base.pred_set.BIJ_DEF"
+  "BIJ_COMPOSE" > "HOL4Base.pred_set.BIJ_COMPOSE"
+  "BIGUNION_def" > "HOL4Base.pred_set.BIGUNION_def"
+  "BIGUNION_UNION" > "HOL4Base.pred_set.BIGUNION_UNION"
+  "BIGUNION_SUBSET" > "HOL4Base.pred_set.BIGUNION_SUBSET"
+  "BIGUNION_SING" > "HOL4Base.pred_set.BIGUNION_SING"
+  "BIGUNION_INSERT" > "HOL4Base.pred_set.BIGUNION_INSERT"
+  "BIGUNION_EMPTY" > "HOL4Base.pred_set.BIGUNION_EMPTY"
+  "BIGUNION" > "HOL4Base.pred_set.BIGUNION"
+  "BIGINTER_def" > "HOL4Base.pred_set.BIGINTER_def"
+  "BIGINTER_SING" > "HOL4Base.pred_set.BIGINTER_SING"
+  "BIGINTER_INTER" > "HOL4Base.pred_set.BIGINTER_INTER"
+  "BIGINTER_INSERT" > "HOL4Base.pred_set.BIGINTER_INSERT"
+  "BIGINTER_EMPTY" > "HOL4Base.pred_set.BIGINTER_EMPTY"
+  "BIGINTER" > "HOL4Base.pred_set.BIGINTER"
+  "ABSORPTION" > "HOL4Base.pred_set.ABSORPTION"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prim_rec.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,72 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "wellfounded" > "wellfounded_primdef"
+  "measure" > "measure_primdef"
+  "SIMP_REC_REL" > "SIMP_REC_REL_def"
+  "SIMP_REC" > "SIMP_REC_def"
+  "PRIM_REC_FUN" > "PRIM_REC_FUN_def"
+  "PRIM_REC" > "PRIM_REC_def"
+  "PRE" > "PRE_def"
+
+const_maps
+  "wellfounded" > "HOL4Base.prim_rec.wellfounded"
+  "measure" > "HOL4Base.prim_rec.measure"
+  "SIMP_REC_REL" > "HOL4Base.prim_rec.SIMP_REC_REL"
+  "PRIM_REC_FUN" > "HOL4Base.prim_rec.PRIM_REC_FUN"
+  "PRIM_REC" > "HOL4Base.prim_rec.PRIM_REC"
+  "PRE" > "HOL4Base.prim_rec.PRE"
+  "<" > "op <" :: "nat => nat => bool"
+
+thm_maps
+  "wellfounded_primdef" > "HOL4Base.prim_rec.wellfounded_primdef"
+  "wellfounded_def" > "HOL4Base.prim_rec.wellfounded_def"
+  "num_Axiom_old" > "HOL4Base.prim_rec.num_Axiom_old"
+  "num_Axiom" > "HOL4Base.prim_rec.num_Axiom"
+  "measure_thm" > "HOL4Base.prim_rec.measure_thm"
+  "measure_primdef" > "HOL4Base.prim_rec.measure_primdef"
+  "measure_def" > "HOL4Base.prim_rec.measure_def"
+  "WF_measure" > "HOL4Base.prim_rec.WF_measure"
+  "WF_PRED" > "HOL4Base.prim_rec.WF_PRED"
+  "WF_LESS" > "HOL4Base.prim_rec.WF_LESS"
+  "WF_IFF_WELLFOUNDED" > "HOL4Base.prim_rec.WF_IFF_WELLFOUNDED"
+  "SUC_LESS" > "Nat.Suc_lessD"
+  "SUC_ID" > "Nat.Suc_n_not_n"
+  "SIMP_REC_THM" > "HOL4Base.prim_rec.SIMP_REC_THM"
+  "SIMP_REC_REL_def" > "HOL4Base.prim_rec.SIMP_REC_REL_def"
+  "SIMP_REC_REL_UNIQUE_RESULT" > "HOL4Base.prim_rec.SIMP_REC_REL_UNIQUE_RESULT"
+  "SIMP_REC_REL_UNIQUE" > "HOL4Base.prim_rec.SIMP_REC_REL_UNIQUE"
+  "SIMP_REC_REL" > "HOL4Base.prim_rec.SIMP_REC_REL"
+  "SIMP_REC_EXISTS" > "HOL4Base.prim_rec.SIMP_REC_EXISTS"
+  "SIMP_REC" > "HOL4Base.prim_rec.SIMP_REC"
+  "PRIM_REC_def" > "HOL4Base.prim_rec.PRIM_REC_def"
+  "PRIM_REC_THM" > "HOL4Base.prim_rec.PRIM_REC_THM"
+  "PRIM_REC_FUN_def" > "HOL4Base.prim_rec.PRIM_REC_FUN_def"
+  "PRIM_REC_FUN" > "HOL4Base.prim_rec.PRIM_REC_FUN"
+  "PRIM_REC_EQN" > "HOL4Base.prim_rec.PRIM_REC_EQN"
+  "PRIM_REC" > "HOL4Base.prim_rec.PRIM_REC"
+  "PRE_def" > "HOL4Base.prim_rec.PRE_def"
+  "PRE_DEF" > "HOL4Base.prim_rec.PRE_DEF"
+  "PRE" > "HOL4Base.prim_rec.PRE"
+  "NOT_LESS_EQ" > "HOL4Base.prim_rec.NOT_LESS_EQ"
+  "NOT_LESS_0" > "Nat.not_less0"
+  "LESS_THM" > "HOL4Base.prim_rec.LESS_THM"
+  "LESS_SUC_SUC" > "HOL4Base.prim_rec.LESS_SUC_SUC"
+  "LESS_SUC_REFL" > "Nat.lessI"
+  "LESS_SUC_IMP" > "HOL4Base.prim_rec.LESS_SUC_IMP"
+  "LESS_SUC" > "Nat.less_SucI"
+  "LESS_REFL" > "Nat.less_not_refl"
+  "LESS_NOT_EQ" > "Nat.less_not_refl3"
+  "LESS_MONO" > "Nat.Suc_mono"
+  "LESS_LEMMA2" > "HOL4Base.prim_rec.LESS_LEMMA2"
+  "LESS_LEMMA1" > "HOL4Base.prim_rec.LESS_LEMMA1"
+  "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"
+  "EQ_LESS" > "HOL4Base.prim_rec.EQ_LESS"
+  "DC" > "HOL4Base.prim_rec.DC"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prime.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,17 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "prime" > "prime_primdef"
+
+const_maps
+  "prime" > "HOL4Base.prime.prime"
+
+thm_maps
+  "prime_primdef" > "HOL4Base.prime.prime_primdef"
+  "prime_def" > "HOL4Base.prime.prime_def"
+  "NOT_PRIME_1" > "HOL4Base.prime.NOT_PRIME_1"
+  "NOT_PRIME_0" > "HOL4Base.prime.NOT_PRIME_0"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,61 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "prob" > "prob_primdef"
+  "algebra_measure" > "algebra_measure_primdef"
+  "alg_measure" > "alg_measure_primdef"
+
+const_maps
+  "prob" > "HOL4Prob.prob.prob"
+  "algebra_measure" > "HOL4Prob.prob.algebra_measure"
+
+thm_maps
+  "prob_primdef" > "HOL4Prob.prob.prob_primdef"
+  "prob_def" > "HOL4Prob.prob.prob_def"
+  "algebra_measure_primdef" > "HOL4Prob.prob.algebra_measure_primdef"
+  "algebra_measure_def" > "HOL4Prob.prob.algebra_measure_def"
+  "alg_measure_def" > "HOL4Prob.prob.alg_measure_def"
+  "X_LE_PROB" > "HOL4Prob.prob.X_LE_PROB"
+  "PROB_SUP_EXISTS2" > "HOL4Prob.prob.PROB_SUP_EXISTS2"
+  "PROB_SUP_EXISTS1" > "HOL4Prob.prob.PROB_SUP_EXISTS1"
+  "PROB_SUBSET_MONO" > "HOL4Prob.prob.PROB_SUBSET_MONO"
+  "PROB_STL" > "HOL4Prob.prob.PROB_STL"
+  "PROB_SHD" > "HOL4Prob.prob.PROB_SHD"
+  "PROB_SDROP" > "HOL4Prob.prob.PROB_SDROP"
+  "PROB_RANGE" > "HOL4Prob.prob.PROB_RANGE"
+  "PROB_POS" > "HOL4Prob.prob.PROB_POS"
+  "PROB_MAX" > "HOL4Prob.prob.PROB_MAX"
+  "PROB_LE_X" > "HOL4Prob.prob.PROB_LE_X"
+  "PROB_INTER_SHD" > "HOL4Prob.prob.PROB_INTER_SHD"
+  "PROB_INTER_HALVES" > "HOL4Prob.prob.PROB_INTER_HALVES"
+  "PROB_COMPL_LE1" > "HOL4Prob.prob.PROB_COMPL_LE1"
+  "PROB_COMPL" > "HOL4Prob.prob.PROB_COMPL"
+  "PROB_BASIC" > "HOL4Prob.prob.PROB_BASIC"
+  "PROB_ALGEBRA" > "HOL4Prob.prob.PROB_ALGEBRA"
+  "PROB_ALG" > "HOL4Prob.prob.PROB_ALG"
+  "PROB_ADDITIVE" > "HOL4Prob.prob.PROB_ADDITIVE"
+  "ALG_TWINS_MEASURE" > "HOL4Prob.prob.ALG_TWINS_MEASURE"
+  "ALG_MEASURE_TLS" > "HOL4Prob.prob.ALG_MEASURE_TLS"
+  "ALG_MEASURE_POS" > "HOL4Prob.prob.ALG_MEASURE_POS"
+  "ALG_MEASURE_COMPL" > "HOL4Prob.prob.ALG_MEASURE_COMPL"
+  "ALG_MEASURE_BASIC" > "HOL4Prob.prob.ALG_MEASURE_BASIC"
+  "ALG_MEASURE_APPEND" > "HOL4Prob.prob.ALG_MEASURE_APPEND"
+  "ALG_MEASURE_ADDITIVE" > "HOL4Prob.prob.ALG_MEASURE_ADDITIVE"
+  "ALG_CANON_PREFS_MONO" > "HOL4Prob.prob.ALG_CANON_PREFS_MONO"
+  "ALG_CANON_MONO" > "HOL4Prob.prob.ALG_CANON_MONO"
+  "ALG_CANON_MERGE_MONO" > "HOL4Prob.prob.ALG_CANON_MERGE_MONO"
+  "ALG_CANON_FIND_MONO" > "HOL4Prob.prob.ALG_CANON_FIND_MONO"
+  "ALG_CANON2_MONO" > "HOL4Prob.prob.ALG_CANON2_MONO"
+  "ALG_CANON1_MONO" > "HOL4Prob.prob.ALG_CANON1_MONO"
+  "ALGEBRA_MEASURE_RANGE" > "HOL4Prob.prob.ALGEBRA_MEASURE_RANGE"
+  "ALGEBRA_MEASURE_POS" > "HOL4Prob.prob.ALGEBRA_MEASURE_POS"
+  "ALGEBRA_MEASURE_MONO_EMBED" > "HOL4Prob.prob.ALGEBRA_MEASURE_MONO_EMBED"
+  "ALGEBRA_MEASURE_MAX" > "HOL4Prob.prob.ALGEBRA_MEASURE_MAX"
+  "ALGEBRA_MEASURE_DEF_ALT" > "HOL4Prob.prob.ALGEBRA_MEASURE_DEF_ALT"
+  "ALGEBRA_MEASURE_BASIC" > "HOL4Prob.prob.ALGEBRA_MEASURE_BASIC"
+  "ALGEBRA_CANON_MEASURE_MAX" > "HOL4Prob.prob.ALGEBRA_CANON_MEASURE_MAX"
+  "ABS_PROB" > "HOL4Prob.prob.ABS_PROB"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_algebra.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,54 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "measurable" > "measurable_primdef"
+  "algebra_embed" > "algebra_embed_primdef"
+  "alg_embed" > "alg_embed_primdef"
+
+const_maps
+  "measurable" > "HOL4Prob.prob_algebra.measurable"
+
+thm_maps
+  "measurable_primdef" > "HOL4Prob.prob_algebra.measurable_primdef"
+  "measurable_def" > "HOL4Prob.prob_algebra.measurable_def"
+  "algebra_embed_def" > "HOL4Prob.prob_algebra.algebra_embed_def"
+  "alg_embed_def" > "HOL4Prob.prob_algebra.alg_embed_def"
+  "MEASURABLE_UNION" > "HOL4Prob.prob_algebra.MEASURABLE_UNION"
+  "MEASURABLE_STL" > "HOL4Prob.prob_algebra.MEASURABLE_STL"
+  "MEASURABLE_SHD" > "HOL4Prob.prob_algebra.MEASURABLE_SHD"
+  "MEASURABLE_SDROP" > "HOL4Prob.prob_algebra.MEASURABLE_SDROP"
+  "MEASURABLE_INTER_SHD" > "HOL4Prob.prob_algebra.MEASURABLE_INTER_SHD"
+  "MEASURABLE_INTER_HALVES" > "HOL4Prob.prob_algebra.MEASURABLE_INTER_HALVES"
+  "MEASURABLE_INTER" > "HOL4Prob.prob_algebra.MEASURABLE_INTER"
+  "MEASURABLE_HALVES" > "HOL4Prob.prob_algebra.MEASURABLE_HALVES"
+  "MEASURABLE_COMPL" > "HOL4Prob.prob_algebra.MEASURABLE_COMPL"
+  "MEASURABLE_BASIC" > "HOL4Prob.prob_algebra.MEASURABLE_BASIC"
+  "MEASURABLE_ALGEBRA" > "HOL4Prob.prob_algebra.MEASURABLE_ALGEBRA"
+  "INTER_STL" > "HOL4Prob.prob_algebra.INTER_STL"
+  "HALVES_INTER" > "HOL4Prob.prob_algebra.HALVES_INTER"
+  "COMPL_SHD" > "HOL4Prob.prob_algebra.COMPL_SHD"
+  "ALG_EMBED_TWINS" > "HOL4Prob.prob_algebra.ALG_EMBED_TWINS"
+  "ALG_EMBED_PREFIX_SUBSET" > "HOL4Prob.prob_algebra.ALG_EMBED_PREFIX_SUBSET"
+  "ALG_EMBED_PREFIX" > "HOL4Prob.prob_algebra.ALG_EMBED_PREFIX"
+  "ALG_EMBED_POPULATED" > "HOL4Prob.prob_algebra.ALG_EMBED_POPULATED"
+  "ALG_EMBED_NIL" > "HOL4Prob.prob_algebra.ALG_EMBED_NIL"
+  "ALG_EMBED_BASIC" > "HOL4Prob.prob_algebra.ALG_EMBED_BASIC"
+  "ALG_CANON_REP" > "HOL4Prob.prob_algebra.ALG_CANON_REP"
+  "ALG_CANON_PREFS_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON_PREFS_EMBED"
+  "ALG_CANON_MERGE_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON_MERGE_EMBED"
+  "ALG_CANON_FIND_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON_FIND_EMBED"
+  "ALG_CANON_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON_EMBED"
+  "ALG_CANON2_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON2_EMBED"
+  "ALG_CANON1_EMBED" > "HOL4Prob.prob_algebra.ALG_CANON1_EMBED"
+  "ALGEBRA_EMBED_TLS" > "HOL4Prob.prob_algebra.ALGEBRA_EMBED_TLS"
+  "ALGEBRA_EMBED_MEM" > "HOL4Prob.prob_algebra.ALGEBRA_EMBED_MEM"
+  "ALGEBRA_EMBED_COMPL" > "HOL4Prob.prob_algebra.ALGEBRA_EMBED_COMPL"
+  "ALGEBRA_EMBED_BASIC" > "HOL4Prob.prob_algebra.ALGEBRA_EMBED_BASIC"
+  "ALGEBRA_EMBED_APPEND" > "HOL4Prob.prob_algebra.ALGEBRA_EMBED_APPEND"
+  "ALGEBRA_CANON_UNIV" > "HOL4Prob.prob_algebra.ALGEBRA_CANON_UNIV"
+  "ALGEBRA_CANON_EMBED_UNIV" > "HOL4Prob.prob_algebra.ALGEBRA_CANON_EMBED_UNIV"
+  "ALGEBRA_CANON_EMBED_EMPTY" > "HOL4Prob.prob_algebra.ALGEBRA_CANON_EMBED_EMPTY"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_canon.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,147 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "algebra_canon" > "algebra_canon_primdef"
+  "alg_twinfree" > "alg_twinfree_primdef"
+  "alg_twin" > "alg_twin_primdef"
+  "alg_sorted" > "alg_sorted_primdef"
+  "alg_prefixfree" > "alg_prefixfree_primdef"
+  "alg_order_tupled" > "alg_order_tupled_def"
+  "alg_order" > "alg_order_primdef"
+  "alg_longest" > "alg_longest_primdef"
+  "alg_canon_prefs" > "alg_canon_prefs_primdef"
+  "alg_canon_merge" > "alg_canon_merge_primdef"
+  "alg_canon_find" > "alg_canon_find_primdef"
+  "alg_canon2" > "alg_canon2_primdef"
+  "alg_canon1" > "alg_canon1_primdef"
+  "alg_canon" > "alg_canon_primdef"
+
+const_maps
+  "algebra_canon" > "HOL4Prob.prob_canon.algebra_canon"
+  "alg_twinfree" > "HOL4Prob.prob_canon.alg_twinfree"
+  "alg_twin" > "HOL4Prob.prob_canon.alg_twin"
+  "alg_sorted" > "HOL4Prob.prob_canon.alg_sorted"
+  "alg_prefixfree" > "HOL4Prob.prob_canon.alg_prefixfree"
+  "alg_order_tupled" > "HOL4Prob.prob_canon.alg_order_tupled"
+  "alg_order" > "HOL4Prob.prob_canon.alg_order"
+  "alg_longest" > "HOL4Prob.prob_canon.alg_longest"
+  "alg_canon2" > "HOL4Prob.prob_canon.alg_canon2"
+  "alg_canon1" > "HOL4Prob.prob_canon.alg_canon1"
+  "alg_canon" > "HOL4Prob.prob_canon.alg_canon"
+
+thm_maps
+  "algebra_canon_primdef" > "HOL4Prob.prob_canon.algebra_canon_primdef"
+  "algebra_canon_def" > "HOL4Prob.prob_canon.algebra_canon_def"
+  "alg_twinfree_primitive_def" > "HOL4Prob.prob_canon.alg_twinfree_primitive_def"
+  "alg_twinfree_primdef" > "HOL4Prob.prob_canon.alg_twinfree_primdef"
+  "alg_twinfree_ind" > "HOL4Prob.prob_canon.alg_twinfree_ind"
+  "alg_twinfree_def" > "HOL4Prob.prob_canon.alg_twinfree_def"
+  "alg_twin_primdef" > "HOL4Prob.prob_canon.alg_twin_primdef"
+  "alg_twin_def" > "HOL4Prob.prob_canon.alg_twin_def"
+  "alg_sorted_primitive_def" > "HOL4Prob.prob_canon.alg_sorted_primitive_def"
+  "alg_sorted_primdef" > "HOL4Prob.prob_canon.alg_sorted_primdef"
+  "alg_sorted_ind" > "HOL4Prob.prob_canon.alg_sorted_ind"
+  "alg_sorted_def" > "HOL4Prob.prob_canon.alg_sorted_def"
+  "alg_prefixfree_primitive_def" > "HOL4Prob.prob_canon.alg_prefixfree_primitive_def"
+  "alg_prefixfree_primdef" > "HOL4Prob.prob_canon.alg_prefixfree_primdef"
+  "alg_prefixfree_ind" > "HOL4Prob.prob_canon.alg_prefixfree_ind"
+  "alg_prefixfree_def" > "HOL4Prob.prob_canon.alg_prefixfree_def"
+  "alg_order_tupled_primitive_def" > "HOL4Prob.prob_canon.alg_order_tupled_primitive_def"
+  "alg_order_tupled_def" > "HOL4Prob.prob_canon.alg_order_tupled_def"
+  "alg_order_primdef" > "HOL4Prob.prob_canon.alg_order_primdef"
+  "alg_order_ind" > "HOL4Prob.prob_canon.alg_order_ind"
+  "alg_order_def" > "HOL4Prob.prob_canon.alg_order_def"
+  "alg_order_curried_def" > "HOL4Prob.prob_canon.alg_order_curried_def"
+  "alg_longest_primdef" > "HOL4Prob.prob_canon.alg_longest_primdef"
+  "alg_longest_def" > "HOL4Prob.prob_canon.alg_longest_def"
+  "alg_canon_primdef" > "HOL4Prob.prob_canon.alg_canon_primdef"
+  "alg_canon_prefs_def" > "HOL4Prob.prob_canon.alg_canon_prefs_def"
+  "alg_canon_merge_def" > "HOL4Prob.prob_canon.alg_canon_merge_def"
+  "alg_canon_find_def" > "HOL4Prob.prob_canon.alg_canon_find_def"
+  "alg_canon_def" > "HOL4Prob.prob_canon.alg_canon_def"
+  "alg_canon2_primdef" > "HOL4Prob.prob_canon.alg_canon2_primdef"
+  "alg_canon2_def" > "HOL4Prob.prob_canon.alg_canon2_def"
+  "alg_canon1_primdef" > "HOL4Prob.prob_canon.alg_canon1_primdef"
+  "alg_canon1_def" > "HOL4Prob.prob_canon.alg_canon1_def"
+  "MEM_NIL_STEP" > "HOL4Prob.prob_canon.MEM_NIL_STEP"
+  "ALG_TWIN_SING" > "HOL4Prob.prob_canon.ALG_TWIN_SING"
+  "ALG_TWIN_REDUCE" > "HOL4Prob.prob_canon.ALG_TWIN_REDUCE"
+  "ALG_TWIN_NIL" > "HOL4Prob.prob_canon.ALG_TWIN_NIL"
+  "ALG_TWIN_CONS" > "HOL4Prob.prob_canon.ALG_TWIN_CONS"
+  "ALG_TWINS_PREFIX" > "HOL4Prob.prob_canon.ALG_TWINS_PREFIX"
+  "ALG_TWINFREE_TLS" > "HOL4Prob.prob_canon.ALG_TWINFREE_TLS"
+  "ALG_TWINFREE_TL" > "HOL4Prob.prob_canon.ALG_TWINFREE_TL"
+  "ALG_TWINFREE_STEP2" > "HOL4Prob.prob_canon.ALG_TWINFREE_STEP2"
+  "ALG_TWINFREE_STEP1" > "HOL4Prob.prob_canon.ALG_TWINFREE_STEP1"
+  "ALG_TWINFREE_STEP" > "HOL4Prob.prob_canon.ALG_TWINFREE_STEP"
+  "ALG_SORTED_TLS" > "HOL4Prob.prob_canon.ALG_SORTED_TLS"
+  "ALG_SORTED_TL" > "HOL4Prob.prob_canon.ALG_SORTED_TL"
+  "ALG_SORTED_STEP" > "HOL4Prob.prob_canon.ALG_SORTED_STEP"
+  "ALG_SORTED_PREFIXFREE_MEM_NIL" > "HOL4Prob.prob_canon.ALG_SORTED_PREFIXFREE_MEM_NIL"
+  "ALG_SORTED_PREFIXFREE_EQUALITY" > "HOL4Prob.prob_canon.ALG_SORTED_PREFIXFREE_EQUALITY"
+  "ALG_SORTED_MONO" > "HOL4Prob.prob_canon.ALG_SORTED_MONO"
+  "ALG_SORTED_MIN" > "HOL4Prob.prob_canon.ALG_SORTED_MIN"
+  "ALG_SORTED_FILTER" > "HOL4Prob.prob_canon.ALG_SORTED_FILTER"
+  "ALG_SORTED_DEF_ALT" > "HOL4Prob.prob_canon.ALG_SORTED_DEF_ALT"
+  "ALG_SORTED_APPEND" > "HOL4Prob.prob_canon.ALG_SORTED_APPEND"
+  "ALG_PREFIXFREE_TLS" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_TLS"
+  "ALG_PREFIXFREE_TL" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_TL"
+  "ALG_PREFIXFREE_STEP" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_STEP"
+  "ALG_PREFIXFREE_MONO" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_MONO"
+  "ALG_PREFIXFREE_FILTER" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_FILTER"
+  "ALG_PREFIXFREE_ELT" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_ELT"
+  "ALG_PREFIXFREE_APPEND" > "HOL4Prob.prob_canon.ALG_PREFIXFREE_APPEND"
+  "ALG_ORDER_TRANS" > "HOL4Prob.prob_canon.ALG_ORDER_TRANS"
+  "ALG_ORDER_TOTAL" > "HOL4Prob.prob_canon.ALG_ORDER_TOTAL"
+  "ALG_ORDER_SNOC" > "HOL4Prob.prob_canon.ALG_ORDER_SNOC"
+  "ALG_ORDER_REFL" > "HOL4Prob.prob_canon.ALG_ORDER_REFL"
+  "ALG_ORDER_PREFIX_TRANS" > "HOL4Prob.prob_canon.ALG_ORDER_PREFIX_TRANS"
+  "ALG_ORDER_PREFIX_MONO" > "HOL4Prob.prob_canon.ALG_ORDER_PREFIX_MONO"
+  "ALG_ORDER_PREFIX_ANTI" > "HOL4Prob.prob_canon.ALG_ORDER_PREFIX_ANTI"
+  "ALG_ORDER_PREFIX" > "HOL4Prob.prob_canon.ALG_ORDER_PREFIX"
+  "ALG_ORDER_NIL" > "HOL4Prob.prob_canon.ALG_ORDER_NIL"
+  "ALG_ORDER_ANTISYM" > "HOL4Prob.prob_canon.ALG_ORDER_ANTISYM"
+  "ALG_LONGEST_TLS" > "HOL4Prob.prob_canon.ALG_LONGEST_TLS"
+  "ALG_LONGEST_TL" > "HOL4Prob.prob_canon.ALG_LONGEST_TL"
+  "ALG_LONGEST_HD" > "HOL4Prob.prob_canon.ALG_LONGEST_HD"
+  "ALG_LONGEST_APPEND" > "HOL4Prob.prob_canon.ALG_LONGEST_APPEND"
+  "ALG_CANON_SORTED_PREFIXFREE_TWINFREE" > "HOL4Prob.prob_canon.ALG_CANON_SORTED_PREFIXFREE_TWINFREE"
+  "ALG_CANON_PREFS_SORTED" > "HOL4Prob.prob_canon.ALG_CANON_PREFS_SORTED"
+  "ALG_CANON_PREFS_PREFIXFREE" > "HOL4Prob.prob_canon.ALG_CANON_PREFS_PREFIXFREE"
+  "ALG_CANON_PREFS_HD" > "HOL4Prob.prob_canon.ALG_CANON_PREFS_HD"
+  "ALG_CANON_PREFS_DELETES" > "HOL4Prob.prob_canon.ALG_CANON_PREFS_DELETES"
+  "ALG_CANON_PREFS_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON_PREFS_CONSTANT"
+  "ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE" > "HOL4Prob.prob_canon.ALG_CANON_MERGE_SORTED_PREFIXFREE_TWINFREE"
+  "ALG_CANON_MERGE_SHORTENS" > "HOL4Prob.prob_canon.ALG_CANON_MERGE_SHORTENS"
+  "ALG_CANON_MERGE_PREFIXFREE_PRESERVE" > "HOL4Prob.prob_canon.ALG_CANON_MERGE_PREFIXFREE_PRESERVE"
+  "ALG_CANON_MERGE_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON_MERGE_CONSTANT"
+  "ALG_CANON_IDEMPOT" > "HOL4Prob.prob_canon.ALG_CANON_IDEMPOT"
+  "ALG_CANON_FIND_SORTED" > "HOL4Prob.prob_canon.ALG_CANON_FIND_SORTED"
+  "ALG_CANON_FIND_PREFIXFREE" > "HOL4Prob.prob_canon.ALG_CANON_FIND_PREFIXFREE"
+  "ALG_CANON_FIND_HD" > "HOL4Prob.prob_canon.ALG_CANON_FIND_HD"
+  "ALG_CANON_FIND_DELETES" > "HOL4Prob.prob_canon.ALG_CANON_FIND_DELETES"
+  "ALG_CANON_FIND_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON_FIND_CONSTANT"
+  "ALG_CANON_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON_CONSTANT"
+  "ALG_CANON_BASIC" > "HOL4Prob.prob_canon.ALG_CANON_BASIC"
+  "ALG_CANON2_SORTED_PREFIXFREE_TWINFREE" > "HOL4Prob.prob_canon.ALG_CANON2_SORTED_PREFIXFREE_TWINFREE"
+  "ALG_CANON2_SHORTENS" > "HOL4Prob.prob_canon.ALG_CANON2_SHORTENS"
+  "ALG_CANON2_PREFIXFREE_PRESERVE" > "HOL4Prob.prob_canon.ALG_CANON2_PREFIXFREE_PRESERVE"
+  "ALG_CANON2_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON2_CONSTANT"
+  "ALG_CANON1_SORTED" > "HOL4Prob.prob_canon.ALG_CANON1_SORTED"
+  "ALG_CANON1_PREFIXFREE" > "HOL4Prob.prob_canon.ALG_CANON1_PREFIXFREE"
+  "ALG_CANON1_CONSTANT" > "HOL4Prob.prob_canon.ALG_CANON1_CONSTANT"
+  "ALGEBRA_CANON_TLS" > "HOL4Prob.prob_canon.ALGEBRA_CANON_TLS"
+  "ALGEBRA_CANON_TL" > "HOL4Prob.prob_canon.ALGEBRA_CANON_TL"
+  "ALGEBRA_CANON_STEP2" > "HOL4Prob.prob_canon.ALGEBRA_CANON_STEP2"
+  "ALGEBRA_CANON_STEP1" > "HOL4Prob.prob_canon.ALGEBRA_CANON_STEP1"
+  "ALGEBRA_CANON_STEP" > "HOL4Prob.prob_canon.ALGEBRA_CANON_STEP"
+  "ALGEBRA_CANON_NIL_MEM" > "HOL4Prob.prob_canon.ALGEBRA_CANON_NIL_MEM"
+  "ALGEBRA_CANON_INDUCTION" > "HOL4Prob.prob_canon.ALGEBRA_CANON_INDUCTION"
+  "ALGEBRA_CANON_DEF_ALT" > "HOL4Prob.prob_canon.ALGEBRA_CANON_DEF_ALT"
+  "ALGEBRA_CANON_CASES_THM" > "HOL4Prob.prob_canon.ALGEBRA_CANON_CASES_THM"
+  "ALGEBRA_CANON_CASES" > "HOL4Prob.prob_canon.ALGEBRA_CANON_CASES"
+  "ALGEBRA_CANON_BASIC" > "HOL4Prob.prob_canon.ALGEBRA_CANON_BASIC"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_extra.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,96 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "inf" > "inf_primdef"
+
+const_maps
+  "inf" > "HOL4Prob.prob_extra.inf"
+  "COMPL" > "HOL4Base.pred_set.COMPL"
+
+thm_maps
+  "inf_primdef" > "HOL4Prob.prob_extra.inf_primdef"
+  "inf_def" > "HOL4Prob.prob_extra.inf_def"
+  "X_HALF_HALF" > "HOL4Prob.prob_extra.X_HALF_HALF"
+  "UNION_DISJOINT_SPLIT" > "HOL4Prob.prob_extra.UNION_DISJOINT_SPLIT"
+  "UNION_DEF_ALT" > "HOL4Prob.prob_extra.UNION_DEF_ALT"
+  "SUBSET_EQ_DECOMP" > "HOL4Base.pred_set.SUBSET_ANTISYM"
+  "SUBSET_EQ" > "HOL4Prob.prob_extra.SUBSET_EQ"
+  "SET_EQ_EXT" > "HOL4Base.pred_set.EXTENSION"
+  "REAL_X_LE_SUP" > "HOL4Prob.prob_extra.REAL_X_LE_SUP"
+  "REAL_SUP_MAX" > "HOL4Prob.prob_extra.REAL_SUP_MAX"
+  "REAL_SUP_LE_X" > "HOL4Prob.prob_extra.REAL_SUP_LE_X"
+  "REAL_SUP_EXISTS_UNIQUE" > "HOL4Prob.prob_extra.REAL_SUP_EXISTS_UNIQUE"
+  "REAL_POW" > "RealPow.realpow_real_of_nat"
+  "REAL_LE_INV_LE" > "Ring_and_Field.le_imp_inverse_le"
+  "REAL_LE_EQ" > "Set.basic_trans_rules_26"
+  "REAL_INVINV_ALL" > "Ring_and_Field.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"
+  "POW_HALF_POS" > "HOL4Prob.prob_extra.POW_HALF_POS"
+  "POW_HALF_MONO" > "HOL4Prob.prob_extra.POW_HALF_MONO"
+  "POW_HALF_EXP" > "HOL4Prob.prob_extra.POW_HALF_EXP"
+  "ONE_MINUS_HALF" > "HOL4Prob.prob_extra.ONE_MINUS_HALF"
+  "MOD_TWO" > "HOL4Prob.prob_extra.MOD_TWO"
+  "MEM_NIL_MAP_CONS" > "HOL4Prob.prob_extra.MEM_NIL_MAP_CONS"
+  "MEM_NIL" > "HOL4Prob.prob_extra.MEM_NIL"
+  "MEM_FILTER" > "HOL4Prob.prob_extra.MEM_FILTER"
+  "MAP_MEM" > "HOL4Prob.prob_extra.MAP_MEM"
+  "MAP_ID" > "List.map_ident"
+  "LENGTH_FILTER" > "List.length_filter"
+  "LAST_MEM" > "HOL4Prob.prob_extra.LAST_MEM"
+  "LAST_MAP_CONS" > "HOL4Prob.prob_extra.LAST_MAP_CONS"
+  "IS_PREFIX_TRANS" > "HOL4Prob.prob_extra.IS_PREFIX_TRANS"
+  "IS_PREFIX_SNOC" > "HOL4Prob.prob_extra.IS_PREFIX_SNOC"
+  "IS_PREFIX_REFL" > "HOL4Prob.prob_extra.IS_PREFIX_REFL"
+  "IS_PREFIX_NIL" > "HOL4Prob.prob_extra.IS_PREFIX_NIL"
+  "IS_PREFIX_LENGTH_ANTI" > "HOL4Prob.prob_extra.IS_PREFIX_LENGTH_ANTI"
+  "IS_PREFIX_LENGTH" > "HOL4Prob.prob_extra.IS_PREFIX_LENGTH"
+  "IS_PREFIX_BUTLAST" > "HOL4Prob.prob_extra.IS_PREFIX_BUTLAST"
+  "IS_PREFIX_ANTISYM" > "HOL4Prob.prob_extra.IS_PREFIX_ANTISYM"
+  "IN_EMPTY" > "HOL4Base.pred_set.NOT_IN_EMPTY"
+  "IN_COMPL" > "HOL4Base.pred_set.IN_COMPL"
+  "INV_SUC_POS" > "HOL4Prob.prob_extra.INV_SUC_POS"
+  "INV_SUC_MAX" > "HOL4Prob.prob_extra.INV_SUC_MAX"
+  "INV_SUC" > "HOL4Prob.prob_extra.INV_SUC"
+  "INTER_UNION_RDISTRIB" > "HOL4Prob.prob_extra.INTER_UNION_RDISTRIB"
+  "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_LT_1" > "HOL4Prob.prob_extra.HALF_LT_1"
+  "HALF_CANCEL" > "HOL4Prob.prob_extra.HALF_CANCEL"
+  "GSPEC_DEF_ALT" > "HOL4Prob.prob_extra.GSPEC_DEF_ALT"
+  "FOLDR_MAP" > "HOL4Prob.prob_extra.FOLDR_MAP"
+  "FILTER_TRUE" > "HOL4Prob.prob_extra.FILTER_TRUE"
+  "FILTER_OUT_ELT" > "HOL4Prob.prob_extra.FILTER_OUT_ELT"
+  "FILTER_MEM" > "HOL4Prob.prob_extra.FILTER_MEM"
+  "FILTER_FALSE" > "HOL4Prob.prob_extra.FILTER_FALSE"
+  "EXP_DIV_TWO" > "HOL4Prob.prob_extra.EXP_DIV_TWO"
+  "EXISTS_LONGEST" > "HOL4Prob.prob_extra.EXISTS_LONGEST"
+  "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.expand_fun_eq"
+  "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"
+  "DIV_TWO_EXP" > "HOL4Prob.prob_extra.DIV_TWO_EXP"
+  "DIV_TWO_CANCEL" > "HOL4Prob.prob_extra.DIV_TWO_CANCEL"
+  "DIV_TWO_BASIC" > "HOL4Prob.prob_extra.DIV_TWO_BASIC"
+  "DIV_TWO" > "HOL4Prob.prob_extra.DIV_TWO"
+  "DIV_THEN_MULT" > "HOL4Prob.prob_extra.DIV_THEN_MULT"
+  "DIVISION_TWO" > "HOL4Prob.prob_extra.DIVISION_TWO"
+  "COMPL_def" > "HOL4Base.pred_set.COMPL_DEF"
+  "COMPL_SPLITS" > "HOL4Base.pred_set.COMPL_SPLITS"
+  "COMPL_COMPL" > "HOL4Base.pred_set.COMPL_COMPL"
+  "COMPL_CLAUSES" > "HOL4Base.pred_set.COMPL_CLAUSES"
+  "BOOL_BOOL_CASES_THM" > "HOL4Prob.prob_extra.BOOL_BOOL_CASES_THM"
+  "BOOL_BOOL_CASES" > "HOL4Base.bool.BOOL_FUN_INDUCT"
+  "APPEND_MEM" > "HOL4Base.list.MEM_APPEND"
+  "ABS_UNIT_INTERVAL" > "HOL4Prob.prob_extra.ABS_UNIT_INTERVAL"
+  "ABS_BETWEEN_LE" > "HOL4Prob.prob_extra.ABS_BETWEEN_LE"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_indep.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,56 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "indep_set" > "indep_set_primdef"
+  "indep" > "indep_primdef"
+  "alg_cover_set" > "alg_cover_set_primdef"
+  "alg_cover" > "alg_cover_primdef"
+
+const_maps
+  "indep_set" > "HOL4Prob.prob_indep.indep_set"
+  "indep" > "HOL4Prob.prob_indep.indep"
+  "alg_cover_set" > "HOL4Prob.prob_indep.alg_cover_set"
+  "alg_cover" > "HOL4Prob.prob_indep.alg_cover"
+
+thm_maps
+  "indep_set_primdef" > "HOL4Prob.prob_indep.indep_set_primdef"
+  "indep_set_def" > "HOL4Prob.prob_indep.indep_set_def"
+  "indep_primdef" > "HOL4Prob.prob_indep.indep_primdef"
+  "indep_def" > "HOL4Prob.prob_indep.indep_def"
+  "alg_cover_set_primdef" > "HOL4Prob.prob_indep.alg_cover_set_primdef"
+  "alg_cover_set_def" > "HOL4Prob.prob_indep.alg_cover_set_def"
+  "alg_cover_primdef" > "HOL4Prob.prob_indep.alg_cover_primdef"
+  "alg_cover_def" > "HOL4Prob.prob_indep.alg_cover_def"
+  "PROB_INDEP_BOUND" > "HOL4Prob.prob_indep.PROB_INDEP_BOUND"
+  "MAP_CONS_TL_FILTER" > "HOL4Prob.prob_indep.MAP_CONS_TL_FILTER"
+  "INDEP_UNIT" > "HOL4Prob.prob_indep.INDEP_UNIT"
+  "INDEP_SET_SYM" > "HOL4Prob.prob_indep.INDEP_SET_SYM"
+  "INDEP_SET_LIST" > "HOL4Prob.prob_indep.INDEP_SET_LIST"
+  "INDEP_SET_DISJOINT_DECOMP" > "HOL4Prob.prob_indep.INDEP_SET_DISJOINT_DECOMP"
+  "INDEP_SET_BASIC" > "HOL4Prob.prob_indep.INDEP_SET_BASIC"
+  "INDEP_SDEST" > "HOL4Prob.prob_indep.INDEP_SDEST"
+  "INDEP_PROB" > "HOL4Prob.prob_indep.INDEP_PROB"
+  "INDEP_MEASURABLE2" > "HOL4Prob.prob_indep.INDEP_MEASURABLE2"
+  "INDEP_MEASURABLE1" > "HOL4Prob.prob_indep.INDEP_MEASURABLE1"
+  "INDEP_INDEP_SET_LEMMA" > "HOL4Prob.prob_indep.INDEP_INDEP_SET_LEMMA"
+  "INDEP_INDEP_SET" > "HOL4Prob.prob_indep.INDEP_INDEP_SET"
+  "INDEP_BIND_SDEST" > "HOL4Prob.prob_indep.INDEP_BIND_SDEST"
+  "INDEP_BIND" > "HOL4Prob.prob_indep.INDEP_BIND"
+  "BIND_STEP" > "HOL4Prob.prob_indep.BIND_STEP"
+  "ALG_COVER_WELL_DEFINED" > "HOL4Prob.prob_indep.ALG_COVER_WELL_DEFINED"
+  "ALG_COVER_UNIV" > "HOL4Prob.prob_indep.ALG_COVER_UNIV"
+  "ALG_COVER_UNIQUE" > "HOL4Prob.prob_indep.ALG_COVER_UNIQUE"
+  "ALG_COVER_TAIL_STEP" > "HOL4Prob.prob_indep.ALG_COVER_TAIL_STEP"
+  "ALG_COVER_TAIL_PROB" > "HOL4Prob.prob_indep.ALG_COVER_TAIL_PROB"
+  "ALG_COVER_TAIL_MEASURABLE" > "HOL4Prob.prob_indep.ALG_COVER_TAIL_MEASURABLE"
+  "ALG_COVER_STEP" > "HOL4Prob.prob_indep.ALG_COVER_STEP"
+  "ALG_COVER_SET_INDUCTION" > "HOL4Prob.prob_indep.ALG_COVER_SET_INDUCTION"
+  "ALG_COVER_SET_CASES_THM" > "HOL4Prob.prob_indep.ALG_COVER_SET_CASES_THM"
+  "ALG_COVER_SET_CASES" > "HOL4Prob.prob_indep.ALG_COVER_SET_CASES"
+  "ALG_COVER_SET_BASIC" > "HOL4Prob.prob_indep.ALG_COVER_SET_BASIC"
+  "ALG_COVER_HEAD" > "HOL4Prob.prob_indep.ALG_COVER_HEAD"
+  "ALG_COVER_EXISTS_UNIQUE" > "HOL4Prob.prob_indep.ALG_COVER_EXISTS_UNIQUE"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_pseudo.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,26 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "pseudo_linear_tl" > "pseudo_linear_tl_primdef"
+  "pseudo_linear_hd" > "pseudo_linear_hd_primdef"
+  "pseudo_linear1" > "pseudo_linear1_primdef"
+  "pseudo" > "pseudo_primdef"
+
+const_maps
+  "pseudo_linear_tl" > "HOL4Prob.prob_pseudo.pseudo_linear_tl"
+  "pseudo_linear_hd" > "HOL4Prob.prob_pseudo.pseudo_linear_hd"
+  "pseudo" > "HOL4Prob.prob_pseudo.pseudo"
+
+thm_maps
+  "pseudo_primdef" > "HOL4Prob.prob_pseudo.pseudo_primdef"
+  "pseudo_linear_tl_primdef" > "HOL4Prob.prob_pseudo.pseudo_linear_tl_primdef"
+  "pseudo_linear_tl_def" > "HOL4Prob.prob_pseudo.pseudo_linear_tl_def"
+  "pseudo_linear_hd_primdef" > "HOL4Prob.prob_pseudo.pseudo_linear_hd_primdef"
+  "pseudo_linear_hd_def" > "HOL4Prob.prob_pseudo.pseudo_linear_hd_def"
+  "pseudo_linear1_def" > "HOL4Prob.prob_pseudo.pseudo_linear1_def"
+  "pseudo_def" > "HOL4Prob.prob_pseudo.pseudo_def"
+  "PSEUDO_LINEAR1_EXECUTE" > "HOL4Prob.prob_pseudo.PSEUDO_LINEAR1_EXECUTE"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/prob_uniform.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,57 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "uniform_tupled" > "uniform_tupled_def"
+  "uniform" > "uniform_primdef"
+  "unif_tupled" > "unif_tupled_def"
+  "unif_bound" > "unif_bound_primdef"
+  "unif" > "unif_primdef"
+
+const_maps
+  "uniform_tupled" > "HOL4Prob.prob_uniform.uniform_tupled"
+  "uniform" > "HOL4Prob.prob_uniform.uniform"
+  "unif_tupled" > "HOL4Prob.prob_uniform.unif_tupled"
+  "unif_bound" > "HOL4Prob.prob_uniform.unif_bound"
+  "unif" > "HOL4Prob.prob_uniform.unif"
+
+thm_maps
+  "uniform_tupled_primitive_def" > "HOL4Prob.prob_uniform.uniform_tupled_primitive_def"
+  "uniform_tupled_def" > "HOL4Prob.prob_uniform.uniform_tupled_def"
+  "uniform_primdef" > "HOL4Prob.prob_uniform.uniform_primdef"
+  "uniform_ind" > "HOL4Prob.prob_uniform.uniform_ind"
+  "uniform_def" > "HOL4Prob.prob_uniform.uniform_def"
+  "uniform_curried_def" > "HOL4Prob.prob_uniform.uniform_curried_def"
+  "unif_tupled_primitive_def" > "HOL4Prob.prob_uniform.unif_tupled_primitive_def"
+  "unif_tupled_def" > "HOL4Prob.prob_uniform.unif_tupled_def"
+  "unif_primdef" > "HOL4Prob.prob_uniform.unif_primdef"
+  "unif_ind" > "HOL4Prob.prob_uniform.unif_ind"
+  "unif_def" > "HOL4Prob.prob_uniform.unif_def"
+  "unif_curried_def" > "HOL4Prob.prob_uniform.unif_curried_def"
+  "unif_bound_primitive_def" > "HOL4Prob.prob_uniform.unif_bound_primitive_def"
+  "unif_bound_primdef" > "HOL4Prob.prob_uniform.unif_bound_primdef"
+  "unif_bound_ind" > "HOL4Prob.prob_uniform.unif_bound_ind"
+  "unif_bound_def" > "HOL4Prob.prob_uniform.unif_bound_def"
+  "UNIF_RANGE" > "HOL4Prob.prob_uniform.UNIF_RANGE"
+  "UNIF_DEF_MONAD" > "HOL4Prob.prob_uniform.UNIF_DEF_MONAD"
+  "UNIF_BOUND_UPPER_SUC" > "HOL4Prob.prob_uniform.UNIF_BOUND_UPPER_SUC"
+  "UNIF_BOUND_UPPER" > "HOL4Prob.prob_uniform.UNIF_BOUND_UPPER"
+  "UNIF_BOUND_LOWER_SUC" > "HOL4Prob.prob_uniform.UNIF_BOUND_LOWER_SUC"
+  "UNIF_BOUND_LOWER" > "HOL4Prob.prob_uniform.UNIF_BOUND_LOWER"
+  "UNIFORM_RANGE" > "HOL4Prob.prob_uniform.UNIFORM_RANGE"
+  "UNIFORM_DEF_MONAD" > "HOL4Prob.prob_uniform.UNIFORM_DEF_MONAD"
+  "SUC_DIV_TWO_ZERO" > "HOL4Prob.prob_uniform.SUC_DIV_TWO_ZERO"
+  "PROB_UNIF_PAIR" > "HOL4Prob.prob_uniform.PROB_UNIF_PAIR"
+  "PROB_UNIF_GOOD" > "HOL4Prob.prob_uniform.PROB_UNIF_GOOD"
+  "PROB_UNIF_BOUND" > "HOL4Prob.prob_uniform.PROB_UNIF_BOUND"
+  "PROB_UNIFORM_UPPER_BOUND" > "HOL4Prob.prob_uniform.PROB_UNIFORM_UPPER_BOUND"
+  "PROB_UNIFORM_SUC" > "HOL4Prob.prob_uniform.PROB_UNIFORM_SUC"
+  "PROB_UNIFORM_PAIR_SUC" > "HOL4Prob.prob_uniform.PROB_UNIFORM_PAIR_SUC"
+  "PROB_UNIFORM_LOWER_BOUND" > "HOL4Prob.prob_uniform.PROB_UNIFORM_LOWER_BOUND"
+  "PROB_UNIFORM" > "HOL4Prob.prob_uniform.PROB_UNIFORM"
+  "PROB_UNIF" > "HOL4Prob.prob_uniform.PROB_UNIF"
+  "INDEP_UNIFORM" > "HOL4Prob.prob_uniform.INDEP_UNIFORM"
+  "INDEP_UNIF" > "HOL4Prob.prob_uniform.INDEP_UNIF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/real.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,359 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "sup" > "sup_def"
+  "sumc" > "sumc_def"
+  "sum" > "sum_def"
+
+const_maps
+  "sup" > "HOL4Real.real.sup"
+  "sum" > "HOL4Real.real.sum"
+  "real_sub" > "op -" :: "real => real => real"
+  "real_of_num" > "RealDef.real" :: "nat => real"
+  "real_lte" > "op <=" :: "real => real => bool"
+  "real_gt" > "HOL4Compat.real_gt"
+  "real_ge" > "HOL4Compat.real_ge"
+  "pow" > "Nat.power" :: "real => nat => real"
+  "abs" > "HOL.abs" :: "real => real"
+  "/" > "HOL.divide" :: "real => real => real"
+
+thm_maps
+  "sup_def" > "HOL4Real.real.sup_def"
+  "sup" > "HOL4Real.real.sup"
+  "sumc" > "HOL4Real.real.sumc"
+  "sum_def" > "HOL4Real.real.sum_def"
+  "sum" > "HOL4Real.real.sum"
+  "real_sub" > "Ring_and_Field.compare_rls_1"
+  "real_of_num" > "HOL4Compat.real_of_num"
+  "real_lte" > "HOL4Compat.real_lte"
+  "real_lt" > "HOL.linorder_not_le"
+  "real_gt" > "HOL4Compat.real_gt"
+  "real_ge" > "HOL4Compat.real_ge"
+  "real_div" > "Ring_and_Field.field.divide_inverse"
+  "pow" > "HOL4Compat.pow"
+  "abs" > "HOL4Compat.abs"
+  "SUP_LEMMA3" > "HOL4Real.real.SUP_LEMMA3"
+  "SUP_LEMMA2" > "HOL4Real.real.SUP_LEMMA2"
+  "SUP_LEMMA1" > "HOL4Real.real.SUP_LEMMA1"
+  "SUM_ZERO" > "HOL4Real.real.SUM_ZERO"
+  "SUM_TWO" > "HOL4Real.real.SUM_TWO"
+  "SUM_SUBST" > "HOL4Real.real.SUM_SUBST"
+  "SUM_SUB" > "HOL4Real.real.SUM_SUB"
+  "SUM_REINDEX" > "HOL4Real.real.SUM_REINDEX"
+  "SUM_POS_GEN" > "HOL4Real.real.SUM_POS_GEN"
+  "SUM_POS" > "HOL4Real.real.SUM_POS"
+  "SUM_PERMUTE_0" > "HOL4Real.real.SUM_PERMUTE_0"
+  "SUM_OFFSET" > "HOL4Real.real.SUM_OFFSET"
+  "SUM_NSUB" > "HOL4Real.real.SUM_NSUB"
+  "SUM_NEG" > "HOL4Real.real.SUM_NEG"
+  "SUM_LE" > "HOL4Real.real.SUM_LE"
+  "SUM_GROUP" > "HOL4Real.real.SUM_GROUP"
+  "SUM_EQ" > "HOL4Real.real.SUM_EQ"
+  "SUM_DIFF" > "HOL4Real.real.SUM_DIFF"
+  "SUM_DEF" > "HOL4Real.real.SUM_DEF"
+  "SUM_CMUL" > "HOL4Real.real.SUM_CMUL"
+  "SUM_CANCEL" > "HOL4Real.real.SUM_CANCEL"
+  "SUM_BOUND" > "HOL4Real.real.SUM_BOUND"
+  "SUM_ADD" > "HOL4Real.real.SUM_ADD"
+  "SUM_ABS_LE" > "HOL4Real.real.SUM_ABS_LE"
+  "SUM_ABS" > "HOL4Real.real.SUM_ABS"
+  "SUM_2" > "HOL4Real.real.SUM_2"
+  "SUM_1" > "HOL4Real.real.SUM_1"
+  "SUM_0" > "HOL4Real.real.SUM_0"
+  "SETOK_LE_LT" > "HOL4Real.real.SETOK_LE_LT"
+  "REAL_SUP_UBOUND_LE" > "HOL4Real.real.REAL_SUP_UBOUND_LE"
+  "REAL_SUP_UBOUND" > "HOL4Real.real.REAL_SUP_UBOUND"
+  "REAL_SUP_SOMEPOS" > "HOL4Real.real.REAL_SUP_SOMEPOS"
+  "REAL_SUP_LE" > "HOL4Real.real.REAL_SUP_LE"
+  "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_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" > "Ring_and_Field.diff_0_right"
+  "REAL_SUB_RNEG" > "Ring_and_Field.diff_minus_eq_add"
+  "REAL_SUB_REFL" > "Ring_and_Field.diff_self"
+  "REAL_SUB_RDISTRIB" > "Ring_and_Field.left_diff_distrib"
+  "REAL_SUB_NEG2" > "HOL4Real.real.REAL_SUB_NEG2"
+  "REAL_SUB_LZERO" > "Ring_and_Field.diff_0"
+  "REAL_SUB_LT" > "RealDef.real_0_less_diff_iff"
+  "REAL_SUB_LNEG" > "HOL4Real.real.REAL_SUB_LNEG"
+  "REAL_SUB_LE" > "RealDef.real_0_le_diff_iff"
+  "REAL_SUB_LDISTRIB" > "Ring_and_Field.right_diff_distrib"
+  "REAL_SUB_INV2" > "HOL4Real.real.REAL_SUB_INV2"
+  "REAL_SUB_ADD2" > "Ring_and_Field.add_minus_self_left"
+  "REAL_SUB_ADD" > "Ring_and_Field.minus_add_self"
+  "REAL_SUB_ABS" > "HOL4Real.real.REAL_SUB_ABS"
+  "REAL_SUB_0" > "Ring_and_Field.eq_iff_diff_eq_0"
+  "REAL_RNEG_UNIQ" > "RealDef.real_add_eq_0_iff"
+  "REAL_RINV_UNIQ" > "RealDef.real_inverse_unique"
+  "REAL_RDISTRIB" > "Ring_and_Field.ring_distrib_2"
+  "REAL_POW_POW" > "Power.power_mult"
+  "REAL_POW_MONO_LT" > "Power.power_strict_increasing"
+  "REAL_POW_LT2" > "HOL4Real.real.REAL_POW_LT2"
+  "REAL_POW_LT" > "Power.zero_less_power"
+  "REAL_POW_INV" > "Power.power_inverse"
+  "REAL_POW_DIV" > "Power.power_divide"
+  "REAL_POW_ADD" > "Power.power_add"
+  "REAL_POW2_ABS" > "NatBin.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" > "Ring_and_Field.divide_1"
+  "REAL_OF_NUM_SUC" > "RealDef.real_of_nat_Suc"
+  "REAL_OF_NUM_POW" > "RealPow.realpow_real_of_nat"
+  "REAL_OF_NUM_MUL" > "RealDef.real_of_nat_mult"
+  "REAL_OF_NUM_LE" > "RealDef.real_of_nat_le_iff"
+  "REAL_OF_NUM_EQ" > "RealDef.real_of_nat_inject"
+  "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" > "HOL.linorder_not_le"
+  "REAL_NEG_SUB" > "Ring_and_Field.minus_diff_eq"
+  "REAL_NEG_RMUL" > "Ring_and_Field.minus_mult_right"
+  "REAL_NEG_NEG" > "Ring_and_Field.minus_minus"
+  "REAL_NEG_MUL2" > "Ring_and_Field.minus_mult_minus"
+  "REAL_NEG_MINUS1" > "HOL4Real.real.REAL_NEG_MINUS1"
+  "REAL_NEG_LT0" > "Ring_and_Field.neg_less_0_iff_less"
+  "REAL_NEG_LMUL" > "Ring_and_Field.minus_mult_left"
+  "REAL_NEG_LE0" > "Ring_and_Field.neg_le_0_iff_le"
+  "REAL_NEG_INV" > "Ring_and_Field.nonzero_inverse_minus_eq"
+  "REAL_NEG_GT0" > "Ring_and_Field.neg_0_less_iff_less"
+  "REAL_NEG_GE0" > "Ring_and_Field.neg_0_le_iff_le"
+  "REAL_NEG_EQ0" > "Ring_and_Field.neg_equal_0_iff_equal"
+  "REAL_NEG_EQ" > "HOL4Real.real.REAL_NEG_EQ"
+  "REAL_NEG_ADD" > "Ring_and_Field.minus_add_distrib"
+  "REAL_NEG_0" > "Ring_and_Field.minus_zero"
+  "REAL_NEGNEG" > "Ring_and_Field.minus_minus"
+  "REAL_MUL_SYM" > "Ring_and_Field.mult_ac_2"
+  "REAL_MUL_RZERO" > "Ring_and_Field.mult_zero_right"
+  "REAL_MUL_RNEG" > "Ring_and_Field.minus_mult_right"
+  "REAL_MUL_RINV" > "Ring_and_Field.right_inverse"
+  "REAL_MUL_RID" > "Ring_and_Field.mult_1_right"
+  "REAL_MUL_LZERO" > "Ring_and_Field.mult_zero_left"
+  "REAL_MUL_LNEG" > "Ring_and_Field.minus_mult_left"
+  "REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV"
+  "REAL_MUL_LID" > "Ring_and_Field.almost_semiring.mult_1"
+  "REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC"
+  "REAL_MUL" > "RealDef.real_of_nat_mult"
+  "REAL_MIDDLE2" > "HOL4Real.real.REAL_MIDDLE2"
+  "REAL_MIDDLE1" > "HOL4Real.real.REAL_MIDDLE1"
+  "REAL_MEAN" > "Ring_and_Field.dense"
+  "REAL_LT_TRANS" > "Set.basic_trans_rules_21"
+  "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
+  "REAL_LT_SUB_RADD" > "Ring_and_Field.compare_rls_6"
+  "REAL_LT_SUB_LADD" > "Ring_and_Field.compare_rls_7"
+  "REAL_LT_RMUL_IMP" > "Ring_and_Field.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" > "HOL.order_less_irrefl"
+  "REAL_LT_RDIV_EQ" > "Ring_and_Field.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" > "Ring_and_Field.add_less_cancel_right"
+  "REAL_LT_NZ" > "HOL4Real.real.REAL_LT_NZ"
+  "REAL_LT_NEGTOTAL" > "HOL4Real.real.REAL_LT_NEGTOTAL"
+  "REAL_LT_NEG" > "Ring_and_Field.neg_less_iff_less"
+  "REAL_LT_MULTIPLE" > "HOL4Real.real.REAL_LT_MULTIPLE"
+  "REAL_LT_MUL2" > "Ring_and_Field.mult_strict_mono'"
+  "REAL_LT_MUL" > "Ring_and_Field.mult_pos"
+  "REAL_LT_LMUL_IMP" > "Ring_and_Field.almost_ordered_semiring.mult_strict_left_mono"
+  "REAL_LT_LMUL_0" > "HOL4Real.real.REAL_LT_LMUL_0"
+  "REAL_LT_LMUL" > "HOL4Real.real.REAL_LT_LMUL"
+  "REAL_LT_LE" > "HOL.order.order_less_le"
+  "REAL_LT_LDIV_EQ" > "Ring_and_Field.pos_divide_less_eq"
+  "REAL_LT_LADD" > "Ring_and_Field.add_less_cancel_left"
+  "REAL_LT_INV_EQ" > "Ring_and_Field.inverse_positive_iff_positive"
+  "REAL_LT_INV" > "Ring_and_Field.less_imp_inverse_less"
+  "REAL_LT_IMP_NE" > "HOL.less_imp_neq"
+  "REAL_LT_IMP_LE" > "HOL.order_less_imp_le"
+  "REAL_LT_IADD" > "Ring_and_Field.add_strict_left_mono"
+  "REAL_LT_HALF2" > "HOL4Real.real.REAL_LT_HALF2"
+  "REAL_LT_HALF1" > "NatSimprocs.half_gt_zero_iff"
+  "REAL_LT_GT" > "HOL.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" > "HOL4Real.real.REAL_LT_DIV"
+  "REAL_LT_ANTISYM" > "HOL4Real.real.REAL_LT_ANTISYM"
+  "REAL_LT_ADD_SUB" > "Ring_and_Field.compare_rls_7"
+  "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" > "Ring_and_Field.add_strict_mono"
+  "REAL_LT_ADD1" > "HOL4Real.real.REAL_LT_ADD1"
+  "REAL_LT_ADD" > "RealDef.real_add_order"
+  "REAL_LT_1" > "HOL4Real.real.REAL_LT_1"
+  "REAL_LT_01" > "Ring_and_Field.ordered_semiring.zero_less_one"
+  "REAL_LTE_TRANS" > "Set.basic_trans_rules_24"
+  "REAL_LTE_TOTAL" > "HOL4Real.real.REAL_LTE_TOTAL"
+  "REAL_LTE_ANTSYM" > "HOL4Real.real.REAL_LTE_ANTSYM"
+  "REAL_LTE_ADD2" > "Ring_and_Field.add_less_le_mono"
+  "REAL_LTE_ADD" > "HOL4Real.real.REAL_LTE_ADD"
+  "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_LINV_UNIQ" > "HOL4Real.real.REAL_LINV_UNIQ"
+  "REAL_LE_TRANS" > "Set.basic_trans_rules_25"
+  "REAL_LE_TOTAL" > "HOL.linorder.linorder_linear"
+  "REAL_LE_SUB_RADD" > "Ring_and_Field.compare_rls_8"
+  "REAL_LE_SUB_LADD" > "Ring_and_Field.compare_rls_9"
+  "REAL_LE_SQUARE" > "Ring_and_Field.zero_le_square"
+  "REAL_LE_RNEG" > "HOL4Real.real.REAL_LE_RNEG"
+  "REAL_LE_RMUL_IMP" > "Ring_and_Field.mult_right_mono"
+  "REAL_LE_RMUL" > "RealDef.real_mult_le_cancel_iff1"
+  "REAL_LE_REFL" > "HOL.order.order_refl"
+  "REAL_LE_RDIV_EQ" > "Ring_and_Field.pos_le_divide_eq"
+  "REAL_LE_RDIV" > "HOL4Real.real.REAL_LE_RDIV"
+  "REAL_LE_RADD" > "Ring_and_Field.add_le_cancel_right"
+  "REAL_LE_POW2" > "NatBin.zero_le_power2"
+  "REAL_LE_NEGTOTAL" > "HOL4Real.real.REAL_LE_NEGTOTAL"
+  "REAL_LE_NEGR" > "Ring_and_Field.le_minus_self_iff"
+  "REAL_LE_NEGL" > "Ring_and_Field.minus_le_self_iff"
+  "REAL_LE_NEG2" > "Ring_and_Field.neg_le_iff_le"
+  "REAL_LE_NEG" > "Ring_and_Field.neg_le_iff_le"
+  "REAL_LE_MUL2" > "HOL4Real.real.REAL_LE_MUL2"
+  "REAL_LE_MUL" > "HOL4Real.real.REAL_LE_MUL"
+  "REAL_LE_LT" > "HOL.order_le_less"
+  "REAL_LE_LNEG" > "RealDef.real_0_le_add_iff"
+  "REAL_LE_LMUL_IMP" > "Ring_and_Field.mult_left_mono"
+  "REAL_LE_LMUL" > "RealDef.real_mult_le_cancel_iff2"
+  "REAL_LE_LDIV_EQ" > "Ring_and_Field.pos_divide_le_eq"
+  "REAL_LE_LDIV" > "HOL4Real.real.REAL_LE_LDIV"
+  "REAL_LE_LADD_IMP" > "Ring_and_Field.almost_ordered_semiring.add_left_mono"
+  "REAL_LE_LADD" > "Ring_and_Field.add_le_cancel_left"
+  "REAL_LE_INV_EQ" > "Ring_and_Field.inverse_nonnegative_iff_nonnegative"
+  "REAL_LE_INV" > "HOL4Real.real.REAL_LE_INV"
+  "REAL_LE_DOUBLE" > "HOL4Real.real.REAL_LE_DOUBLE"
+  "REAL_LE_DIV" > "HOL4Real.real.REAL_LE_DIV"
+  "REAL_LE_ANTISYM" > "HOL.order_eq_iff"
+  "REAL_LE_ADDR" > "HOL4Real.real.REAL_LE_ADDR"
+  "REAL_LE_ADDL" > "HOL4Real.real.REAL_LE_ADDL"
+  "REAL_LE_ADD2" > "Ring_and_Field.add_mono"
+  "REAL_LE_ADD" > "RealDef.real_le_add_order"
+  "REAL_LE_01" > "Ring_and_Field.zero_le_one"
+  "REAL_LET_TRANS" > "Set.basic_trans_rules_23"
+  "REAL_LET_TOTAL" > "HOL4Real.real.REAL_LET_TOTAL"
+  "REAL_LET_ANTISYM" > "HOL4Real.real.REAL_LET_ANTISYM"
+  "REAL_LET_ADD2" > "Ring_and_Field.add_le_less_mono"
+  "REAL_LET_ADD" > "HOL4Real.real.REAL_LET_ADD"
+  "REAL_LE1_POW2" > "HOL4Real.real.REAL_LE1_POW2"
+  "REAL_LE" > "RealDef.real_of_nat_le_iff"
+  "REAL_LDISTRIB" > "Ring_and_Field.ring_distrib_1"
+  "REAL_INV_POS" > "Ring_and_Field.positive_imp_inverse_positive"
+  "REAL_INV_NZ" > "Ring_and_Field.nonzero_imp_inverse_nonzero"
+  "REAL_INV_MUL" > "HOL4Real.real.REAL_INV_MUL"
+  "REAL_INV_LT1" > "RealDef.real_inverse_gt_one"
+  "REAL_INV_INV" > "Ring_and_Field.inverse_inverse_eq"
+  "REAL_INV_EQ_0" > "Ring_and_Field.inverse_nonzero_iff_nonzero"
+  "REAL_INV_1OVER" > "Ring_and_Field.inverse_eq_divide"
+  "REAL_INV_0" > "Ring_and_Field.division_by_zero.inverse_zero"
+  "REAL_INVINV" > "Ring_and_Field.nonzero_inverse_inverse_eq"
+  "REAL_INV1" > "Ring_and_Field.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" > "Ring_and_Field.compare_rls_10"
+  "REAL_EQ_SUB_LADD" > "Ring_and_Field.compare_rls_11"
+  "REAL_EQ_RMUL_IMP" > "Ring_and_Field.field_mult_cancel_right_lemma"
+  "REAL_EQ_RMUL" > "Ring_and_Field.field_mult_cancel_right"
+  "REAL_EQ_RDIV_EQ" > "HOL4Real.real.REAL_EQ_RDIV_EQ"
+  "REAL_EQ_RADD" > "Ring_and_Field.add_right_cancel"
+  "REAL_EQ_NEG" > "Ring_and_Field.neg_equal_iff_equal"
+  "REAL_EQ_MUL_LCANCEL" > "Ring_and_Field.field_mult_cancel_left"
+  "REAL_EQ_LMUL_IMP" > "HOL4Real.real.REAL_EQ_LMUL_IMP"
+  "REAL_EQ_LMUL2" > "RealDef.real_mult_left_cancel"
+  "REAL_EQ_LMUL" > "Ring_and_Field.field_mult_cancel_left"
+  "REAL_EQ_LDIV_EQ" > "HOL4Real.real.REAL_EQ_LDIV_EQ"
+  "REAL_EQ_LADD" > "Ring_and_Field.add_left_cancel"
+  "REAL_EQ_IMP_LE" > "HOL.order_eq_refl"
+  "REAL_ENTIRE" > "Ring_and_Field.field_mult_eq_0_iff"
+  "REAL_DOWN2" > "RealDef.real_lbound_gt_zero"
+  "REAL_DOWN" > "HOL4Real.real.REAL_DOWN"
+  "REAL_DOUBLE" > "IntArith.mult_2"
+  "REAL_DIV_RMUL" > "HOL4Real.real.REAL_DIV_RMUL"
+  "REAL_DIV_REFL" > "Ring_and_Field.divide_self"
+  "REAL_DIV_MUL2" > "Ring_and_Field.nonzero_mult_divide_cancel_left"
+  "REAL_DIV_LZERO" > "Ring_and_Field.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_ARCH" > "RComplete.reals_Archimedean3"
+  "REAL_ADD_SYM" > "Ring_and_Field.add_ac_2"
+  "REAL_ADD_SUB2" > "HOL4Real.real.REAL_ADD_SUB2"
+  "REAL_ADD_SUB" > "Ring_and_Field.add_minus_self_right"
+  "REAL_ADD_RINV" > "Ring_and_Field.right_minus"
+  "REAL_ADD_RID_UNIQ" > "HOL4Real.real.REAL_ADD_RID_UNIQ"
+  "REAL_ADD_RID" > "Ring_and_Field.add_0_right"
+  "REAL_ADD_RDISTRIB" > "Ring_and_Field.ring_distrib_2"
+  "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
+  "REAL_ADD_LID_UNIQ" > "HOL4Real.real.REAL_ADD_LID_UNIQ"
+  "REAL_ADD_LID" > "Ring_and_Field.abelian_semigroup.add_0"
+  "REAL_ADD_LDISTRIB" > "Ring_and_Field.ring_distrib_1"
+  "REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC"
+  "REAL_ADD2_SUB2" > "HOL4Real.real.REAL_ADD2_SUB2"
+  "REAL_ADD" > "RealDef.real_of_nat_add"
+  "REAL_ABS_TRIANGLE" > "Ring_and_Field.abs_triangle_ineq"
+  "REAL_ABS_POS" > "Ring_and_Field.abs_ge_zero"
+  "REAL_ABS_MUL" > "Ring_and_Field.abs_mult"
+  "REAL_ABS_0" > "Ring_and_Field.abs_zero"
+  "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_POS_LT" > "HOL4Real.real.POW_POS_LT"
+  "POW_POS" > "Power.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" > "NatBin.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" > "NatBin.power2_eq_square"
+  "POW_1" > "Power.power_one_right"
+  "POW_0" > "Power.power_0_Suc"
+  "ABS_ZERO" > "Ring_and_Field.abs_zero_iff"
+  "ABS_TRIANGLE" > "Ring_and_Field.abs_triangle_ineq"
+  "ABS_SUM" > "HOL4Real.real.ABS_SUM"
+  "ABS_SUB_ABS" > "HOL4Real.real.ABS_SUB_ABS"
+  "ABS_SUB" > "HOL4Real.real.ABS_SUB"
+  "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" > "NatBin.abs_power2"
+  "ABS_POS" > "Ring_and_Field.abs_ge_zero"
+  "ABS_NZ" > "Ring_and_Field.zero_less_abs_iff"
+  "ABS_NEG" > "Ring_and_Field.abs_minus_cancel"
+  "ABS_N" > "RealDef.abs_real_of_nat_cancel"
+  "ABS_MUL" > "Ring_and_Field.abs_mult"
+  "ABS_LT_MUL2" > "HOL4Real.real.ABS_LT_MUL2"
+  "ABS_LE" > "Ring_and_Field.abs_ge_self"
+  "ABS_INV" > "Ring_and_Field.nonzero_abs_inverse"
+  "ABS_DIV" > "Ring_and_Field.nonzero_abs_divide"
+  "ABS_CIRCLE" > "HOL4Real.real.ABS_CIRCLE"
+  "ABS_CASES" > "HOL4Real.real.ABS_CASES"
+  "ABS_BOUNDS" > "RealDef.abs_le_interval_iff"
+  "ABS_BOUND" > "HOL4Real.real.ABS_BOUND"
+  "ABS_BETWEEN2" > "HOL4Real.real.ABS_BETWEEN2"
+  "ABS_BETWEEN1" > "HOL4Real.real.ABS_BETWEEN1"
+  "ABS_BETWEEN" > "HOL4Real.real.ABS_BETWEEN"
+  "ABS_ABS" > "Ring_and_Field.abs_idempotent"
+  "ABS_1" > "Ring_and_Field.abs_one"
+  "ABS_0" > "Ring_and_Field.abs_zero"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/realax.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,141 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "treal_of_hreal" > "treal_of_hreal_def"
+  "treal_neg" > "treal_neg_def"
+  "treal_mul" > "treal_mul_def"
+  "treal_lt" > "treal_lt_def"
+  "treal_inv" > "treal_inv_def"
+  "treal_eq" > "treal_eq_def"
+  "treal_add" > "treal_add_def"
+  "treal_1" > "treal_1_def"
+  "treal_0" > "treal_0_def"
+  "hreal_of_treal" > "hreal_of_treal_def"
+
+type_maps
+  "real" > "RealDef.real"
+
+const_maps
+  "treal_of_hreal" > "HOL4Real.realax.treal_of_hreal"
+  "treal_neg" > "HOL4Real.realax.treal_neg"
+  "treal_mul" > "HOL4Real.realax.treal_mul"
+  "treal_lt" > "HOL4Real.realax.treal_lt"
+  "treal_inv" > "HOL4Real.realax.treal_inv"
+  "treal_eq" > "HOL4Real.realax.treal_eq"
+  "treal_add" > "HOL4Real.realax.treal_add"
+  "treal_1" > "HOL4Real.realax.treal_1"
+  "treal_0" > "HOL4Real.realax.treal_0"
+  "real_neg" > "uminus" :: "real => real"
+  "real_mul" > "op *" :: "real => real => real"
+  "real_lt" > "op <" :: "real => real => bool"
+  "real_add" > "op +" :: "real => real => real"
+  "real_1" > "1" :: "real"
+  "real_0" > "0" :: "real"
+  "inv" > "HOL.inverse" :: "real => real"
+  "hreal_of_treal" > "HOL4Real.realax.hreal_of_treal"
+
+thm_maps
+  "treal_of_hreal_def" > "HOL4Real.realax.treal_of_hreal_def"
+  "treal_of_hreal" > "HOL4Real.realax.treal_of_hreal"
+  "treal_neg_def" > "HOL4Real.realax.treal_neg_def"
+  "treal_neg" > "HOL4Real.realax.treal_neg"
+  "treal_mul_def" > "HOL4Real.realax.treal_mul_def"
+  "treal_mul" > "HOL4Real.realax.treal_mul"
+  "treal_lt_def" > "HOL4Real.realax.treal_lt_def"
+  "treal_lt" > "HOL4Real.realax.treal_lt"
+  "treal_inv_def" > "HOL4Real.realax.treal_inv_def"
+  "treal_inv" > "HOL4Real.realax.treal_inv"
+  "treal_eq_def" > "HOL4Real.realax.treal_eq_def"
+  "treal_eq" > "HOL4Real.realax.treal_eq"
+  "treal_add_def" > "HOL4Real.realax.treal_add_def"
+  "treal_add" > "HOL4Real.realax.treal_add"
+  "treal_1_def" > "HOL4Real.realax.treal_1_def"
+  "treal_1" > "HOL4Real.realax.treal_1"
+  "treal_0_def" > "HOL4Real.realax.treal_0_def"
+  "treal_0" > "HOL4Real.realax.treal_0"
+  "hreal_of_treal_def" > "HOL4Real.realax.hreal_of_treal_def"
+  "hreal_of_treal" > "HOL4Real.realax.hreal_of_treal"
+  "TREAL_NEG_WELLDEF" > "HOL4Real.realax.TREAL_NEG_WELLDEF"
+  "TREAL_MUL_WELLDEFR" > "HOL4Real.realax.TREAL_MUL_WELLDEFR"
+  "TREAL_MUL_WELLDEF" > "HOL4Real.realax.TREAL_MUL_WELLDEF"
+  "TREAL_MUL_SYM" > "HOL4Real.realax.TREAL_MUL_SYM"
+  "TREAL_MUL_LINV" > "HOL4Real.realax.TREAL_MUL_LINV"
+  "TREAL_MUL_LID" > "HOL4Real.realax.TREAL_MUL_LID"
+  "TREAL_MUL_ASSOC" > "HOL4Real.realax.TREAL_MUL_ASSOC"
+  "TREAL_LT_WELLDEFR" > "HOL4Real.realax.TREAL_LT_WELLDEFR"
+  "TREAL_LT_WELLDEFL" > "HOL4Real.realax.TREAL_LT_WELLDEFL"
+  "TREAL_LT_WELLDEF" > "HOL4Real.realax.TREAL_LT_WELLDEF"
+  "TREAL_LT_TRANS" > "HOL4Real.realax.TREAL_LT_TRANS"
+  "TREAL_LT_TOTAL" > "HOL4Real.realax.TREAL_LT_TOTAL"
+  "TREAL_LT_REFL" > "HOL4Real.realax.TREAL_LT_REFL"
+  "TREAL_LT_MUL" > "HOL4Real.realax.TREAL_LT_MUL"
+  "TREAL_LT_ADD" > "HOL4Real.realax.TREAL_LT_ADD"
+  "TREAL_LDISTRIB" > "HOL4Real.realax.TREAL_LDISTRIB"
+  "TREAL_ISO" > "HOL4Real.realax.TREAL_ISO"
+  "TREAL_INV_WELLDEF" > "HOL4Real.realax.TREAL_INV_WELLDEF"
+  "TREAL_INV_0" > "HOL4Real.realax.TREAL_INV_0"
+  "TREAL_EQ_TRANS" > "HOL4Real.realax.TREAL_EQ_TRANS"
+  "TREAL_EQ_SYM" > "HOL4Real.realax.TREAL_EQ_SYM"
+  "TREAL_EQ_REFL" > "HOL4Real.realax.TREAL_EQ_REFL"
+  "TREAL_EQ_EQUIV" > "HOL4Real.realax.TREAL_EQ_EQUIV"
+  "TREAL_EQ_AP" > "HOL4Real.realax.TREAL_EQ_AP"
+  "TREAL_BIJ_WELLDEF" > "HOL4Real.realax.TREAL_BIJ_WELLDEF"
+  "TREAL_BIJ" > "HOL4Real.realax.TREAL_BIJ"
+  "TREAL_ADD_WELLDEFR" > "HOL4Real.realax.TREAL_ADD_WELLDEFR"
+  "TREAL_ADD_WELLDEF" > "HOL4Real.realax.TREAL_ADD_WELLDEF"
+  "TREAL_ADD_SYM" > "HOL4Real.realax.TREAL_ADD_SYM"
+  "TREAL_ADD_LINV" > "HOL4Real.realax.TREAL_ADD_LINV"
+  "TREAL_ADD_LID" > "HOL4Real.realax.TREAL_ADD_LID"
+  "TREAL_ADD_ASSOC" > "HOL4Real.realax.TREAL_ADD_ASSOC"
+  "TREAL_10" > "HOL4Real.realax.TREAL_10"
+  "REAL_SUP_ALLPOS" > "HOL4Compat.REAL_SUP_ALLPOS"
+  "REAL_MUL_SYM" > "Ring_and_Field.mult_ac_2"
+  "REAL_MUL_LINV" > "HOL4Compat.REAL_MUL_LINV"
+  "REAL_MUL_LID" > "Ring_and_Field.almost_semiring.mult_1"
+  "REAL_MUL_ASSOC" > "HOL4Compat.REAL_MUL_ASSOC"
+  "REAL_LT_TRANS" > "Set.basic_trans_rules_21"
+  "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL"
+  "REAL_LT_REFL" > "HOL.order_less_irrefl"
+  "REAL_LT_MUL" > "Ring_and_Field.mult_pos"
+  "REAL_LT_IADD" > "Ring_and_Field.add_strict_left_mono"
+  "REAL_LDISTRIB" > "Ring_and_Field.ring_distrib_1"
+  "REAL_INV_0" > "Ring_and_Field.division_by_zero.inverse_zero"
+  "REAL_ADD_SYM" > "Ring_and_Field.add_ac_2"
+  "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
+  "REAL_ADD_LID" > "Ring_and_Field.abelian_semigroup.add_0"
+  "REAL_ADD_ASSOC" > "HOL4Compat.REAL_ADD_ASSOC"
+  "REAL_10" > "HOL4Compat.REAL_10"
+  "HREAL_RDISTRIB" > "HOL4Real.realax.HREAL_RDISTRIB"
+  "HREAL_LT_REFL" > "HOL4Real.realax.HREAL_LT_REFL"
+  "HREAL_LT_NE" > "HOL4Real.realax.HREAL_LT_NE"
+  "HREAL_LT_LADD" > "HOL4Real.realax.HREAL_LT_LADD"
+  "HREAL_LT_GT" > "HOL4Real.realax.HREAL_LT_GT"
+  "HREAL_LT_ADDR" > "HOL4Real.realax.HREAL_LT_ADDR"
+  "HREAL_LT_ADDL" > "HOL4Real.realax.HREAL_LT_ADDL"
+  "HREAL_LT_ADD2" > "HOL4Real.realax.HREAL_LT_ADD2"
+  "HREAL_EQ_LADD" > "HOL4Real.realax.HREAL_EQ_LADD"
+  "HREAL_EQ_ADDR" > "HOL4Base.hreal.HREAL_NOZERO"
+  "HREAL_EQ_ADDL" > "HOL4Real.realax.HREAL_EQ_ADDL"
+
+ignore_thms
+  "real_tybij"
+  "real_of_hreal"
+  "real_neg"
+  "real_mul"
+  "real_lt"
+  "real_inv"
+  "real_add"
+  "real_TY_DEF"
+  "real_1"
+  "real_0"
+  "hreal_of_real"
+  "SUP_ALLPOS_LEMMA4"
+  "SUP_ALLPOS_LEMMA3"
+  "SUP_ALLPOS_LEMMA2"
+  "SUP_ALLPOS_LEMMA1"
+  "REAL_POS"
+  "REAL_ISO_EQ"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/relation.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,102 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "transitive" > "transitive_primdef"
+  "the_fun" > "the_fun_primdef"
+  "pred_reflexive" > "pred_reflexive_def"
+  "inv_image" > "inv_image_primdef"
+  "approx" > "approx_primdef"
+  "WFREC" > "WFREC_def"
+  "WF" > "WF_def"
+  "TC" > "TC_def"
+  "RTC" > "RTC_def"
+  "RESTRICT" > "RESTRICT_def"
+  "RC" > "RC_primdef"
+  "EMPTY_REL" > "EMPTY_REL_def"
+
+const_maps
+  "transitive" > "HOL4Base.relation.transitive"
+  "the_fun" > "HOL4Base.relation.the_fun"
+  "pred_reflexive" > "HOL4Base.relation.pred_reflexive"
+  "inv_image" > "HOL4Base.relation.inv_image"
+  "approx" > "HOL4Base.relation.approx"
+  "WFREC" > "HOL4Base.relation.WFREC"
+  "WF" > "HOL4Base.relation.WF"
+  "TC" > "HOL4Base.relation.TC"
+  "RTC" > "HOL4Base.relation.RTC"
+  "RESTRICT" > "HOL4Base.relation.RESTRICT"
+  "RC" > "HOL4Base.relation.RC"
+  "EMPTY_REL" > "HOL4Base.relation.EMPTY_REL"
+
+const_renames
+  "reflexive" > "pred_reflexive"
+
+thm_maps
+  "transitive_primdef" > "HOL4Base.relation.transitive_primdef"
+  "transitive_def" > "HOL4Base.relation.transitive_def"
+  "the_fun_primdef" > "HOL4Base.relation.the_fun_primdef"
+  "the_fun_def" > "HOL4Base.relation.the_fun_def"
+  "reflexive_def" > "HOL4Base.relation.reflexive_def"
+  "pred_reflexive_def" > "HOL4Base.relation.pred_reflexive_def"
+  "inv_image_primdef" > "HOL4Base.relation.inv_image_primdef"
+  "inv_image_def" > "HOL4Base.relation.inv_image_def"
+  "approx_primdef" > "HOL4Base.relation.approx_primdef"
+  "approx_def" > "HOL4Base.relation.approx_def"
+  "WF_inv_image" > "HOL4Base.relation.WF_inv_image"
+  "WF_def" > "HOL4Base.relation.WF_def"
+  "WF_TC" > "HOL4Base.relation.WF_TC"
+  "WF_SUBSET" > "HOL4Base.relation.WF_SUBSET"
+  "WF_RECURSION_THM" > "HOL4Base.relation.WF_RECURSION_THM"
+  "WF_NOT_REFL" > "HOL4Base.relation.WF_NOT_REFL"
+  "WF_INDUCTION_THM" > "HOL4Base.relation.WF_INDUCTION_THM"
+  "WF_EMPTY_REL" > "HOL4Base.relation.WF_EMPTY_REL"
+  "WF_DEF" > "HOL4Base.relation.WF_DEF"
+  "WFREC_def" > "HOL4Base.relation.WFREC_def"
+  "WFREC_THM" > "HOL4Base.relation.WFREC_THM"
+  "WFREC_DEF" > "HOL4Base.relation.WFREC_DEF"
+  "WFREC_COROLLARY" > "HOL4Base.relation.WFREC_COROLLARY"
+  "TC_def" > "HOL4Base.relation.TC_def"
+  "TC_TRANSITIVE" > "HOL4Base.relation.TC_TRANSITIVE"
+  "TC_SUBSET" > "HOL4Base.relation.TC_SUBSET"
+  "TC_STRONG_INDUCT_LEFT1" > "HOL4Base.relation.TC_STRONG_INDUCT_LEFT1"
+  "TC_STRONG_INDUCT" > "HOL4Base.relation.TC_STRONG_INDUCT"
+  "TC_RULES" > "HOL4Base.relation.TC_RULES"
+  "TC_RTC" > "HOL4Base.relation.TC_RTC"
+  "TC_RC_EQNS" > "HOL4Base.relation.TC_RC_EQNS"
+  "TC_MONOTONE" > "HOL4Base.relation.TC_MONOTONE"
+  "TC_INDUCT_LEFT1" > "HOL4Base.relation.TC_INDUCT_LEFT1"
+  "TC_INDUCT" > "HOL4Base.relation.TC_INDUCT"
+  "TC_IDEM" > "HOL4Base.relation.TC_IDEM"
+  "TC_DEF" > "HOL4Base.relation.TC_DEF"
+  "TC_CASES2" > "HOL4Base.relation.TC_CASES2"
+  "TC_CASES1" > "HOL4Base.relation.TC_CASES1"
+  "RTC_def" > "HOL4Base.relation.RTC_def"
+  "RTC_TRANSITIVE" > "HOL4Base.relation.RTC_TRANSITIVE"
+  "RTC_TC_RC" > "HOL4Base.relation.RTC_TC_RC"
+  "RTC_SUBSET" > "HOL4Base.relation.RTC_SUBSET"
+  "RTC_STRONG_INDUCT" > "HOL4Base.relation.RTC_STRONG_INDUCT"
+  "RTC_RULES" > "HOL4Base.relation.RTC_RULES"
+  "RTC_RTC" > "HOL4Base.relation.RTC_RTC"
+  "RTC_REFLEXIVE" > "HOL4Base.relation.RTC_REFLEXIVE"
+  "RTC_MONOTONE" > "HOL4Base.relation.RTC_MONOTONE"
+  "RTC_INDUCT" > "HOL4Base.relation.RTC_INDUCT"
+  "RTC_IDEM" > "HOL4Base.relation.RTC_IDEM"
+  "RTC_DEF" > "HOL4Base.relation.RTC_DEF"
+  "RTC_CASES_RTC_TWICE" > "HOL4Base.relation.RTC_CASES_RTC_TWICE"
+  "RTC_CASES2" > "HOL4Base.relation.RTC_CASES2"
+  "RTC_CASES1" > "HOL4Base.relation.RTC_CASES1"
+  "RESTRICT_def" > "HOL4Base.relation.RESTRICT_def"
+  "RESTRICT_LEMMA" > "HOL4Base.relation.RESTRICT_LEMMA"
+  "RESTRICT_DEF" > "HOL4Base.relation.RESTRICT_DEF"
+  "RC_primdef" > "HOL4Base.relation.RC_primdef"
+  "RC_def" > "HOL4Base.relation.RC_def"
+  "RC_SUBSET" > "HOL4Base.relation.RC_SUBSET"
+  "RC_RTC" > "HOL4Base.relation.RC_RTC"
+  "RC_REFLEXIVE" > "HOL4Base.relation.RC_REFLEXIVE"
+  "RC_IDEM" > "HOL4Base.relation.RC_IDEM"
+  "EMPTY_REL_def" > "HOL4Base.relation.EMPTY_REL_def"
+  "EMPTY_REL_DEF" > "HOL4Base.relation.EMPTY_REL_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/res_quan.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,37 @@
+import
+
+import_segment "hol4"
+
+thm_maps
+  "RES_SELECT_UNIV" > "HOL4Vec.res_quan.RES_SELECT_UNIV"
+  "RES_SELECT_EMPTY" > "HOL4Vec.res_quan.RES_SELECT_EMPTY"
+  "RES_SELECT" > "HOL4Base.bool.RES_SELECT_DEF"
+  "RES_FORALL_UNIV" > "HOL4Vec.res_quan.RES_FORALL_UNIV"
+  "RES_FORALL_UNIQUE" > "HOL4Vec.res_quan.RES_FORALL_UNIQUE"
+  "RES_FORALL_REORDER" > "HOL4Vec.res_quan.RES_FORALL_REORDER"
+  "RES_FORALL_NULL" > "HOL4Vec.res_quan.RES_FORALL_NULL"
+  "RES_FORALL_FORALL" > "HOL4Vec.res_quan.RES_FORALL_FORALL"
+  "RES_FORALL_EMPTY" > "HOL4Vec.res_quan.RES_FORALL_EMPTY"
+  "RES_FORALL_DISJ_DIST" > "HOL4Vec.res_quan.RES_FORALL_DISJ_DIST"
+  "RES_FORALL_CONJ_DIST" > "HOL4Vec.res_quan.RES_FORALL_CONJ_DIST"
+  "RES_FORALL" > "HOL4Base.bool.RES_FORALL_DEF"
+  "RES_EXISTS_UNIV" > "HOL4Vec.res_quan.RES_EXISTS_UNIV"
+  "RES_EXISTS_UNIQUE_UNIV" > "HOL4Vec.res_quan.RES_EXISTS_UNIQUE_UNIV"
+  "RES_EXISTS_UNIQUE_NULL" > "HOL4Vec.res_quan.RES_EXISTS_UNIQUE_NULL"
+  "RES_EXISTS_UNIQUE_EMPTY" > "HOL4Vec.res_quan.RES_EXISTS_UNIQUE_EMPTY"
+  "RES_EXISTS_UNIQUE_ALT" > "HOL4Vec.res_quan.RES_EXISTS_UNIQUE_ALT"
+  "RES_EXISTS_UNIQUE" > "HOL4Base.bool.RES_EXISTS_UNIQUE_DEF"
+  "RES_EXISTS_REORDER" > "HOL4Vec.res_quan.RES_EXISTS_REORDER"
+  "RES_EXISTS_NULL" > "HOL4Vec.res_quan.RES_EXISTS_NULL"
+  "RES_EXISTS_EQUAL" > "HOL4Vec.res_quan.RES_EXISTS_EQUAL"
+  "RES_EXISTS_EMPTY" > "HOL4Vec.res_quan.RES_EXISTS_EMPTY"
+  "RES_EXISTS_DISJ_DIST" > "HOL4Vec.res_quan.RES_EXISTS_DISJ_DIST"
+  "RES_EXISTS_ALT" > "HOL4Vec.res_quan.RES_EXISTS_ALT"
+  "RES_EXISTS" > "HOL4Base.bool.RES_EXISTS_DEF"
+  "RES_DISJ_EXISTS_DIST" > "HOL4Vec.res_quan.RES_DISJ_EXISTS_DIST"
+  "RES_ABSTRACT_IDEMPOT" > "HOL4Vec.res_quan.RES_ABSTRACT_IDEMPOT"
+  "RES_ABSTRACT_EQUAL_EQ" > "HOL4Vec.res_quan.RES_ABSTRACT_EQUAL_EQ"
+  "RES_ABSTRACT_EQUAL" > "HOL4Vec.res_quan.RES_ABSTRACT_EQUAL"
+  "RES_ABSTRACT" > "HOL4Vec.res_quan.RES_ABSTRACT"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/rich_list.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,375 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "UNZIP_SND" > "UNZIP_SND_def"
+  "UNZIP_FST" > "UNZIP_FST_def"
+  "SUFFIX" > "SUFFIX_def"
+  "SPLITP" > "SPLITP_def"
+  "SNOC" > "SNOC_def"
+  "SEG" > "SEG_def"
+  "SCANR" > "SCANR_def"
+  "SCANL" > "SCANL_def"
+  "REPLICATE" > "REPLICATE_def"
+  "PREFIX" > "PREFIX_def"
+  "OR_EL" > "OR_EL_def"
+  "LASTN" > "LASTN_def"
+  "IS_SUFFIX" > "IS_SUFFIX_def"
+  "IS_SUBLIST" > "IS_SUBLIST_def"
+  "IS_PREFIX" > "IS_PREFIX_def"
+  "GENLIST" > "GENLIST_def"
+  "FIRSTN" > "FIRSTN_def"
+  "ELL" > "ELL_def"
+  "BUTLASTN" > "BUTLASTN_def"
+  "BUTFIRSTN" > "BUTFIRSTN_def"
+  "AND_EL" > "AND_EL_def"
+
+const_maps
+  "UNZIP_SND" > "HOL4Base.rich_list.UNZIP_SND"
+  "UNZIP_FST" > "HOL4Base.rich_list.UNZIP_FST"
+  "SUFFIX" > "HOL4Base.rich_list.SUFFIX"
+  "PREFIX" > "HOL4Base.rich_list.PREFIX"
+  "OR_EL" > "HOL4Base.rich_list.OR_EL"
+  "AND_EL" > "HOL4Base.rich_list.AND_EL"
+
+thm_maps
+  "list_INDUCT" > "HOL4Compat.unzip.induct"
+  "list_CASES" > "HOL4Compat.list_CASES"
+  "ZIP_UNZIP" > "HOL4Base.list.ZIP_UNZIP"
+  "ZIP_SNOC" > "HOL4Base.rich_list.ZIP_SNOC"
+  "ZIP" > "HOL4Compat.ZIP"
+  "UNZIP_ZIP" > "HOL4Base.list.UNZIP_ZIP"
+  "UNZIP_SNOC" > "HOL4Base.rich_list.UNZIP_SNOC"
+  "UNZIP_SND_def" > "HOL4Base.rich_list.UNZIP_SND_def"
+  "UNZIP_SND_DEF" > "HOL4Base.rich_list.UNZIP_SND_DEF"
+  "UNZIP_FST_def" > "HOL4Base.rich_list.UNZIP_FST_def"
+  "UNZIP_FST_DEF" > "HOL4Base.rich_list.UNZIP_FST_DEF"
+  "UNZIP" > "HOL4Compat.UNZIP"
+  "TL_SNOC" > "HOL4Base.rich_list.TL_SNOC"
+  "TL" > "List.tl.simps_2"
+  "SUM_SNOC" > "HOL4Base.rich_list.SUM_SNOC"
+  "SUM_REVERSE" > "HOL4Base.rich_list.SUM_REVERSE"
+  "SUM_FOLDR" > "HOL4Compat.sum_def"
+  "SUM_FOLDL" > "HOL4Base.rich_list.SUM_FOLDL"
+  "SUM_FLAT" > "HOL4Base.rich_list.SUM_FLAT"
+  "SUM_APPEND" > "HOL4Base.rich_list.SUM_APPEND"
+  "SUM" > "HOL4Compat.SUM"
+  "SUFFIX_def" > "HOL4Base.rich_list.SUFFIX_def"
+  "SUFFIX_DEF" > "HOL4Base.rich_list.SUFFIX_DEF"
+  "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_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"
+  "SOME_EL_FOLDR" > "HOL4Base.rich_list.SOME_EL_FOLDR"
+  "SOME_EL_FOLDL_MAP" > "HOL4Base.rich_list.SOME_EL_FOLDL_MAP"
+  "SOME_EL_FOLDL" > "HOL4Base.rich_list.SOME_EL_FOLDL"
+  "SOME_EL_FIRSTN" > "HOL4Base.rich_list.SOME_EL_FIRSTN"
+  "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" > "HOL4Compat.list_exists_DEF"
+  "SNOC_REVERSE_CONS" > "HOL4Base.rich_list.SNOC_REVERSE_CONS"
+  "SNOC_INDUCT" > "HOL4Base.rich_list.SNOC_INDUCT"
+  "SNOC_FOLDR" > "HOL4Base.rich_list.SNOC_FOLDR"
+  "SNOC_EQ_LENGTH_EQ" > "HOL4Base.rich_list.SNOC_EQ_LENGTH_EQ"
+  "SNOC_CASES" > "HOL4Base.rich_list.SNOC_CASES"
+  "SNOC_Axiom" > "HOL4Base.rich_list.SNOC_Axiom"
+  "SNOC_APPEND" > "HOL4Base.rich_list.SNOC_APPEND"
+  "SNOC_11" > "HOL4Base.rich_list.SNOC_11"
+  "SNOC" > "HOL4Base.rich_list.SNOC"
+  "SEG_SUC_CONS" > "HOL4Base.rich_list.SEG_SUC_CONS"
+  "SEG_SNOC" > "HOL4Base.rich_list.SEG_SNOC"
+  "SEG_SEG" > "HOL4Base.rich_list.SEG_SEG"
+  "SEG_REVERSE" > "HOL4Base.rich_list.SEG_REVERSE"
+  "SEG_LENGTH_SNOC" > "HOL4Base.rich_list.SEG_LENGTH_SNOC"
+  "SEG_LENGTH_ID" > "HOL4Base.rich_list.SEG_LENGTH_ID"
+  "SEG_LASTN_BUTLASTN" > "HOL4Base.rich_list.SEG_LASTN_BUTLASTN"
+  "SEG_FIRSTN_BUTFISTN" > "HOL4Base.rich_list.SEG_FIRSTN_BUTFISTN"
+  "SEG_APPEND2" > "HOL4Base.rich_list.SEG_APPEND2"
+  "SEG_APPEND1" > "HOL4Base.rich_list.SEG_APPEND1"
+  "SEG_APPEND" > "HOL4Base.rich_list.SEG_APPEND"
+  "SEG_0_SNOC" > "HOL4Base.rich_list.SEG_0_SNOC"
+  "SEG" > "HOL4Base.rich_list.SEG"
+  "SCANR" > "HOL4Base.rich_list.SCANR"
+  "SCANL" > "HOL4Base.rich_list.SCANL"
+  "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_FLAT" > "HOL4Base.rich_list.REVERSE_FLAT"
+  "REVERSE_EQ_NIL" > "List.rev_is_Nil_conv"
+  "REVERSE_APPEND" > "List.rev_append"
+  "REVERSE" > "HOL4Base.rich_list.REVERSE"
+  "REPLICATE" > "HOL4Base.rich_list.REPLICATE"
+  "PREFIX_def" > "HOL4Base.rich_list.PREFIX_def"
+  "PREFIX_FOLDR" > "HOL4Base.rich_list.PREFIX_FOLDR"
+  "PREFIX_DEF" > "HOL4Base.rich_list.PREFIX_DEF"
+  "PREFIX" > "HOL4Base.rich_list.PREFIX"
+  "OR_EL_def" > "HOL4Base.rich_list.OR_EL_def"
+  "OR_EL_FOLDR" > "HOL4Base.rich_list.OR_EL_FOLDR"
+  "OR_EL_FOLDL" > "HOL4Base.rich_list.OR_EL_FOLDL"
+  "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_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_EQ_LIST" > "HOL4Base.list.NOT_EQ_LIST"
+  "NOT_CONS_NIL" > "List.list.simps_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_SNOC" > "HOL4Base.rich_list.MAP_SNOC"
+  "MAP_REVERSE" > "List.rev_map"
+  "MAP_MAP_o" > "List.map_compose"
+  "MAP_FOLDR" > "HOL4Base.rich_list.MAP_FOLDR"
+  "MAP_FOLDL" > "HOL4Base.rich_list.MAP_FOLDL"
+  "MAP_FLAT" > "List.map_concat"
+  "MAP_FILTER" > "HOL4Base.rich_list.MAP_FILTER"
+  "MAP_APPEND" > "List.map_append"
+  "MAP2_ZIP" > "HOL4Base.list.MAP2_ZIP"
+  "MAP2" > "HOL4Compat.MAP2"
+  "MAP" > "HOL4Compat.MAP"
+  "LIST_NOT_EQ" > "HOL4Base.list.LIST_NOT_EQ"
+  "LENGTH_ZIP" > "HOL4Base.list.LENGTH_ZIP"
+  "LENGTH_UNZIP_SND" > "HOL4Base.rich_list.LENGTH_UNZIP_SND"
+  "LENGTH_UNZIP_FST" > "HOL4Base.rich_list.LENGTH_UNZIP_FST"
+  "LENGTH_SNOC" > "HOL4Base.rich_list.LENGTH_SNOC"
+  "LENGTH_SEG" > "HOL4Base.rich_list.LENGTH_SEG"
+  "LENGTH_SCANR" > "HOL4Base.rich_list.LENGTH_SCANR"
+  "LENGTH_SCANL" > "HOL4Base.rich_list.LENGTH_SCANL"
+  "LENGTH_REVERSE" > "List.length_rev"
+  "LENGTH_REPLICATE" > "HOL4Base.rich_list.LENGTH_REPLICATE"
+  "LENGTH_NOT_NULL" > "HOL4Base.rich_list.LENGTH_NOT_NULL"
+  "LENGTH_NIL" > "List.length_0_conv"
+  "LENGTH_MAP2" > "HOL4Base.rich_list.LENGTH_MAP2"
+  "LENGTH_MAP" > "List.length_map"
+  "LENGTH_LASTN" > "HOL4Base.rich_list.LENGTH_LASTN"
+  "LENGTH_GENLIST" > "HOL4Base.rich_list.LENGTH_GENLIST"
+  "LENGTH_FOLDR" > "HOL4Base.rich_list.LENGTH_FOLDR"
+  "LENGTH_FOLDL" > "HOL4Base.rich_list.LENGTH_FOLDL"
+  "LENGTH_FLAT" > "HOL4Base.rich_list.LENGTH_FLAT"
+  "LENGTH_FIRSTN" > "HOL4Base.rich_list.LENGTH_FIRSTN"
+  "LENGTH_EQ_NIL" > "HOL4Base.list.LENGTH_EQ_NIL"
+  "LENGTH_EQ" > "HOL4Base.rich_list.LENGTH_EQ"
+  "LENGTH_CONS" > "HOL4Base.list.LENGTH_CONS"
+  "LENGTH_BUTLASTN" > "HOL4Base.rich_list.LENGTH_BUTLASTN"
+  "LENGTH_BUTLAST" > "HOL4Base.rich_list.LENGTH_BUTLAST"
+  "LENGTH_BUTFIRSTN" > "HOL4Base.rich_list.LENGTH_BUTFIRSTN"
+  "LENGTH_APPEND" > "List.length_append"
+  "LENGTH" > "HOL4Compat.LENGTH"
+  "LAST_LASTN_LAST" > "HOL4Base.rich_list.LAST_LASTN_LAST"
+  "LAST_CONS" > "HOL4Base.list.LAST_CONS"
+  "LASTN_SEG" > "HOL4Base.rich_list.LASTN_SEG"
+  "LASTN_REVERSE" > "HOL4Base.rich_list.LASTN_REVERSE"
+  "LASTN_MAP" > "HOL4Base.rich_list.LASTN_MAP"
+  "LASTN_LENGTH_ID" > "HOL4Base.rich_list.LASTN_LENGTH_ID"
+  "LASTN_LENGTH_APPEND" > "HOL4Base.rich_list.LASTN_LENGTH_APPEND"
+  "LASTN_LASTN" > "HOL4Base.rich_list.LASTN_LASTN"
+  "LASTN_CONS" > "HOL4Base.rich_list.LASTN_CONS"
+  "LASTN_BUTLASTN" > "HOL4Base.rich_list.LASTN_BUTLASTN"
+  "LASTN_BUTFIRSTN" > "HOL4Base.rich_list.LASTN_BUTFIRSTN"
+  "LASTN_APPEND2" > "HOL4Base.rich_list.LASTN_APPEND2"
+  "LASTN_APPEND1" > "HOL4Base.rich_list.LASTN_APPEND1"
+  "LASTN_1" > "HOL4Base.rich_list.LASTN_1"
+  "LASTN" > "HOL4Base.rich_list.LASTN"
+  "LAST" > "HOL4Base.rich_list.LAST"
+  "IS_SUFFIX_REVERSE" > "HOL4Base.rich_list.IS_SUFFIX_REVERSE"
+  "IS_SUFFIX_IS_SUBLIST" > "HOL4Base.rich_list.IS_SUFFIX_IS_SUBLIST"
+  "IS_SUFFIX_APPEND" > "HOL4Base.rich_list.IS_SUFFIX_APPEND"
+  "IS_SUFFIX" > "HOL4Base.rich_list.IS_SUFFIX"
+  "IS_SUBLIST_REVERSE" > "HOL4Base.rich_list.IS_SUBLIST_REVERSE"
+  "IS_SUBLIST_APPEND" > "HOL4Base.rich_list.IS_SUBLIST_APPEND"
+  "IS_SUBLIST" > "HOL4Base.rich_list.IS_SUBLIST"
+  "IS_PREFIX_REVERSE" > "HOL4Base.rich_list.IS_PREFIX_REVERSE"
+  "IS_PREFIX_PREFIX" > "HOL4Base.rich_list.IS_PREFIX_PREFIX"
+  "IS_PREFIX_IS_SUBLIST" > "HOL4Base.rich_list.IS_PREFIX_IS_SUBLIST"
+  "IS_PREFIX_APPEND" > "HOL4Base.rich_list.IS_PREFIX_APPEND"
+  "IS_PREFIX" > "HOL4Base.rich_list.IS_PREFIX"
+  "IS_EL_SOME_EL" > "HOL4Base.rich_list.IS_EL_SOME_EL"
+  "IS_EL_SNOC" > "HOL4Base.rich_list.IS_EL_SNOC"
+  "IS_EL_SEG" > "HOL4Base.rich_list.IS_EL_SEG"
+  "IS_EL_REVERSE" > "HOL4Base.rich_list.IS_EL_REVERSE"
+  "IS_EL_REPLICATE" > "HOL4Base.rich_list.IS_EL_REPLICATE"
+  "IS_EL_LASTN" > "HOL4Base.rich_list.IS_EL_LASTN"
+  "IS_EL_FOLDR_MAP" > "HOL4Base.rich_list.IS_EL_FOLDR_MAP"
+  "IS_EL_FOLDR" > "HOL4Base.rich_list.IS_EL_FOLDR"
+  "IS_EL_FOLDL_MAP" > "HOL4Base.rich_list.IS_EL_FOLDL_MAP"
+  "IS_EL_FOLDL" > "HOL4Base.rich_list.IS_EL_FOLDL"
+  "IS_EL_FIRSTN" > "HOL4Base.rich_list.IS_EL_FIRSTN"
+  "IS_EL_FILTER" > "HOL4Base.rich_list.IS_EL_FILTER"
+  "IS_EL_DEF" > "HOL4Base.rich_list.IS_EL_DEF"
+  "IS_EL_BUTLASTN" > "HOL4Base.rich_list.IS_EL_BUTLASTN"
+  "IS_EL_BUTFIRSTN" > "HOL4Base.rich_list.IS_EL_BUTFIRSTN"
+  "IS_EL_APPEND" > "HOL4Base.list.MEM_APPEND"
+  "IS_EL" > "HOL4Compat.MEM"
+  "HD" > "List.hd.simps"
+  "GENLIST" > "HOL4Base.rich_list.GENLIST"
+  "FOLDR_SNOC" > "HOL4Base.rich_list.FOLDR_SNOC"
+  "FOLDR_SINGLE" > "HOL4Base.rich_list.FOLDR_SINGLE"
+  "FOLDR_REVERSE" > "HOL4Base.rich_list.FOLDR_REVERSE"
+  "FOLDR_MAP_REVERSE" > "HOL4Base.rich_list.FOLDR_MAP_REVERSE"
+  "FOLDR_MAP" > "HOL4Base.rich_list.FOLDR_MAP"
+  "FOLDR_FOLDL_REVERSE" > "List.foldr_foldl"
+  "FOLDR_FOLDL" > "HOL4Base.rich_list.FOLDR_FOLDL"
+  "FOLDR_FILTER_REVERSE" > "HOL4Base.rich_list.FOLDR_FILTER_REVERSE"
+  "FOLDR_FILTER" > "HOL4Base.rich_list.FOLDR_FILTER"
+  "FOLDR_CONS_NIL" > "HOL4Base.rich_list.FOLDR_CONS_NIL"
+  "FOLDR_APPEND" > "List.foldr_append"
+  "FOLDR" > "HOL4Compat.FOLDR"
+  "FOLDL_SNOC_NIL" > "HOL4Base.rich_list.FOLDL_SNOC_NIL"
+  "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_FOLDR_REVERSE" > "List.foldl_foldr"
+  "FOLDL_FILTER" > "HOL4Base.rich_list.FOLDL_FILTER"
+  "FOLDL_APPEND" > "List.foldl_append"
+  "FOLDL" > "HOL4Compat.FOLDL"
+  "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_FLAT" > "HOL4Base.rich_list.FLAT_FLAT"
+  "FLAT_APPEND" > "List.concat_append"
+  "FLAT" > "HOL4Compat.FLAT"
+  "FIRSTN_SNOC" > "HOL4Base.rich_list.FIRSTN_SNOC"
+  "FIRSTN_SEG" > "HOL4Base.rich_list.FIRSTN_SEG"
+  "FIRSTN_REVERSE" > "HOL4Base.rich_list.FIRSTN_REVERSE"
+  "FIRSTN_LENGTH_ID" > "HOL4Base.rich_list.FIRSTN_LENGTH_ID"
+  "FIRSTN_LENGTH_APPEND" > "HOL4Base.rich_list.FIRSTN_LENGTH_APPEND"
+  "FIRSTN_FIRSTN" > "HOL4Base.rich_list.FIRSTN_FIRSTN"
+  "FIRSTN_BUTLASTN" > "HOL4Base.rich_list.FIRSTN_BUTLASTN"
+  "FIRSTN_APPEND2" > "HOL4Base.rich_list.FIRSTN_APPEND2"
+  "FIRSTN_APPEND1" > "HOL4Base.rich_list.FIRSTN_APPEND1"
+  "FIRSTN" > "HOL4Base.rich_list.FIRSTN"
+  "FILTER_SNOC" > "HOL4Base.rich_list.FILTER_SNOC"
+  "FILTER_REVERSE" > "HOL4Base.rich_list.FILTER_REVERSE"
+  "FILTER_MAP" > "HOL4Base.rich_list.FILTER_MAP"
+  "FILTER_IDEM" > "HOL4Base.rich_list.FILTER_IDEM"
+  "FILTER_FOLDR" > "HOL4Base.rich_list.FILTER_FOLDR"
+  "FILTER_FOLDL" > "HOL4Base.rich_list.FILTER_FOLDL"
+  "FILTER_FLAT" > "List.filter_concat"
+  "FILTER_FILTER" > "HOL4Base.rich_list.FILTER_FILTER"
+  "FILTER_COMM" > "HOL4Base.rich_list.FILTER_COMM"
+  "FILTER_APPEND" > "List.filter_append"
+  "FILTER" > "HOL4Compat.FILTER"
+  "FCOMM_FOLDR_FLAT" > "HOL4Base.rich_list.FCOMM_FOLDR_FLAT"
+  "FCOMM_FOLDR_APPEND" > "HOL4Base.rich_list.FCOMM_FOLDR_APPEND"
+  "FCOMM_FOLDL_FLAT" > "HOL4Base.rich_list.FCOMM_FOLDL_FLAT"
+  "FCOMM_FOLDL_APPEND" > "HOL4Base.rich_list.FCOMM_FOLDL_APPEND"
+  "EQ_LIST" > "HOL4Base.list.EQ_LIST"
+  "EL_SNOC" > "HOL4Base.rich_list.EL_SNOC"
+  "EL_SEG" > "HOL4Base.rich_list.EL_SEG"
+  "EL_REVERSE_ELL" > "HOL4Base.rich_list.EL_REVERSE_ELL"
+  "EL_REVERSE" > "HOL4Base.rich_list.EL_REVERSE"
+  "EL_PRE_LENGTH" > "HOL4Base.rich_list.EL_PRE_LENGTH"
+  "EL_MAP" > "HOL4Base.rich_list.EL_MAP"
+  "EL_LENGTH_SNOC" > "HOL4Base.rich_list.EL_LENGTH_SNOC"
+  "EL_LENGTH_APPEND" > "HOL4Base.rich_list.EL_LENGTH_APPEND"
+  "EL_IS_EL" > "HOL4Base.rich_list.EL_IS_EL"
+  "EL_ELL" > "HOL4Base.rich_list.EL_ELL"
+  "EL_CONS" > "HOL4Base.rich_list.EL_CONS"
+  "EL_APPEND2" > "HOL4Base.rich_list.EL_APPEND2"
+  "EL_APPEND1" > "HOL4Base.rich_list.EL_APPEND1"
+  "ELL_SUC_SNOC" > "HOL4Base.rich_list.ELL_SUC_SNOC"
+  "ELL_SNOC" > "HOL4Base.rich_list.ELL_SNOC"
+  "ELL_SEG" > "HOL4Base.rich_list.ELL_SEG"
+  "ELL_REVERSE_EL" > "HOL4Base.rich_list.ELL_REVERSE_EL"
+  "ELL_REVERSE" > "HOL4Base.rich_list.ELL_REVERSE"
+  "ELL_PRE_LENGTH" > "HOL4Base.rich_list.ELL_PRE_LENGTH"
+  "ELL_MAP" > "HOL4Base.rich_list.ELL_MAP"
+  "ELL_LENGTH_SNOC" > "HOL4Base.rich_list.ELL_LENGTH_SNOC"
+  "ELL_LENGTH_CONS" > "HOL4Base.rich_list.ELL_LENGTH_CONS"
+  "ELL_LENGTH_APPEND" > "HOL4Base.rich_list.ELL_LENGTH_APPEND"
+  "ELL_LAST" > "HOL4Base.rich_list.ELL_LAST"
+  "ELL_IS_EL" > "HOL4Base.rich_list.ELL_IS_EL"
+  "ELL_EL" > "HOL4Base.rich_list.ELL_EL"
+  "ELL_CONS" > "HOL4Base.rich_list.ELL_CONS"
+  "ELL_APPEND2" > "HOL4Base.rich_list.ELL_APPEND2"
+  "ELL_APPEND1" > "HOL4Base.rich_list.ELL_APPEND1"
+  "ELL_0_SNOC" > "HOL4Base.rich_list.ELL_0_SNOC"
+  "ELL" > "HOL4Base.rich_list.ELL"
+  "EL" > "HOL4Base.rich_list.EL"
+  "CONS_APPEND" > "HOL4Base.rich_list.CONS_APPEND"
+  "CONS_11" > "List.list.simps_3"
+  "CONS" > "HOL4Base.list.CONS"
+  "COMM_MONOID_FOLDR" > "HOL4Base.rich_list.COMM_MONOID_FOLDR"
+  "COMM_MONOID_FOLDL" > "HOL4Base.rich_list.COMM_MONOID_FOLDL"
+  "COMM_ASSOC_FOLDR_REVERSE" > "HOL4Base.rich_list.COMM_ASSOC_FOLDR_REVERSE"
+  "COMM_ASSOC_FOLDL_REVERSE" > "HOL4Base.rich_list.COMM_ASSOC_FOLDL_REVERSE"
+  "BUTLAST_CONS" > "HOL4Base.list.FRONT_CONS"
+  "BUTLASTN_SUC_BUTLAST" > "HOL4Base.rich_list.BUTLASTN_SUC_BUTLAST"
+  "BUTLASTN_SEG" > "HOL4Base.rich_list.BUTLASTN_SEG"
+  "BUTLASTN_REVERSE" > "HOL4Base.rich_list.BUTLASTN_REVERSE"
+  "BUTLASTN_MAP" > "HOL4Base.rich_list.BUTLASTN_MAP"
+  "BUTLASTN_LENGTH_NIL" > "HOL4Base.rich_list.BUTLASTN_LENGTH_NIL"
+  "BUTLASTN_LENGTH_CONS" > "HOL4Base.rich_list.BUTLASTN_LENGTH_CONS"
+  "BUTLASTN_LENGTH_APPEND" > "HOL4Base.rich_list.BUTLASTN_LENGTH_APPEND"
+  "BUTLASTN_LASTN_NIL" > "HOL4Base.rich_list.BUTLASTN_LASTN_NIL"
+  "BUTLASTN_LASTN" > "HOL4Base.rich_list.BUTLASTN_LASTN"
+  "BUTLASTN_FIRSTN" > "HOL4Base.rich_list.BUTLASTN_FIRSTN"
+  "BUTLASTN_CONS" > "HOL4Base.rich_list.BUTLASTN_CONS"
+  "BUTLASTN_BUTLASTN" > "HOL4Base.rich_list.BUTLASTN_BUTLASTN"
+  "BUTLASTN_BUTLAST" > "HOL4Base.rich_list.BUTLASTN_BUTLAST"
+  "BUTLASTN_APPEND2" > "HOL4Base.rich_list.BUTLASTN_APPEND2"
+  "BUTLASTN_APPEND1" > "HOL4Base.rich_list.BUTLASTN_APPEND1"
+  "BUTLASTN_1" > "HOL4Base.rich_list.BUTLASTN_1"
+  "BUTLASTN" > "HOL4Base.rich_list.BUTLASTN"
+  "BUTLAST" > "HOL4Base.rich_list.BUTLAST"
+  "BUTFIRSTN_SNOC" > "HOL4Base.rich_list.BUTFIRSTN_SNOC"
+  "BUTFIRSTN_SEG" > "HOL4Base.rich_list.BUTFIRSTN_SEG"
+  "BUTFIRSTN_REVERSE" > "HOL4Base.rich_list.BUTFIRSTN_REVERSE"
+  "BUTFIRSTN_LENGTH_NIL" > "HOL4Base.rich_list.BUTFIRSTN_LENGTH_NIL"
+  "BUTFIRSTN_LENGTH_APPEND" > "HOL4Base.rich_list.BUTFIRSTN_LENGTH_APPEND"
+  "BUTFIRSTN_LASTN" > "HOL4Base.rich_list.BUTFIRSTN_LASTN"
+  "BUTFIRSTN_BUTFIRSTN" > "HOL4Base.rich_list.BUTFIRSTN_BUTFIRSTN"
+  "BUTFIRSTN_APPEND2" > "HOL4Base.rich_list.BUTFIRSTN_APPEND2"
+  "BUTFIRSTN_APPEND1" > "HOL4Base.rich_list.BUTFIRSTN_APPEND1"
+  "BUTFIRSTN" > "HOL4Base.rich_list.BUTFIRSTN"
+  "ASSOC_FOLDR_FLAT" > "HOL4Base.rich_list.ASSOC_FOLDR_FLAT"
+  "ASSOC_FOLDL_FLAT" > "HOL4Base.rich_list.ASSOC_FOLDL_FLAT"
+  "ASSOC_APPEND" > "HOL4Base.rich_list.ASSOC_APPEND"
+  "APPEND_SNOC" > "HOL4Base.rich_list.APPEND_SNOC"
+  "APPEND_NIL" > "HOL4Base.rich_list.APPEND_NIL"
+  "APPEND_LENGTH_EQ" > "HOL4Base.rich_list.APPEND_LENGTH_EQ"
+  "APPEND_FOLDR" > "HOL4Base.rich_list.APPEND_FOLDR"
+  "APPEND_FOLDL" > "HOL4Base.rich_list.APPEND_FOLDL"
+  "APPEND_FIRSTN_LASTN" > "HOL4Base.rich_list.APPEND_FIRSTN_LASTN"
+  "APPEND_FIRSTN_BUTFIRSTN" > "HOL4Base.rich_list.APPEND_FIRSTN_BUTFIRSTN"
+  "APPEND_BUTLAST_LAST" > "List.append_butlast_last_id"
+  "APPEND_BUTLASTN_LASTN" > "HOL4Base.rich_list.APPEND_BUTLASTN_LASTN"
+  "APPEND_BUTLASTN_BUTFIRSTN" > "HOL4Base.rich_list.APPEND_BUTLASTN_BUTFIRSTN"
+  "APPEND_ASSOC" > "List.append_assoc"
+  "APPEND" > "HOL4Compat.APPEND"
+  "AND_EL_def" > "HOL4Base.rich_list.AND_EL_def"
+  "AND_EL_FOLDR" > "HOL4Base.rich_list.AND_EL_FOLDR"
+  "AND_EL_FOLDL" > "HOL4Base.rich_list.AND_EL_FOLDL"
+  "AND_EL_DEF" > "HOL4Base.rich_list.AND_EL_DEF"
+  "ALL_EL_SNOC" > "HOL4Base.rich_list.ALL_EL_SNOC"
+  "ALL_EL_SEG" > "HOL4Base.rich_list.ALL_EL_SEG"
+  "ALL_EL_REVERSE" > "HOL4Base.rich_list.ALL_EL_REVERSE"
+  "ALL_EL_REPLICATE" > "HOL4Base.rich_list.ALL_EL_REPLICATE"
+  "ALL_EL_MAP" > "HOL4Base.rich_list.ALL_EL_MAP"
+  "ALL_EL_LASTN" > "HOL4Base.rich_list.ALL_EL_LASTN"
+  "ALL_EL_FOLDR_MAP" > "HOL4Base.rich_list.ALL_EL_FOLDR_MAP"
+  "ALL_EL_FOLDR" > "HOL4Base.rich_list.ALL_EL_FOLDR"
+  "ALL_EL_FOLDL_MAP" > "HOL4Base.rich_list.ALL_EL_FOLDL_MAP"
+  "ALL_EL_FOLDL" > "HOL4Base.rich_list.ALL_EL_FOLDL"
+  "ALL_EL_FIRSTN" > "HOL4Base.rich_list.ALL_EL_FIRSTN"
+  "ALL_EL_CONJ" > "HOL4Base.list.EVERY_CONJ"
+  "ALL_EL_BUTLASTN" > "HOL4Base.rich_list.ALL_EL_BUTLASTN"
+  "ALL_EL_BUTFIRSTN" > "HOL4Base.rich_list.ALL_EL_BUTFIRSTN"
+  "ALL_EL_APPEND" > "List.list_all_append"
+  "ALL_EL" > "HOL4Compat.EVERY_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/seq.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,109 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "sums" > "sums_def"
+  "summable" > "summable_def"
+  "suminf" > "suminf_def"
+  "subseq" > "subseq_def"
+  "mono" > "mono_def"
+  "lim" > "lim_def"
+  "convergent" > "convergent_def"
+  "cauchy" > "cauchy_def"
+  "-->" > "-->_def"
+
+const_maps
+  "sums" > "HOL4Real.seq.sums"
+  "summable" > "HOL4Real.seq.summable"
+  "suminf" > "HOL4Real.seq.suminf"
+  "subseq" > "HOL4Real.seq.subseq"
+  "mono" > "HOL4Real.seq.mono"
+  "lim" > "HOL4Real.seq.lim"
+  "convergent" > "HOL4Real.seq.convergent"
+  "cauchy" > "HOL4Real.seq.cauchy"
+  "-->" > "HOL4Real.seq.-->"
+
+thm_maps
+  "tends_num_real" > "HOL4Real.seq.tends_num_real"
+  "sums_def" > "HOL4Real.seq.sums_def"
+  "sums" > "HOL4Real.seq.sums"
+  "summable_def" > "HOL4Real.seq.summable_def"
+  "summable" > "HOL4Real.seq.summable"
+  "suminf_def" > "HOL4Real.seq.suminf_def"
+  "suminf" > "HOL4Real.seq.suminf"
+  "subseq_def" > "HOL4Real.seq.subseq_def"
+  "subseq" > "HOL4Real.seq.subseq"
+  "mono_def" > "HOL4Real.seq.mono_def"
+  "mono" > "HOL4Real.seq.mono"
+  "lim_def" > "HOL4Real.seq.lim_def"
+  "lim" > "HOL4Real.seq.lim"
+  "convergent_def" > "HOL4Real.seq.convergent_def"
+  "convergent" > "HOL4Real.seq.convergent"
+  "cauchy_def" > "HOL4Real.seq.cauchy_def"
+  "cauchy" > "HOL4Real.seq.cauchy"
+  "SUM_UNIQ" > "HOL4Real.seq.SUM_UNIQ"
+  "SUM_SUMMABLE" > "HOL4Real.seq.SUM_SUMMABLE"
+  "SUMMABLE_SUM" > "HOL4Real.seq.SUMMABLE_SUM"
+  "SUBSEQ_SUC" > "HOL4Real.seq.SUBSEQ_SUC"
+  "SER_ZERO" > "HOL4Real.seq.SER_ZERO"
+  "SER_SUB" > "HOL4Real.seq.SER_SUB"
+  "SER_RATIO" > "HOL4Real.seq.SER_RATIO"
+  "SER_POS_LT_PAIR" > "HOL4Real.seq.SER_POS_LT_PAIR"
+  "SER_POS_LT" > "HOL4Real.seq.SER_POS_LT"
+  "SER_POS_LE" > "HOL4Real.seq.SER_POS_LE"
+  "SER_PAIR" > "HOL4Real.seq.SER_PAIR"
+  "SER_OFFSET" > "HOL4Real.seq.SER_OFFSET"
+  "SER_NEG" > "HOL4Real.seq.SER_NEG"
+  "SER_LE2" > "HOL4Real.seq.SER_LE2"
+  "SER_LE" > "HOL4Real.seq.SER_LE"
+  "SER_GROUP" > "HOL4Real.seq.SER_GROUP"
+  "SER_COMPARA" > "HOL4Real.seq.SER_COMPARA"
+  "SER_COMPAR" > "HOL4Real.seq.SER_COMPAR"
+  "SER_CMUL" > "HOL4Real.seq.SER_CMUL"
+  "SER_CDIV" > "HOL4Real.seq.SER_CDIV"
+  "SER_CAUCHY" > "HOL4Real.seq.SER_CAUCHY"
+  "SER_ADD" > "HOL4Real.seq.SER_ADD"
+  "SER_ACONV" > "HOL4Real.seq.SER_ACONV"
+  "SER_ABS" > "HOL4Real.seq.SER_ABS"
+  "SER_0" > "HOL4Real.seq.SER_0"
+  "SEQ_UNIQ" > "HOL4Real.seq.SEQ_UNIQ"
+  "SEQ_SUC" > "HOL4Real.seq.SEQ_SUC"
+  "SEQ_SUBLE" > "HOL4Real.seq.SEQ_SUBLE"
+  "SEQ_SUB" > "HOL4Real.seq.SEQ_SUB"
+  "SEQ_SBOUNDED" > "HOL4Real.seq.SEQ_SBOUNDED"
+  "SEQ_POWER_ABS" > "HOL4Real.seq.SEQ_POWER_ABS"
+  "SEQ_POWER" > "HOL4Real.seq.SEQ_POWER"
+  "SEQ_NEG_CONV" > "HOL4Real.seq.SEQ_NEG_CONV"
+  "SEQ_NEG_BOUNDED" > "HOL4Real.seq.SEQ_NEG_BOUNDED"
+  "SEQ_NEG" > "HOL4Real.seq.SEQ_NEG"
+  "SEQ_MUL" > "HOL4Real.seq.SEQ_MUL"
+  "SEQ_MONOSUB" > "HOL4Real.seq.SEQ_MONOSUB"
+  "SEQ_LIM" > "HOL4Real.seq.SEQ_LIM"
+  "SEQ_LE" > "HOL4Real.seq.SEQ_LE"
+  "SEQ_INV0" > "HOL4Real.seq.SEQ_INV0"
+  "SEQ_INV" > "HOL4Real.seq.SEQ_INV"
+  "SEQ_ICONV" > "HOL4Real.seq.SEQ_ICONV"
+  "SEQ_DIV" > "HOL4Real.seq.SEQ_DIV"
+  "SEQ_DIRECT" > "HOL4Real.seq.SEQ_DIRECT"
+  "SEQ_CONST" > "HOL4Real.seq.SEQ_CONST"
+  "SEQ_CBOUNDED" > "HOL4Real.seq.SEQ_CBOUNDED"
+  "SEQ_CAUCHY" > "HOL4Real.seq.SEQ_CAUCHY"
+  "SEQ_BOUNDED_2" > "HOL4Real.seq.SEQ_BOUNDED_2"
+  "SEQ_BOUNDED" > "HOL4Real.seq.SEQ_BOUNDED"
+  "SEQ_BCONV" > "HOL4Real.seq.SEQ_BCONV"
+  "SEQ_ADD" > "HOL4Real.seq.SEQ_ADD"
+  "SEQ_ABS_IMP" > "HOL4Real.seq.SEQ_ABS_IMP"
+  "SEQ_ABS" > "HOL4Real.seq.SEQ_ABS"
+  "SEQ" > "HOL4Real.seq.SEQ"
+  "NEST_LEMMA_UNIQ" > "HOL4Real.seq.NEST_LEMMA_UNIQ"
+  "NEST_LEMMA" > "HOL4Real.seq.NEST_LEMMA"
+  "MONO_SUC" > "HOL4Real.seq.MONO_SUC"
+  "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"
+  "-->_def" > "HOL4Real.seq.-->_def"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/state_transformer.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,42 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "UNIT" > "UNIT_def"
+  "MMAP" > "MMAP_def"
+  "JOIN" > "JOIN_def"
+  "BIND" > "BIND_def"
+
+const_maps
+  "UNIT" > "HOL4Base.state_transformer.UNIT"
+  "MMAP" > "HOL4Base.state_transformer.MMAP"
+  "JOIN" > "HOL4Base.state_transformer.JOIN"
+  "BIND" > "HOL4Base.state_transformer.BIND"
+
+thm_maps
+  "UNIT_def" > "HOL4Base.state_transformer.UNIT_def"
+  "UNIT_UNCURRY" > "HOL4Base.state_transformer.UNIT_UNCURRY"
+  "UNIT_DEF" > "HOL4Base.state_transformer.UNIT_DEF"
+  "SND_o_UNIT" > "HOL4Base.state_transformer.SND_o_UNIT"
+  "MMAP_def" > "HOL4Base.state_transformer.MMAP_def"
+  "MMAP_UNIT" > "HOL4Base.state_transformer.MMAP_UNIT"
+  "MMAP_JOIN" > "HOL4Base.state_transformer.MMAP_JOIN"
+  "MMAP_ID" > "HOL4Base.state_transformer.MMAP_ID"
+  "MMAP_DEF" > "HOL4Base.state_transformer.MMAP_DEF"
+  "MMAP_COMP" > "HOL4Base.state_transformer.MMAP_COMP"
+  "JOIN_def" > "HOL4Base.state_transformer.JOIN_def"
+  "JOIN_UNIT" > "HOL4Base.state_transformer.JOIN_UNIT"
+  "JOIN_MMAP_UNIT" > "HOL4Base.state_transformer.JOIN_MMAP_UNIT"
+  "JOIN_MAP_JOIN" > "HOL4Base.state_transformer.JOIN_MAP_JOIN"
+  "JOIN_MAP" > "HOL4Base.state_transformer.JOIN_MAP"
+  "JOIN_DEF" > "HOL4Base.state_transformer.JOIN_DEF"
+  "FST_o_UNIT" > "HOL4Base.state_transformer.FST_o_UNIT"
+  "FST_o_MMAP" > "HOL4Base.state_transformer.FST_o_MMAP"
+  "BIND_def" > "HOL4Base.state_transformer.BIND_def"
+  "BIND_RIGHT_UNIT" > "HOL4Base.state_transformer.BIND_RIGHT_UNIT"
+  "BIND_LEFT_UNIT" > "HOL4Base.state_transformer.BIND_LEFT_UNIT"
+  "BIND_DEF" > "HOL4Base.state_transformer.BIND_DEF"
+  "BIND_ASSOC" > "HOL4Base.state_transformer.BIND_ASSOC"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/sum.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,43 @@
+import
+
+import_segment "hol4"
+
+type_maps
+  "sum" > "+"
+
+const_maps
+  "sum_case" > "Datatype.sum.sum_case"
+  "OUTR" > "HOL4Compat.OUTR"
+  "OUTL" > "HOL4Compat.OUTL"
+  "ISR" > "HOL4Compat.ISR"
+  "ISL" > "HOL4Compat.ISL"
+  "INR" > "Inr"
+  "INL" > "Inl"
+
+thm_maps
+  "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"
+  "OUTR" > "HOL4Compat.OUTR"
+  "OUTL" > "HOL4Compat.OUTL"
+  "ISR" > "HOL4Compat.ISR"
+  "ISL_OR_ISR" > "HOL4Base.sum.ISL_OR_ISR"
+  "ISL" > "HOL4Compat.ISL"
+  "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"
+  "IS_SUM_REP"
+  "INR_DEF"
+  "INL_DEF"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/topology.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,116 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "topology" > "topology_def"
+  "re_universe" > "re_universe_def"
+  "re_union" > "re_union_def"
+  "re_subset" > "re_subset_def"
+  "re_null" > "re_null_def"
+  "re_intersect" > "re_intersect_def"
+  "re_compl" > "re_compl_def"
+  "re_Union" > "re_Union_def"
+  "open" > "open_def"
+  "neigh" > "neigh_def"
+  "mtop" > "mtop_def"
+  "mr1" > "mr1_def"
+  "metric" > "metric_def"
+  "limpt" > "limpt_def"
+  "istopology" > "istopology_def"
+  "ismet" > "ismet_def"
+  "dist" > "dist_def"
+  "closed" > "closed_def"
+  "B" > "B_def"
+
+type_maps
+  "topology" > "HOL4Real.topology.topology"
+  "metric" > "HOL4Real.topology.metric"
+
+const_maps
+  "re_universe" > "HOL4Real.topology.re_universe"
+  "re_union" > "HOL4Real.topology.re_union"
+  "re_subset" > "HOL4Real.topology.re_subset"
+  "re_null" > "HOL4Real.topology.re_null"
+  "re_intersect" > "HOL4Real.topology.re_intersect"
+  "re_compl" > "HOL4Real.topology.re_compl"
+  "re_Union" > "HOL4Real.topology.re_Union"
+  "neigh" > "HOL4Real.topology.neigh"
+  "mtop" > "HOL4Real.topology.mtop"
+  "mr1" > "HOL4Real.topology.mr1"
+  "limpt" > "HOL4Real.topology.limpt"
+  "istopology" > "HOL4Real.topology.istopology"
+  "ismet" > "HOL4Real.topology.ismet"
+  "closed" > "HOL4Real.topology.closed"
+  "B" > "HOL4Real.topology.B"
+
+thm_maps
+  "topology_tybij" > "HOL4Real.topology.topology_tybij"
+  "topology_TY_DEF" > "HOL4Real.topology.topology_TY_DEF"
+  "re_universe_def" > "HOL4Real.topology.re_universe_def"
+  "re_universe" > "HOL4Real.topology.re_universe"
+  "re_union_def" > "HOL4Real.topology.re_union_def"
+  "re_union" > "HOL4Real.topology.re_union"
+  "re_subset_def" > "HOL4Real.topology.re_subset_def"
+  "re_subset" > "HOL4Real.topology.re_subset"
+  "re_null_def" > "HOL4Real.topology.re_null_def"
+  "re_null" > "HOL4Real.topology.re_null"
+  "re_intersect_def" > "HOL4Real.topology.re_intersect_def"
+  "re_intersect" > "HOL4Real.topology.re_intersect"
+  "re_compl_def" > "HOL4Real.topology.re_compl_def"
+  "re_compl" > "HOL4Real.topology.re_compl"
+  "re_Union_def" > "HOL4Real.topology.re_Union_def"
+  "re_Union" > "HOL4Real.topology.re_Union"
+  "neigh_def" > "HOL4Real.topology.neigh_def"
+  "neigh" > "HOL4Real.topology.neigh"
+  "mtop_istopology" > "HOL4Real.topology.mtop_istopology"
+  "mtop_def" > "HOL4Real.topology.mtop_def"
+  "mtop" > "HOL4Real.topology.mtop"
+  "mr1_def" > "HOL4Real.topology.mr1_def"
+  "mr1" > "HOL4Real.topology.mr1"
+  "metric_tybij" > "HOL4Real.topology.metric_tybij"
+  "metric_TY_DEF" > "HOL4Real.topology.metric_TY_DEF"
+  "limpt_def" > "HOL4Real.topology.limpt_def"
+  "limpt" > "HOL4Real.topology.limpt"
+  "istopology_def" > "HOL4Real.topology.istopology_def"
+  "istopology" > "HOL4Real.topology.istopology"
+  "ismet_def" > "HOL4Real.topology.ismet_def"
+  "ismet" > "HOL4Real.topology.ismet"
+  "closed_def" > "HOL4Real.topology.closed_def"
+  "closed" > "HOL4Real.topology.closed"
+  "ball" > "HOL4Real.topology.ball"
+  "TOPOLOGY_UNION" > "HOL4Real.topology.TOPOLOGY_UNION"
+  "TOPOLOGY" > "HOL4Real.topology.TOPOLOGY"
+  "SUBSET_TRANS" > "HOL4Real.topology.SUBSET_TRANS"
+  "SUBSET_REFL" > "HOL4Real.topology.SUBSET_REFL"
+  "SUBSET_ANTISYM" > "HOL4Real.topology.SUBSET_ANTISYM"
+  "OPEN_UNOPEN" > "HOL4Real.topology.OPEN_UNOPEN"
+  "OPEN_SUBOPEN" > "HOL4Real.topology.OPEN_SUBOPEN"
+  "OPEN_OWN_NEIGH" > "HOL4Real.topology.OPEN_OWN_NEIGH"
+  "OPEN_NEIGH" > "HOL4Real.topology.OPEN_NEIGH"
+  "MTOP_OPEN" > "HOL4Real.topology.MTOP_OPEN"
+  "MTOP_LIMPT" > "HOL4Real.topology.MTOP_LIMPT"
+  "MR1_SUB_LT" > "HOL4Real.topology.MR1_SUB_LT"
+  "MR1_SUB_LE" > "HOL4Real.topology.MR1_SUB_LE"
+  "MR1_SUB" > "HOL4Real.topology.MR1_SUB"
+  "MR1_LIMPT" > "HOL4Real.topology.MR1_LIMPT"
+  "MR1_DEF" > "HOL4Real.topology.MR1_DEF"
+  "MR1_BETWEEN1" > "HOL4Real.topology.MR1_BETWEEN1"
+  "MR1_ADD_POS" > "HOL4Real.topology.MR1_ADD_POS"
+  "MR1_ADD_LT" > "HOL4Real.topology.MR1_ADD_LT"
+  "MR1_ADD" > "HOL4Real.topology.MR1_ADD"
+  "METRIC_ZERO" > "HOL4Real.topology.METRIC_ZERO"
+  "METRIC_TRIANGLE" > "HOL4Real.topology.METRIC_TRIANGLE"
+  "METRIC_SYM" > "HOL4Real.topology.METRIC_SYM"
+  "METRIC_SAME" > "HOL4Real.topology.METRIC_SAME"
+  "METRIC_POS" > "HOL4Real.topology.METRIC_POS"
+  "METRIC_NZ" > "HOL4Real.topology.METRIC_NZ"
+  "METRIC_ISMET" > "HOL4Real.topology.METRIC_ISMET"
+  "ISMET_R1" > "HOL4Real.topology.ISMET_R1"
+  "COMPL_MEM" > "HOL4Real.topology.COMPL_MEM"
+  "CLOSED_LIMPT" > "HOL4Real.topology.CLOSED_LIMPT"
+  "B_def" > "HOL4Real.topology.B_def"
+  "BALL_OPEN" > "HOL4Real.topology.BALL_OPEN"
+  "BALL_NEIGH" > "HOL4Real.topology.BALL_NEIGH"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/transc.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,273 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "tdiv" > "tdiv_def"
+  "tan" > "tan_def"
+  "sqrt" > "sqrt_def"
+  "sin" > "sin_def"
+  "rsum" > "rsum_def"
+  "root" > "root_def"
+  "pi" > "pi_def"
+  "ln" > "ln_def"
+  "gauge" > "gauge_def"
+  "fine" > "fine_def"
+  "exp" > "exp_def"
+  "dsize" > "dsize_def"
+  "division" > "division_def"
+  "cos" > "cos_def"
+  "atn" > "atn_def"
+  "asn" > "asn_def"
+  "acs" > "acs_def"
+  "Dint" > "Dint_def"
+
+const_maps
+  "tdiv" > "HOL4Real.transc.tdiv"
+  "tan" > "HOL4Real.transc.tan"
+  "sqrt" > "HOL4Real.transc.sqrt"
+  "sin" > "HOL4Real.transc.sin"
+  "rsum" > "HOL4Real.transc.rsum"
+  "root" > "HOL4Real.transc.root"
+  "pi" > "HOL4Real.transc.pi"
+  "ln" > "HOL4Real.transc.ln"
+  "gauge" > "HOL4Real.transc.gauge"
+  "fine" > "HOL4Real.transc.fine"
+  "exp" > "HOL4Real.transc.exp"
+  "dsize" > "HOL4Real.transc.dsize"
+  "division" > "HOL4Real.transc.division"
+  "cos" > "HOL4Real.transc.cos"
+  "atn" > "HOL4Real.transc.atn"
+  "asn" > "HOL4Real.transc.asn"
+  "acs" > "HOL4Real.transc.acs"
+  "Dint" > "HOL4Real.transc.Dint"
+
+thm_maps
+  "tdiv_def" > "HOL4Real.transc.tdiv_def"
+  "tdiv" > "HOL4Real.transc.tdiv"
+  "tan_def" > "HOL4Real.transc.tan_def"
+  "tan" > "HOL4Real.transc.tan"
+  "sqrt_def" > "HOL4Real.transc.sqrt_def"
+  "sqrt" > "HOL4Real.transc.sqrt"
+  "sin_def" > "HOL4Real.transc.sin_def"
+  "sin" > "HOL4Real.transc.sin"
+  "rsum_def" > "HOL4Real.transc.rsum_def"
+  "rsum" > "HOL4Real.transc.rsum"
+  "root_def" > "HOL4Real.transc.root_def"
+  "root" > "HOL4Real.transc.root"
+  "pi_def" > "HOL4Real.transc.pi_def"
+  "pi" > "HOL4Real.transc.pi"
+  "ln_def" > "HOL4Real.transc.ln_def"
+  "ln" > "HOL4Real.transc.ln"
+  "gauge_def" > "HOL4Real.transc.gauge_def"
+  "gauge" > "HOL4Real.transc.gauge"
+  "fine_def" > "HOL4Real.transc.fine_def"
+  "fine" > "HOL4Real.transc.fine"
+  "exp_def" > "HOL4Real.transc.exp_def"
+  "exp" > "HOL4Real.transc.exp"
+  "dsize_def" > "HOL4Real.transc.dsize_def"
+  "dsize" > "HOL4Real.transc.dsize"
+  "division_def" > "HOL4Real.transc.division_def"
+  "division" > "HOL4Real.transc.division"
+  "cos_def" > "HOL4Real.transc.cos_def"
+  "cos" > "HOL4Real.transc.cos"
+  "atn_def" > "HOL4Real.transc.atn_def"
+  "atn" > "HOL4Real.transc.atn"
+  "asn_def" > "HOL4Real.transc.asn_def"
+  "asn" > "HOL4Real.transc.asn"
+  "acs_def" > "HOL4Real.transc.acs_def"
+  "acs" > "HOL4Real.transc.acs"
+  "TAN_TOTAL_POS" > "HOL4Real.transc.TAN_TOTAL_POS"
+  "TAN_TOTAL_LEMMA" > "HOL4Real.transc.TAN_TOTAL_LEMMA"
+  "TAN_TOTAL" > "HOL4Real.transc.TAN_TOTAL"
+  "TAN_SEC" > "HOL4Real.transc.TAN_SEC"
+  "TAN_POS_PI2" > "HOL4Real.transc.TAN_POS_PI2"
+  "TAN_PI" > "HOL4Real.transc.TAN_PI"
+  "TAN_PERIODIC" > "HOL4Real.transc.TAN_PERIODIC"
+  "TAN_NPI" > "HOL4Real.transc.TAN_NPI"
+  "TAN_NEG" > "HOL4Real.transc.TAN_NEG"
+  "TAN_DOUBLE" > "HOL4Real.transc.TAN_DOUBLE"
+  "TAN_ATN" > "HOL4Real.transc.TAN_ATN"
+  "TAN_ADD" > "HOL4Real.transc.TAN_ADD"
+  "TAN_0" > "HOL4Real.transc.TAN_0"
+  "SQRT_POW_2" > "HOL4Real.transc.SQRT_POW_2"
+  "SQRT_POW2" > "HOL4Real.transc.SQRT_POW2"
+  "SQRT_POS_UNIQ" > "HOL4Real.transc.SQRT_POS_UNIQ"
+  "SQRT_POS_LT" > "HOL4Real.transc.SQRT_POS_LT"
+  "SQRT_POS_LE" > "HOL4Real.transc.SQRT_POS_LE"
+  "SQRT_MUL" > "HOL4Real.transc.SQRT_MUL"
+  "SQRT_MONO_LE" > "HOL4Real.transc.SQRT_MONO_LE"
+  "SQRT_INV" > "HOL4Real.transc.SQRT_INV"
+  "SQRT_EVEN_POW2" > "HOL4Real.transc.SQRT_EVEN_POW2"
+  "SQRT_EQ" > "HOL4Real.transc.SQRT_EQ"
+  "SQRT_DIV" > "HOL4Real.transc.SQRT_DIV"
+  "SQRT_1" > "HOL4Real.transc.SQRT_1"
+  "SQRT_0" > "HOL4Real.transc.SQRT_0"
+  "SIN_ZERO_LEMMA" > "HOL4Real.transc.SIN_ZERO_LEMMA"
+  "SIN_ZERO" > "HOL4Real.transc.SIN_ZERO"
+  "SIN_TOTAL" > "HOL4Real.transc.SIN_TOTAL"
+  "SIN_POS_PI_LE" > "HOL4Real.transc.SIN_POS_PI_LE"
+  "SIN_POS_PI2_LE" > "HOL4Real.transc.SIN_POS_PI2_LE"
+  "SIN_POS_PI2" > "HOL4Real.transc.SIN_POS_PI2"
+  "SIN_POS_PI" > "HOL4Real.transc.SIN_POS_PI"
+  "SIN_POS" > "HOL4Real.transc.SIN_POS"
+  "SIN_PI2" > "HOL4Real.transc.SIN_PI2"
+  "SIN_PI" > "HOL4Real.transc.SIN_PI"
+  "SIN_PERIODIC_PI" > "HOL4Real.transc.SIN_PERIODIC_PI"
+  "SIN_PERIODIC" > "HOL4Real.transc.SIN_PERIODIC"
+  "SIN_PAIRED" > "HOL4Real.transc.SIN_PAIRED"
+  "SIN_NPI" > "HOL4Real.transc.SIN_NPI"
+  "SIN_NEGLEMMA" > "HOL4Real.transc.SIN_NEGLEMMA"
+  "SIN_NEG" > "HOL4Real.transc.SIN_NEG"
+  "SIN_FDIFF" > "HOL4Real.transc.SIN_FDIFF"
+  "SIN_DOUBLE" > "HOL4Real.transc.SIN_DOUBLE"
+  "SIN_COS_SQRT" > "HOL4Real.transc.SIN_COS_SQRT"
+  "SIN_COS_SQ" > "HOL4Real.transc.SIN_COS_SQ"
+  "SIN_COS_NEG" > "HOL4Real.transc.SIN_COS_NEG"
+  "SIN_COS_ADD" > "HOL4Real.transc.SIN_COS_ADD"
+  "SIN_COS" > "HOL4Real.transc.SIN_COS"
+  "SIN_CONVERGES" > "HOL4Real.transc.SIN_CONVERGES"
+  "SIN_CIRCLE" > "HOL4Real.transc.SIN_CIRCLE"
+  "SIN_BOUNDS" > "HOL4Real.transc.SIN_BOUNDS"
+  "SIN_BOUND" > "HOL4Real.transc.SIN_BOUND"
+  "SIN_ASN" > "HOL4Real.transc.SIN_ASN"
+  "SIN_ADD" > "HOL4Real.transc.SIN_ADD"
+  "SIN_ACS_NZ" > "HOL4Real.transc.SIN_ACS_NZ"
+  "SIN_0" > "HOL4Real.transc.SIN_0"
+  "ROOT_POW_POS" > "HOL4Real.transc.ROOT_POW_POS"
+  "ROOT_POS_UNIQ" > "HOL4Real.transc.ROOT_POS_UNIQ"
+  "ROOT_POS_LT" > "HOL4Real.transc.ROOT_POS_LT"
+  "ROOT_POS" > "HOL4Real.transc.ROOT_POS"
+  "ROOT_MUL" > "HOL4Real.transc.ROOT_MUL"
+  "ROOT_MONO_LE" > "HOL4Real.transc.ROOT_MONO_LE"
+  "ROOT_LT_LEMMA" > "HOL4Real.transc.ROOT_LT_LEMMA"
+  "ROOT_LN" > "HOL4Real.transc.ROOT_LN"
+  "ROOT_INV" > "HOL4Real.transc.ROOT_INV"
+  "ROOT_DIV" > "HOL4Real.transc.ROOT_DIV"
+  "ROOT_1" > "HOL4Real.transc.ROOT_1"
+  "ROOT_0" > "HOL4Real.transc.ROOT_0"
+  "REAL_DIV_SQRT" > "HOL4Real.transc.REAL_DIV_SQRT"
+  "POW_ROOT_POS" > "HOL4Real.transc.POW_ROOT_POS"
+  "POW_2_SQRT" > "HOL4Real.transc.POW_2_SQRT"
+  "PI_POS" > "HOL4Real.transc.PI_POS"
+  "PI2_BOUNDS" > "HOL4Real.transc.PI2_BOUNDS"
+  "PI2" > "HOL4Real.transc.PI2"
+  "MCLAURIN_ZERO" > "HOL4Real.transc.MCLAURIN_ZERO"
+  "MCLAURIN_NEG" > "HOL4Real.transc.MCLAURIN_NEG"
+  "MCLAURIN_EXP_LT" > "HOL4Real.transc.MCLAURIN_EXP_LT"
+  "MCLAURIN_EXP_LE" > "HOL4Real.transc.MCLAURIN_EXP_LE"
+  "MCLAURIN_ALL_LT" > "HOL4Real.transc.MCLAURIN_ALL_LT"
+  "MCLAURIN_ALL_LE" > "HOL4Real.transc.MCLAURIN_ALL_LE"
+  "MCLAURIN" > "HOL4Real.transc.MCLAURIN"
+  "LN_POW" > "HOL4Real.transc.LN_POW"
+  "LN_POS" > "HOL4Real.transc.LN_POS"
+  "LN_MUL" > "HOL4Real.transc.LN_MUL"
+  "LN_MONO_LT" > "HOL4Real.transc.LN_MONO_LT"
+  "LN_MONO_LE" > "HOL4Real.transc.LN_MONO_LE"
+  "LN_LT_X" > "HOL4Real.transc.LN_LT_X"
+  "LN_LE" > "HOL4Real.transc.LN_LE"
+  "LN_INV" > "HOL4Real.transc.LN_INV"
+  "LN_INJ" > "HOL4Real.transc.LN_INJ"
+  "LN_EXP" > "HOL4Real.transc.LN_EXP"
+  "LN_DIV" > "HOL4Real.transc.LN_DIV"
+  "LN_1" > "HOL4Real.transc.LN_1"
+  "INTEGRAL_NULL" > "HOL4Real.transc.INTEGRAL_NULL"
+  "GAUGE_MIN" > "HOL4Real.transc.GAUGE_MIN"
+  "FTC1" > "HOL4Real.transc.FTC1"
+  "FINE_MIN" > "HOL4Real.transc.FINE_MIN"
+  "EXP_TOTAL_LEMMA" > "HOL4Real.transc.EXP_TOTAL_LEMMA"
+  "EXP_TOTAL" > "HOL4Real.transc.EXP_TOTAL"
+  "EXP_SUB" > "HOL4Real.transc.EXP_SUB"
+  "EXP_POS_LT" > "HOL4Real.transc.EXP_POS_LT"
+  "EXP_POS_LE" > "HOL4Real.transc.EXP_POS_LE"
+  "EXP_NZ" > "HOL4Real.transc.EXP_NZ"
+  "EXP_NEG_MUL2" > "HOL4Real.transc.EXP_NEG_MUL2"
+  "EXP_NEG_MUL" > "HOL4Real.transc.EXP_NEG_MUL"
+  "EXP_NEG" > "HOL4Real.transc.EXP_NEG"
+  "EXP_N" > "HOL4Real.transc.EXP_N"
+  "EXP_MONO_LT" > "HOL4Real.transc.EXP_MONO_LT"
+  "EXP_MONO_LE" > "HOL4Real.transc.EXP_MONO_LE"
+  "EXP_MONO_IMP" > "HOL4Real.transc.EXP_MONO_IMP"
+  "EXP_LT_1" > "HOL4Real.transc.EXP_LT_1"
+  "EXP_LN" > "HOL4Real.transc.EXP_LN"
+  "EXP_LE_X" > "HOL4Real.transc.EXP_LE_X"
+  "EXP_INJ" > "HOL4Real.transc.EXP_INJ"
+  "EXP_FDIFF" > "HOL4Real.transc.EXP_FDIFF"
+  "EXP_CONVERGES" > "HOL4Real.transc.EXP_CONVERGES"
+  "EXP_ADD_MUL" > "HOL4Real.transc.EXP_ADD_MUL"
+  "EXP_ADD" > "HOL4Real.transc.EXP_ADD"
+  "EXP_0" > "HOL4Real.transc.EXP_0"
+  "Dint_def" > "HOL4Real.transc.Dint_def"
+  "Dint" > "HOL4Real.transc.Dint"
+  "DIVISION_UBOUND_LT" > "HOL4Real.transc.DIVISION_UBOUND_LT"
+  "DIVISION_UBOUND" > "HOL4Real.transc.DIVISION_UBOUND"
+  "DIVISION_THM" > "HOL4Real.transc.DIVISION_THM"
+  "DIVISION_SINGLE" > "HOL4Real.transc.DIVISION_SINGLE"
+  "DIVISION_RHS" > "HOL4Real.transc.DIVISION_RHS"
+  "DIVISION_LT_GEN" > "HOL4Real.transc.DIVISION_LT_GEN"
+  "DIVISION_LT" > "HOL4Real.transc.DIVISION_LT"
+  "DIVISION_LHS" > "HOL4Real.transc.DIVISION_LHS"
+  "DIVISION_LE" > "HOL4Real.transc.DIVISION_LE"
+  "DIVISION_LBOUND_LT" > "HOL4Real.transc.DIVISION_LBOUND_LT"
+  "DIVISION_LBOUND" > "HOL4Real.transc.DIVISION_LBOUND"
+  "DIVISION_GT" > "HOL4Real.transc.DIVISION_GT"
+  "DIVISION_EXISTS" > "HOL4Real.transc.DIVISION_EXISTS"
+  "DIVISION_EQ" > "HOL4Real.transc.DIVISION_EQ"
+  "DIVISION_APPEND" > "HOL4Real.transc.DIVISION_APPEND"
+  "DIVISION_1" > "HOL4Real.transc.DIVISION_1"
+  "DIVISION_0" > "HOL4Real.transc.DIVISION_0"
+  "DINT_UNIQ" > "HOL4Real.transc.DINT_UNIQ"
+  "DIFF_TAN" > "HOL4Real.transc.DIFF_TAN"
+  "DIFF_SIN" > "HOL4Real.transc.DIFF_SIN"
+  "DIFF_LN_COMPOSITE" > "HOL4Real.transc.DIFF_LN_COMPOSITE"
+  "DIFF_LN" > "HOL4Real.transc.DIFF_LN"
+  "DIFF_EXP" > "HOL4Real.transc.DIFF_EXP"
+  "DIFF_COS" > "HOL4Real.transc.DIFF_COS"
+  "DIFF_COMPOSITE" > "HOL4Real.transc.DIFF_COMPOSITE"
+  "DIFF_ATN" > "HOL4Real.transc.DIFF_ATN"
+  "DIFF_ASN_LEMMA" > "HOL4Real.transc.DIFF_ASN_LEMMA"
+  "DIFF_ASN" > "HOL4Real.transc.DIFF_ASN"
+  "DIFF_ACS_LEMMA" > "HOL4Real.transc.DIFF_ACS_LEMMA"
+  "DIFF_ACS" > "HOL4Real.transc.DIFF_ACS"
+  "COS_ZERO_LEMMA" > "HOL4Real.transc.COS_ZERO_LEMMA"
+  "COS_ZERO" > "HOL4Real.transc.COS_ZERO"
+  "COS_TOTAL" > "HOL4Real.transc.COS_TOTAL"
+  "COS_SIN_SQRT" > "HOL4Real.transc.COS_SIN_SQRT"
+  "COS_SIN_SQ" > "HOL4Real.transc.COS_SIN_SQ"
+  "COS_SIN" > "HOL4Real.transc.COS_SIN"
+  "COS_POS_PI_LE" > "HOL4Real.transc.COS_POS_PI_LE"
+  "COS_POS_PI2_LE" > "HOL4Real.transc.COS_POS_PI2_LE"
+  "COS_POS_PI2" > "HOL4Real.transc.COS_POS_PI2"
+  "COS_POS_PI" > "HOL4Real.transc.COS_POS_PI"
+  "COS_PI2" > "HOL4Real.transc.COS_PI2"
+  "COS_PI" > "HOL4Real.transc.COS_PI"
+  "COS_PERIODIC_PI" > "HOL4Real.transc.COS_PERIODIC_PI"
+  "COS_PERIODIC" > "HOL4Real.transc.COS_PERIODIC"
+  "COS_PAIRED" > "HOL4Real.transc.COS_PAIRED"
+  "COS_NPI" > "HOL4Real.transc.COS_NPI"
+  "COS_NEG" > "HOL4Real.transc.COS_NEG"
+  "COS_ISZERO" > "HOL4Real.transc.COS_ISZERO"
+  "COS_FDIFF" > "HOL4Real.transc.COS_FDIFF"
+  "COS_DOUBLE" > "HOL4Real.transc.COS_DOUBLE"
+  "COS_CONVERGES" > "HOL4Real.transc.COS_CONVERGES"
+  "COS_BOUNDS" > "HOL4Real.transc.COS_BOUNDS"
+  "COS_BOUND" > "HOL4Real.transc.COS_BOUND"
+  "COS_ATN_NZ" > "HOL4Real.transc.COS_ATN_NZ"
+  "COS_ASN_NZ" > "HOL4Real.transc.COS_ASN_NZ"
+  "COS_ADD" > "HOL4Real.transc.COS_ADD"
+  "COS_ACS" > "HOL4Real.transc.COS_ACS"
+  "COS_2" > "HOL4Real.transc.COS_2"
+  "COS_0" > "HOL4Real.transc.COS_0"
+  "ATN_TAN" > "HOL4Real.transc.ATN_TAN"
+  "ATN_BOUNDS" > "HOL4Real.transc.ATN_BOUNDS"
+  "ATN" > "HOL4Real.transc.ATN"
+  "ASN_SIN" > "HOL4Real.transc.ASN_SIN"
+  "ASN_BOUNDS_LT" > "HOL4Real.transc.ASN_BOUNDS_LT"
+  "ASN_BOUNDS" > "HOL4Real.transc.ASN_BOUNDS"
+  "ASN" > "HOL4Real.transc.ASN"
+  "ACS_COS" > "HOL4Real.transc.ACS_COS"
+  "ACS_BOUNDS_LT" > "HOL4Real.transc.ACS_BOUNDS_LT"
+  "ACS_BOUNDS" > "HOL4Real.transc.ACS_BOUNDS"
+  "ACS" > "HOL4Real.transc.ACS"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/word32.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,377 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "word_suc" > "word_suc_def"
+  "word_sub" > "word_sub_def"
+  "word_ror1" > "word_ror1_def"
+  "word_ror" > "word_ror_def"
+  "word_mul" > "word_mul_def"
+  "word_lsr1" > "word_lsr1_def"
+  "word_lsr" > "word_lsr_def"
+  "word_lsl" > "word_lsl_def"
+  "word_asr1" > "word_asr1_def"
+  "word_asr" > "word_asr_def"
+  "word_add" > "word_add_def"
+  "word_2comp" > "word_2comp_def"
+  "word_1comp" > "word_1comp_def"
+  "w_T" > "w_T_primdef"
+  "w_1" > "w_1_primdef"
+  "w_0" > "w_0_primdef"
+  "w2n" > "w2n_primdef"
+  "n2w" > "n2w_primdef"
+  "mk_word32" > "mk_word32_def"
+  "dest_word32" > "dest_word32_def"
+  "bitwise_or" > "bitwise_or_def"
+  "bitwise_eor" > "bitwise_eor_def"
+  "bitwise_and" > "bitwise_and_def"
+  "WL" > "WL_primdef"
+  "TWO_COMP" > "TWO_COMP_primdef"
+  "TOw" > "TOw_primdef"
+  "SLICEw" > "SLICEw_primdef"
+  "RRXn" > "RRXn_primdef"
+  "RRX" > "RRX_primdef"
+  "ROR_ONE" > "ROR_ONE_primdef"
+  "OR" > "OR_primdef"
+  "ONE_COMP" > "ONE_COMP_primdef"
+  "MSBn" > "MSBn_primdef"
+  "MSB" > "MSB_primdef"
+  "MODw" > "MODw_primdef"
+  "LSR_ONE" > "LSR_ONE_primdef"
+  "LSB" > "LSB_primdef"
+  "INw" > "INw_primdef"
+  "HB" > "HB_primdef"
+  "EQUIV" > "EQUIV_primdef"
+  "EOR" > "EOR_primdef"
+  "COMP0" > "COMP0_primdef"
+  "BITw" > "BITw_primdef"
+  "BITSw" > "BITSw_primdef"
+  "ASR_ONE" > "ASR_ONE_primdef"
+  "AONE" > "AONE_primdef"
+  "AND" > "AND_primdef"
+
+type_maps
+  "word32" > "HOL4Word32.word32.word32"
+
+const_maps
+  "word_suc" > "HOL4Word32.word32.word_suc"
+  "word_sub" > "HOL4Word32.word32.word_sub"
+  "word_ror1" > "HOL4Word32.word32.word_ror1"
+  "word_ror" > "HOL4Word32.word32.word_ror"
+  "word_mul" > "HOL4Word32.word32.word_mul"
+  "word_lsr1" > "HOL4Word32.word32.word_lsr1"
+  "word_lsr" > "HOL4Word32.word32.word_lsr"
+  "word_lsl" > "HOL4Word32.word32.word_lsl"
+  "word_asr1" > "HOL4Word32.word32.word_asr1"
+  "word_asr" > "HOL4Word32.word32.word_asr"
+  "word_add" > "HOL4Word32.word32.word_add"
+  "word_2comp" > "HOL4Word32.word32.word_2comp"
+  "word_1comp" > "HOL4Word32.word32.word_1comp"
+  "w_T" > "HOL4Word32.word32.w_T"
+  "w_1" > "HOL4Word32.word32.w_1"
+  "w_0" > "HOL4Word32.word32.w_0"
+  "w2n" > "HOL4Word32.word32.w2n"
+  "n2w" > "HOL4Word32.word32.n2w"
+  "bitwise_or" > "HOL4Word32.word32.bitwise_or"
+  "bitwise_eor" > "HOL4Word32.word32.bitwise_eor"
+  "bitwise_and" > "HOL4Word32.word32.bitwise_and"
+  "WL" > "HOL4Word32.word32.WL"
+  "TWO_COMP" > "HOL4Word32.word32.TWO_COMP"
+  "TOw" > "HOL4Word32.word32.TOw"
+  "SLICEw" > "HOL4Word32.word32.SLICEw"
+  "RRXn" > "HOL4Word32.word32.RRXn"
+  "RRX" > "HOL4Word32.word32.RRX"
+  "ROR_ONE" > "HOL4Word32.word32.ROR_ONE"
+  "OR" > "HOL4Word32.word32.OR"
+  "ONE_COMP" > "HOL4Word32.word32.ONE_COMP"
+  "MSBn" > "HOL4Word32.word32.MSBn"
+  "MSB" > "HOL4Word32.word32.MSB"
+  "MODw" > "HOL4Word32.word32.MODw"
+  "LSR_ONE" > "HOL4Word32.word32.LSR_ONE"
+  "LSB" > "HOL4Word32.word32.LSB"
+  "INw" > "HOL4Word32.word32.INw"
+  "HB" > "HOL4Word32.word32.HB"
+  "EQUIV" > "HOL4Word32.word32.EQUIV"
+  "EOR" > "HOL4Word32.word32.EOR"
+  "COMP0" > "HOL4Word32.word32.COMP0"
+  "BITw" > "HOL4Word32.word32.BITw"
+  "BITSw" > "HOL4Word32.word32.BITSw"
+  "ASR_ONE" > "HOL4Word32.word32.ASR_ONE"
+  "AONE" > "HOL4Word32.word32.AONE"
+  "AND" > "HOL4Word32.word32.AND"
+
+const_renames
+  "==" > "EQUIV"
+
+thm_maps
+  "word_suc_def" > "HOL4Word32.word32.word_suc_def"
+  "word_suc" > "HOL4Word32.word32.word_suc"
+  "word_sub_def" > "HOL4Word32.word32.word_sub_def"
+  "word_sub" > "HOL4Word32.word32.word_sub"
+  "word_ror_def" > "HOL4Word32.word32.word_ror_def"
+  "word_ror1_def" > "HOL4Word32.word32.word_ror1_def"
+  "word_ror1" > "HOL4Word32.word32.word_ror1"
+  "word_ror" > "HOL4Word32.word32.word_ror"
+  "word_nchotomy" > "HOL4Word32.word32.word_nchotomy"
+  "word_mul_def" > "HOL4Word32.word32.word_mul_def"
+  "word_mul" > "HOL4Word32.word32.word_mul"
+  "word_lsr_def" > "HOL4Word32.word32.word_lsr_def"
+  "word_lsr1_def" > "HOL4Word32.word32.word_lsr1_def"
+  "word_lsr1" > "HOL4Word32.word32.word_lsr1"
+  "word_lsr" > "HOL4Word32.word32.word_lsr"
+  "word_lsl_def" > "HOL4Word32.word32.word_lsl_def"
+  "word_lsl" > "HOL4Word32.word32.word_lsl"
+  "word_asr_def" > "HOL4Word32.word32.word_asr_def"
+  "word_asr1_def" > "HOL4Word32.word32.word_asr1_def"
+  "word_asr1" > "HOL4Word32.word32.word_asr1"
+  "word_asr" > "HOL4Word32.word32.word_asr"
+  "word_add_def" > "HOL4Word32.word32.word_add_def"
+  "word_add" > "HOL4Word32.word32.word_add"
+  "word_2comp_def" > "HOL4Word32.word32.word_2comp_def"
+  "word_2comp" > "HOL4Word32.word32.word_2comp"
+  "word_1comp_def" > "HOL4Word32.word32.word_1comp_def"
+  "word_1comp" > "HOL4Word32.word32.word_1comp"
+  "word32_tybij" > "HOL4Word32.word32.word32_tybij"
+  "word32_TY_DEF" > "HOL4Word32.word32.word32_TY_DEF"
+  "w_T_primdef" > "HOL4Word32.word32.w_T_primdef"
+  "w_T_def" > "HOL4Word32.word32.w_T_def"
+  "w_T" > "HOL4Word32.word32.w_T"
+  "w_1_primdef" > "HOL4Word32.word32.w_1_primdef"
+  "w_1_def" > "HOL4Word32.word32.w_1_def"
+  "w_1" > "HOL4Word32.word32.w_1"
+  "w_0_primdef" > "HOL4Word32.word32.w_0_primdef"
+  "w_0_def" > "HOL4Word32.word32.w_0_def"
+  "w_0" > "HOL4Word32.word32.w_0"
+  "w2n_primdef" > "HOL4Word32.word32.w2n_primdef"
+  "w2n_def" > "HOL4Word32.word32.w2n_def"
+  "w2n_EVAL" > "HOL4Word32.word32.w2n_EVAL"
+  "w2n_ELIM" > "HOL4Word32.word32.w2n_ELIM"
+  "n2w_primdef" > "HOL4Word32.word32.n2w_primdef"
+  "n2w_def" > "HOL4Word32.word32.n2w_def"
+  "n2w_11" > "HOL4Word32.word32.n2w_11"
+  "dest_word_mk_word_eq3" > "HOL4Word32.word32.dest_word_mk_word_eq3"
+  "bitwise_or_def" > "HOL4Word32.word32.bitwise_or_def"
+  "bitwise_or" > "HOL4Word32.word32.bitwise_or"
+  "bitwise_eor_def" > "HOL4Word32.word32.bitwise_eor_def"
+  "bitwise_eor" > "HOL4Word32.word32.bitwise_eor"
+  "bitwise_and_def" > "HOL4Word32.word32.bitwise_and_def"
+  "bitwise_and" > "HOL4Word32.word32.bitwise_and"
+  "ZERO_SHIFT2" > "HOL4Word32.word32.ZERO_SHIFT2"
+  "ZERO_SHIFT" > "HOL4Word32.word32.ZERO_SHIFT"
+  "WL_primdef" > "HOL4Word32.word32.WL_primdef"
+  "WL_def" > "HOL4Word32.word32.WL_def"
+  "TWO_COMP_primdef" > "HOL4Word32.word32.TWO_COMP_primdef"
+  "TWO_COMP_def" > "HOL4Word32.word32.TWO_COMP_def"
+  "TWO_COMP_WELLDEF" > "HOL4Word32.word32.TWO_COMP_WELLDEF"
+  "TWO_COMP_ONE_COMP_QT" > "HOL4Word32.word32.TWO_COMP_ONE_COMP_QT"
+  "TWO_COMP_ONE_COMP" > "HOL4Word32.word32.TWO_COMP_ONE_COMP"
+  "TWO_COMP_EVAL2" > "HOL4Word32.word32.TWO_COMP_EVAL2"
+  "TWO_COMP_EVAL" > "HOL4Word32.word32.TWO_COMP_EVAL"
+  "TWO_COMP_ELIM" > "HOL4Word32.word32.TWO_COMP_ELIM"
+  "TWO_COMP_ADD" > "HOL4Word32.word32.TWO_COMP_ADD"
+  "TOw_primdef" > "HOL4Word32.word32.TOw_primdef"
+  "TOw_def" > "HOL4Word32.word32.TOw_def"
+  "TOw_WELLDEF" > "HOL4Word32.word32.TOw_WELLDEF"
+  "TOw_QT" > "HOL4Word32.word32.TOw_QT"
+  "TOw_IDEM" > "HOL4Word32.word32.TOw_IDEM"
+  "SUC_WELLDEF" > "HOL4Word32.word32.SUC_WELLDEF"
+  "SUC_EQUIV_COMP" > "HOL4Word32.word32.SUC_EQUIV_COMP"
+  "SUBw" > "HOL4Word32.word32.SUBw"
+  "SUB_SUBw" > "HOL4Word32.word32.SUB_SUBw"
+  "SUB_PLUSw" > "HOL4Word32.word32.SUB_PLUSw"
+  "SUB_EQUALw" > "HOL4Word32.word32.SUB_EQUALw"
+  "SLICEw_primdef" > "HOL4Word32.word32.SLICEw_primdef"
+  "SLICEw_def" > "HOL4Word32.word32.SLICEw_def"
+  "SLICEw_ZERO_THM" > "HOL4Word32.word32.SLICEw_ZERO_THM"
+  "SLICEw_ZERO" > "HOL4Word32.word32.SLICEw_ZERO"
+  "SLICEw_THM" > "HOL4Word32.word32.SLICEw_THM"
+  "SLICEw_COMP_THM" > "HOL4Word32.word32.SLICEw_COMP_THM"
+  "SLICE_EVAL" > "HOL4Word32.word32.SLICE_EVAL"
+  "RRXn_primdef" > "HOL4Word32.word32.RRXn_primdef"
+  "RRXn_def" > "HOL4Word32.word32.RRXn_def"
+  "RRX_primdef" > "HOL4Word32.word32.RRX_primdef"
+  "RRX_def" > "HOL4Word32.word32.RRX_def"
+  "RRX_WELLDEF" > "HOL4Word32.word32.RRX_WELLDEF"
+  "RRX_EVAL2" > "HOL4Word32.word32.RRX_EVAL2"
+  "RRX_EVAL" > "HOL4Word32.word32.RRX_EVAL"
+  "ROR_w_T" > "HOL4Word32.word32.ROR_w_T"
+  "ROR_THM" > "HOL4Word32.word32.ROR_THM"
+  "ROR_ONE_primdef" > "HOL4Word32.word32.ROR_ONE_primdef"
+  "ROR_ONE_def" > "HOL4Word32.word32.ROR_ONE_def"
+  "ROR_ONE_WELLDEF" > "HOL4Word32.word32.ROR_ONE_WELLDEF"
+  "ROR_ONE_EVAL2" > "HOL4Word32.word32.ROR_ONE_EVAL2"
+  "ROR_ONE_EVAL" > "HOL4Word32.word32.ROR_ONE_EVAL"
+  "ROR_CYCLE" > "HOL4Word32.word32.ROR_CYCLE"
+  "ROR_ADD" > "HOL4Word32.word32.ROR_ADD"
+  "RIGHT_OR_OVER_ANDw" > "HOL4Word32.word32.RIGHT_OR_OVER_ANDw"
+  "RIGHT_OR_OVER_AND_QT" > "HOL4Word32.word32.RIGHT_OR_OVER_AND_QT"
+  "RIGHT_AND_OVER_ORw" > "HOL4Word32.word32.RIGHT_AND_OVER_ORw"
+  "RIGHT_AND_OVER_OR_QT" > "HOL4Word32.word32.RIGHT_AND_OVER_OR_QT"
+  "OR_primdef" > "HOL4Word32.word32.OR_primdef"
+  "OR_def" > "HOL4Word32.word32.OR_def"
+  "OR_IDEMw" > "HOL4Word32.word32.OR_IDEMw"
+  "OR_IDEM_QT" > "HOL4Word32.word32.OR_IDEM_QT"
+  "OR_EVAL2" > "HOL4Word32.word32.OR_EVAL2"
+  "OR_EVAL" > "HOL4Word32.word32.OR_EVAL"
+  "OR_COMP_QT" > "HOL4Word32.word32.OR_COMP_QT"
+  "OR_COMMw" > "HOL4Word32.word32.OR_COMMw"
+  "OR_COMM_QT" > "HOL4Word32.word32.OR_COMM_QT"
+  "OR_ASSOCw" > "HOL4Word32.word32.OR_ASSOCw"
+  "OR_ASSOC_QT" > "HOL4Word32.word32.OR_ASSOC_QT"
+  "OR_ABSORBw" > "HOL4Word32.word32.OR_ABSORBw"
+  "OR_ABSORB_QT" > "HOL4Word32.word32.OR_ABSORB_QT"
+  "ONE_COMPw" > "HOL4Word32.word32.ONE_COMPw"
+  "ONE_COMP_primdef" > "HOL4Word32.word32.ONE_COMP_primdef"
+  "ONE_COMP_def" > "HOL4Word32.word32.ONE_COMP_def"
+  "ONE_COMP_WELLDEF" > "HOL4Word32.word32.ONE_COMP_WELLDEF"
+  "ONE_COMP_TWO_COMP" > "HOL4Word32.word32.ONE_COMP_TWO_COMP"
+  "ONE_COMP_THM" > "HOL4Word32.word32.ONE_COMP_THM"
+  "ONE_COMP_QT" > "HOL4Word32.word32.ONE_COMP_QT"
+  "ONE_COMP_EVAL2" > "HOL4Word32.word32.ONE_COMP_EVAL2"
+  "ONE_COMP_EVAL" > "HOL4Word32.word32.ONE_COMP_EVAL"
+  "MUL_WELLDEF" > "HOL4Word32.word32.MUL_WELLDEF"
+  "MUL_EVAL2" > "HOL4Word32.word32.MUL_EVAL2"
+  "MUL_EVAL" > "HOL4Word32.word32.MUL_EVAL"
+  "MULT_QT" > "HOL4Word32.word32.MULT_QT"
+  "MULT_COMMw" > "HOL4Word32.word32.MULT_COMMw"
+  "MULT_COMM_QT" > "HOL4Word32.word32.MULT_COMM_QT"
+  "MULT_CLAUSESw" > "HOL4Word32.word32.MULT_CLAUSESw"
+  "MULT_CLAUSES_QT" > "HOL4Word32.word32.MULT_CLAUSES_QT"
+  "MULT_ASSOCw" > "HOL4Word32.word32.MULT_ASSOCw"
+  "MULT_ASSOC_QT" > "HOL4Word32.word32.MULT_ASSOC_QT"
+  "MSBn_primdef" > "HOL4Word32.word32.MSBn_primdef"
+  "MSBn_def" > "HOL4Word32.word32.MSBn_def"
+  "MSB_primdef" > "HOL4Word32.word32.MSB_primdef"
+  "MSB_def" > "HOL4Word32.word32.MSB_def"
+  "MSB_WELLDEF" > "HOL4Word32.word32.MSB_WELLDEF"
+  "MSB_EVAL2" > "HOL4Word32.word32.MSB_EVAL2"
+  "MSB_EVAL" > "HOL4Word32.word32.MSB_EVAL"
+  "MODw_primdef" > "HOL4Word32.word32.MODw_primdef"
+  "MODw_def" > "HOL4Word32.word32.MODw_def"
+  "MODw_THM" > "HOL4Word32.word32.MODw_THM"
+  "MODw_MULT" > "HOL4Word32.word32.MODw_MULT"
+  "MODw_IDEM2" > "HOL4Word32.word32.MODw_IDEM2"
+  "MODw_EVAL" > "HOL4Word32.word32.MODw_EVAL"
+  "MODw_ELIM" > "HOL4Word32.word32.MODw_ELIM"
+  "MOD_MOD_DIV_2EXP" > "HOL4Word32.word32.MOD_MOD_DIV_2EXP"
+  "MOD_MOD_DIV" > "HOL4Word32.word32.MOD_MOD_DIV"
+  "MOD_ADD" > "HOL4Word32.word32.MOD_ADD"
+  "LSR_THM" > "HOL4Word32.word32.LSR_THM"
+  "LSR_ONE_primdef" > "HOL4Word32.word32.LSR_ONE_primdef"
+  "LSR_ONE_def" > "HOL4Word32.word32.LSR_ONE_def"
+  "LSR_ONE_WELLDEF" > "HOL4Word32.word32.LSR_ONE_WELLDEF"
+  "LSR_ONE_EVAL2" > "HOL4Word32.word32.LSR_ONE_EVAL2"
+  "LSR_ONE_EVAL" > "HOL4Word32.word32.LSR_ONE_EVAL"
+  "LSR_ONE" > "HOL4Word32.word32.LSR_ONE"
+  "LSR_LIMIT" > "HOL4Word32.word32.LSR_LIMIT"
+  "LSR_EVAL" > "HOL4Word32.word32.LSR_EVAL"
+  "LSR_ADD" > "HOL4Word32.word32.LSR_ADD"
+  "LSL_LIMIT" > "HOL4Word32.word32.LSL_LIMIT"
+  "LSL_ADD" > "HOL4Word32.word32.LSL_ADD"
+  "LSB_primdef" > "HOL4Word32.word32.LSB_primdef"
+  "LSB_def" > "HOL4Word32.word32.LSB_def"
+  "LSB_WELLDEF" > "HOL4Word32.word32.LSB_WELLDEF"
+  "LSB_EVAL2" > "HOL4Word32.word32.LSB_EVAL2"
+  "LSB_EVAL" > "HOL4Word32.word32.LSB_EVAL"
+  "LEFT_SHIFT_LESS" > "HOL4Word32.word32.LEFT_SHIFT_LESS"
+  "LEFT_ADD_DISTRIBw" > "HOL4Word32.word32.LEFT_ADD_DISTRIBw"
+  "LEFT_ADD_DISTRIB_QT" > "HOL4Word32.word32.LEFT_ADD_DISTRIB_QT"
+  "INw_primdef" > "HOL4Word32.word32.INw_primdef"
+  "INw_def" > "HOL4Word32.word32.INw_def"
+  "INw_MODw" > "HOL4Word32.word32.INw_MODw"
+  "INV_SUC_EQ_QT" > "HOL4Word32.word32.INV_SUC_EQ_QT"
+  "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_COMP" > "HOL4Word32.word32.FUNPOW_COMP"
+  "EQ_ADD_RCANCELw" > "HOL4Word32.word32.EQ_ADD_RCANCELw"
+  "EQ_ADD_RCANCEL_QT" > "HOL4Word32.word32.EQ_ADD_RCANCEL_QT"
+  "EQ_ADD_LCANCELw" > "HOL4Word32.word32.EQ_ADD_LCANCELw"
+  "EQ_ADD_LCANCEL_QT" > "HOL4Word32.word32.EQ_ADD_LCANCEL_QT"
+  "EQUIV_primdef" > "HOL4Word32.word32.EQUIV_primdef"
+  "EQUIV_def" > "HOL4Word32.word32.EQUIV_def"
+  "EQUIV_QT" > "HOL4Word32.word32.EQUIV_QT"
+  "EOR_primdef" > "HOL4Word32.word32.EOR_primdef"
+  "EOR_def" > "HOL4Word32.word32.EOR_def"
+  "EOR_EVAL2" > "HOL4Word32.word32.EOR_EVAL2"
+  "EOR_EVAL" > "HOL4Word32.word32.EOR_EVAL"
+  "DE_MORGAN_THMw" > "HOL4Word32.word32.DE_MORGAN_THMw"
+  "DE_MORGAN_THM_QT" > "HOL4Word32.word32.DE_MORGAN_THM_QT"
+  "COMP0_primdef" > "HOL4Word32.word32.COMP0_primdef"
+  "COMP0_def" > "HOL4Word32.word32.COMP0_def"
+  "CANCEL_SUBw" > "HOL4Word32.word32.CANCEL_SUBw"
+  "BITw_primdef" > "HOL4Word32.word32.BITw_primdef"
+  "BITw_def" > "HOL4Word32.word32.BITw_def"
+  "BITw_THM" > "HOL4Word32.word32.BITw_THM"
+  "BIT_EVAL" > "HOL4Word32.word32.BIT_EVAL"
+  "BIT_EQUIV_THM" > "HOL4Word32.word32.BIT_EQUIV_THM"
+  "BIT_EQUIV" > "HOL4Word32.word32.BIT_EQUIV"
+  "BITWISEw_WELLDEF" > "HOL4Word32.word32.BITWISEw_WELLDEF"
+  "BITWISE_WELLDEF" > "HOL4Word32.word32.BITWISE_WELLDEF"
+  "BITWISE_THM2" > "HOL4Word32.word32.BITWISE_THM2"
+  "BITWISE_ONE_COMP_THM" > "HOL4Word32.word32.BITWISE_ONE_COMP_THM"
+  "BITWISE_ISTEP" > "HOL4Word32.word32.BITWISE_ISTEP"
+  "BITWISE_EVAL2" > "HOL4Word32.word32.BITWISE_EVAL2"
+  "BITWISE_EVAL" > "HOL4Word32.word32.BITWISE_EVAL"
+  "BITSw_primdef" > "HOL4Word32.word32.BITSw_primdef"
+  "BITSw_def" > "HOL4Word32.word32.BITSw_def"
+  "BITSw_ZERO" > "HOL4Word32.word32.BITSw_ZERO"
+  "BITSw_DIV_THM" > "HOL4Word32.word32.BITSw_DIV_THM"
+  "BITSw_COMP_THM" > "HOL4Word32.word32.BITSw_COMP_THM"
+  "BITSwLT_THM" > "HOL4Word32.word32.BITSwLT_THM"
+  "BITS_SUC2" > "HOL4Word32.word32.BITS_SUC2"
+  "BITS_SLICEw_THM" > "HOL4Word32.word32.BITS_SLICEw_THM"
+  "BITS_EVAL" > "HOL4Word32.word32.BITS_EVAL"
+  "ASR_w_T" > "HOL4Word32.word32.ASR_w_T"
+  "ASR_THM" > "HOL4Word32.word32.ASR_THM"
+  "ASR_ONE_primdef" > "HOL4Word32.word32.ASR_ONE_primdef"
+  "ASR_ONE_def" > "HOL4Word32.word32.ASR_ONE_def"
+  "ASR_ONE_WELLDEF" > "HOL4Word32.word32.ASR_ONE_WELLDEF"
+  "ASR_ONE_EVAL2" > "HOL4Word32.word32.ASR_ONE_EVAL2"
+  "ASR_ONE_EVAL" > "HOL4Word32.word32.ASR_ONE_EVAL"
+  "ASR_LIMIT" > "HOL4Word32.word32.ASR_LIMIT"
+  "ASR_ADD" > "HOL4Word32.word32.ASR_ADD"
+  "AONE_primdef" > "HOL4Word32.word32.AONE_primdef"
+  "AONE_def" > "HOL4Word32.word32.AONE_def"
+  "AND_primdef" > "HOL4Word32.word32.AND_primdef"
+  "AND_def" > "HOL4Word32.word32.AND_def"
+  "AND_IDEMw" > "HOL4Word32.word32.AND_IDEMw"
+  "AND_IDEM_QT" > "HOL4Word32.word32.AND_IDEM_QT"
+  "AND_EVAL2" > "HOL4Word32.word32.AND_EVAL2"
+  "AND_EVAL" > "HOL4Word32.word32.AND_EVAL"
+  "AND_COMP_QT" > "HOL4Word32.word32.AND_COMP_QT"
+  "AND_COMMw" > "HOL4Word32.word32.AND_COMMw"
+  "AND_COMM_QT" > "HOL4Word32.word32.AND_COMM_QT"
+  "AND_ASSOCw" > "HOL4Word32.word32.AND_ASSOCw"
+  "AND_ASSOC_QT" > "HOL4Word32.word32.AND_ASSOC_QT"
+  "AND_ABSORBw" > "HOL4Word32.word32.AND_ABSORBw"
+  "AND_ABSORB_QT" > "HOL4Word32.word32.AND_ABSORB_QT"
+  "ADDw" > "HOL4Word32.word32.ADDw"
+  "ADD_WELLDEF" > "HOL4Word32.word32.ADD_WELLDEF"
+  "ADD_TWO_COMP_QT" > "HOL4Word32.word32.ADD_TWO_COMP_QT"
+  "ADD_TWO_COMP2" > "HOL4Word32.word32.ADD_TWO_COMP2"
+  "ADD_TWO_COMP" > "HOL4Word32.word32.ADD_TWO_COMP"
+  "ADD_SUBw" > "HOL4Word32.word32.ADD_SUBw"
+  "ADD_SUB_SYM" > "HOL4Word32.word32.ADD_SUB_SYM"
+  "ADD_SUB_ASSOC" > "HOL4Word32.word32.ADD_SUB_ASSOC"
+  "ADD_QT" > "HOL4Word32.word32.ADD_QT"
+  "ADD_INV_0_QT" > "HOL4Word32.word32.ADD_INV_0_QT"
+  "ADD_INV_0_EQw" > "HOL4Word32.word32.ADD_INV_0_EQw"
+  "ADD_INV_0_EQ_QT" > "HOL4Word32.word32.ADD_INV_0_EQ_QT"
+  "ADD_EVAL2" > "HOL4Word32.word32.ADD_EVAL2"
+  "ADD_EVAL" > "HOL4Word32.word32.ADD_EVAL"
+  "ADD_EQ_SUBw" > "HOL4Word32.word32.ADD_EQ_SUBw"
+  "ADD_COMMw" > "HOL4Word32.word32.ADD_COMMw"
+  "ADD_COMM_QT" > "HOL4Word32.word32.ADD_COMM_QT"
+  "ADD_CLAUSESw" > "HOL4Word32.word32.ADD_CLAUSESw"
+  "ADD_CLAUSES_QT" > "HOL4Word32.word32.ADD_CLAUSES_QT"
+  "ADD_ASSOCw" > "HOL4Word32.word32.ADD_ASSOCw"
+  "ADD_ASSOC_QT" > "HOL4Word32.word32.ADD_ASSOC_QT"
+  "ADD_0w" > "HOL4Word32.word32.ADD_0w"
+  "ADD_0_QT" > "HOL4Word32.word32.ADD_0_QT"
+  "ADD1w" > "HOL4Word32.word32.ADD1w"
+  "ADD1_QT" > "HOL4Word32.word32.ADD1_QT"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/word_base.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,102 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "word_size" > "word_size_primdef"
+  "word_case" > "word_case_primdef"
+  "word_base0" > "word_base0_primdef"
+  "mk_word" > "mk_word_def"
+  "dest_word" > "dest_word_def"
+  "bit" > "bit_def"
+  "WSPLIT" > "WSPLIT_def"
+  "WSEG" > "WSEG_def"
+  "WORDLEN" > "WORDLEN_def"
+  "WORD" > "WORD_def"
+  "WCAT" > "WCAT_def"
+  "PWORDLEN" > "PWORDLEN_primdef"
+  "MSB" > "MSB_def"
+  "LSB" > "LSB_def"
+
+type_maps
+  "word" > "HOL4Vec.word_base.word"
+
+const_maps
+  "word_base0" > "HOL4Vec.word_base.word_base0"
+  "WORD" > "HOL4Vec.word_base.WORD"
+  "PWORDLEN" > "HOL4Vec.word_base.PWORDLEN"
+
+const_renames
+  "BIT" > "bit"
+
+thm_maps
+  "word_size_def" > "HOL4Vec.word_base.word_size_def"
+  "word_repfns" > "HOL4Vec.word_base.word_repfns"
+  "word_nchotomy" > "HOL4Vec.word_base.word_nchotomy"
+  "word_induction" > "HOL4Vec.word_base.word_induction"
+  "word_induct" > "HOL4Vec.word_base.word_induct"
+  "word_cases" > "HOL4Vec.word_base.word_cases"
+  "word_case_def" > "HOL4Vec.word_base.word_case_def"
+  "word_case_cong" > "HOL4Vec.word_base.word_case_cong"
+  "word_base0_primdef" > "HOL4Vec.word_base.word_base0_primdef"
+  "word_base0_def" > "HOL4Vec.word_base.word_base0_def"
+  "word_TY_DEF" > "HOL4Vec.word_base.word_TY_DEF"
+  "word_Axiom" > "HOL4Vec.word_base.word_Axiom"
+  "word_Ax" > "HOL4Vec.word_base.word_Ax"
+  "word_11" > "HOL4Vec.word_base.word_11"
+  "WSPLIT_WSEG2" > "HOL4Vec.word_base.WSPLIT_WSEG2"
+  "WSPLIT_WSEG1" > "HOL4Vec.word_base.WSPLIT_WSEG1"
+  "WSPLIT_WSEG" > "HOL4Vec.word_base.WSPLIT_WSEG"
+  "WSPLIT_PWORDLEN" > "HOL4Vec.word_base.WSPLIT_PWORDLEN"
+  "WSPLIT_DEF" > "HOL4Vec.word_base.WSPLIT_DEF"
+  "WSEG_WSEG" > "HOL4Vec.word_base.WSEG_WSEG"
+  "WSEG_WORD_LENGTH" > "HOL4Vec.word_base.WSEG_WORD_LENGTH"
+  "WSEG_WORDLEN" > "HOL4Vec.word_base.WSEG_WORDLEN"
+  "WSEG_WCAT_WSEG2" > "HOL4Vec.word_base.WSEG_WCAT_WSEG2"
+  "WSEG_WCAT_WSEG1" > "HOL4Vec.word_base.WSEG_WCAT_WSEG1"
+  "WSEG_WCAT_WSEG" > "HOL4Vec.word_base.WSEG_WCAT_WSEG"
+  "WSEG_WCAT2" > "HOL4Vec.word_base.WSEG_WCAT2"
+  "WSEG_WCAT1" > "HOL4Vec.word_base.WSEG_WCAT1"
+  "WSEG_SUC" > "HOL4Vec.word_base.WSEG_SUC"
+  "WSEG_PWORDLEN" > "HOL4Vec.word_base.WSEG_PWORDLEN"
+  "WSEG_DEF" > "HOL4Vec.word_base.WSEG_DEF"
+  "WSEG_BIT" > "HOL4Vec.word_base.WSEG_BIT"
+  "WSEG0" > "HOL4Vec.word_base.WSEG0"
+  "WORD_def" > "HOL4Vec.word_base.WORD_def"
+  "WORD_SPLIT" > "HOL4Vec.word_base.WORD_SPLIT"
+  "WORD_SNOC_WCAT" > "HOL4Vec.word_base.WORD_SNOC_WCAT"
+  "WORD_PARTITION" > "HOL4Vec.word_base.WORD_PARTITION"
+  "WORD_CONS_WCAT" > "HOL4Vec.word_base.WORD_CONS_WCAT"
+  "WORD_11" > "HOL4Vec.word_base.WORD_11"
+  "WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT" > "HOL4Vec.word_base.WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT"
+  "WORDLEN_SUC_WCAT_WSEG_WSEG" > "HOL4Vec.word_base.WORDLEN_SUC_WCAT_WSEG_WSEG"
+  "WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT" > "HOL4Vec.word_base.WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT"
+  "WORDLEN_SUC_WCAT_BIT_WSEG" > "HOL4Vec.word_base.WORDLEN_SUC_WCAT_BIT_WSEG"
+  "WORDLEN_SUC_WCAT" > "HOL4Vec.word_base.WORDLEN_SUC_WCAT"
+  "WORDLEN_DEF" > "HOL4Vec.word_base.WORDLEN_DEF"
+  "WORD" > "HOL4Vec.word_base.WORD"
+  "WCAT_WSEG_WSEG" > "HOL4Vec.word_base.WCAT_WSEG_WSEG"
+  "WCAT_PWORDLEN" > "HOL4Vec.word_base.WCAT_PWORDLEN"
+  "WCAT_DEF" > "HOL4Vec.word_base.WCAT_DEF"
+  "WCAT_ASSOC" > "HOL4Vec.word_base.WCAT_ASSOC"
+  "WCAT_11" > "HOL4Vec.word_base.WCAT_11"
+  "WCAT0" > "HOL4Vec.word_base.WCAT0"
+  "PWORDLEN_primdef" > "HOL4Vec.word_base.PWORDLEN_primdef"
+  "PWORDLEN_def" > "HOL4Vec.word_base.PWORDLEN_def"
+  "PWORDLEN1" > "HOL4Vec.word_base.PWORDLEN1"
+  "PWORDLEN0" > "HOL4Vec.word_base.PWORDLEN0"
+  "PWORDLEN" > "HOL4Vec.word_base.PWORDLEN"
+  "MSB_DEF" > "HOL4Vec.word_base.MSB_DEF"
+  "MSB" > "HOL4Vec.word_base.MSB"
+  "LSB_DEF" > "HOL4Vec.word_base.LSB_DEF"
+  "LSB" > "HOL4Vec.word_base.LSB"
+  "IN_PWORDLEN" > "HOL4Vec.word_base.IN_PWORDLEN"
+  "BIT_WSEG" > "HOL4Vec.word_base.BIT_WSEG"
+  "BIT_WCAT_SND" > "HOL4Vec.word_base.BIT_WCAT_SND"
+  "BIT_WCAT_FST" > "HOL4Vec.word_base.BIT_WCAT_FST"
+  "BIT_WCAT1" > "HOL4Vec.word_base.BIT_WCAT1"
+  "BIT_EQ_IMP_WORD_EQ" > "HOL4Vec.word_base.BIT_EQ_IMP_WORD_EQ"
+  "BIT_DEF" > "HOL4Vec.word_base.BIT_DEF"
+  "BIT0" > "HOL4Vec.word_base.BIT0"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/word_bitop.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,64 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "WMAP" > "WMAP_def"
+  "SHR" > "SHR_def"
+  "SHL" > "SHL_def"
+  "PBITOP" > "PBITOP_primdef"
+  "PBITBOP" > "PBITBOP_primdef"
+  "FORALLBITS" > "FORALLBITS_def"
+  "EXISTSABIT" > "EXISTSABIT_def"
+
+const_maps
+  "SHR" > "HOL4Vec.word_bitop.SHR"
+  "SHL" > "HOL4Vec.word_bitop.SHL"
+  "PBITOP" > "HOL4Vec.word_bitop.PBITOP"
+  "PBITBOP" > "HOL4Vec.word_bitop.PBITBOP"
+
+thm_maps
+  "WSEG_SHL_0" > "HOL4Vec.word_bitop.WSEG_SHL_0"
+  "WSEG_SHL" > "HOL4Vec.word_bitop.WSEG_SHL"
+  "WMAP_o" > "HOL4Vec.word_bitop.WMAP_o"
+  "WMAP_WSEG" > "HOL4Vec.word_bitop.WMAP_WSEG"
+  "WMAP_WCAT" > "HOL4Vec.word_bitop.WMAP_WCAT"
+  "WMAP_PWORDLEN" > "HOL4Vec.word_bitop.WMAP_PWORDLEN"
+  "WMAP_PBITOP" > "HOL4Vec.word_bitop.WMAP_PBITOP"
+  "WMAP_DEF" > "HOL4Vec.word_bitop.WMAP_DEF"
+  "WMAP_BIT" > "HOL4Vec.word_bitop.WMAP_BIT"
+  "WMAP_0" > "HOL4Vec.word_bitop.WMAP_0"
+  "SHR_def" > "HOL4Vec.word_bitop.SHR_def"
+  "SHR_WSEG_NF" > "HOL4Vec.word_bitop.SHR_WSEG_NF"
+  "SHR_WSEG_1F" > "HOL4Vec.word_bitop.SHR_WSEG_1F"
+  "SHR_WSEG" > "HOL4Vec.word_bitop.SHR_WSEG"
+  "SHR_DEF" > "HOL4Vec.word_bitop.SHR_DEF"
+  "SHL_def" > "HOL4Vec.word_bitop.SHL_def"
+  "SHL_WSEG_NF" > "HOL4Vec.word_bitop.SHL_WSEG_NF"
+  "SHL_WSEG_1F" > "HOL4Vec.word_bitop.SHL_WSEG_1F"
+  "SHL_WSEG" > "HOL4Vec.word_bitop.SHL_WSEG"
+  "SHL_DEF" > "HOL4Vec.word_bitop.SHL_DEF"
+  "PBITOP_primdef" > "HOL4Vec.word_bitop.PBITOP_primdef"
+  "PBITOP_def" > "HOL4Vec.word_bitop.PBITOP_def"
+  "PBITOP_WSEG" > "HOL4Vec.word_bitop.PBITOP_WSEG"
+  "PBITOP_PWORDLEN" > "HOL4Vec.word_bitop.PBITOP_PWORDLEN"
+  "PBITOP_BIT" > "HOL4Vec.word_bitop.PBITOP_BIT"
+  "PBITBOP_primdef" > "HOL4Vec.word_bitop.PBITBOP_primdef"
+  "PBITBOP_def" > "HOL4Vec.word_bitop.PBITBOP_def"
+  "PBITBOP_WSEG" > "HOL4Vec.word_bitop.PBITBOP_WSEG"
+  "PBITBOP_PWORDLEN" > "HOL4Vec.word_bitop.PBITBOP_PWORDLEN"
+  "PBITBOP_EXISTS" > "HOL4Vec.word_bitop.PBITBOP_EXISTS"
+  "NOT_FORALLBITS" > "HOL4Vec.word_bitop.NOT_FORALLBITS"
+  "NOT_EXISTSABIT" > "HOL4Vec.word_bitop.NOT_EXISTSABIT"
+  "IN_PBITOP" > "HOL4Vec.word_bitop.IN_PBITOP"
+  "IN_PBITBOP" > "HOL4Vec.word_bitop.IN_PBITBOP"
+  "FORALLBITS_WSEG" > "HOL4Vec.word_bitop.FORALLBITS_WSEG"
+  "FORALLBITS_WCAT" > "HOL4Vec.word_bitop.FORALLBITS_WCAT"
+  "FORALLBITS_DEF" > "HOL4Vec.word_bitop.FORALLBITS_DEF"
+  "FORALLBITS" > "HOL4Vec.word_bitop.FORALLBITS"
+  "EXISTSABIT_WSEG" > "HOL4Vec.word_bitop.EXISTSABIT_WSEG"
+  "EXISTSABIT_WCAT" > "HOL4Vec.word_bitop.EXISTSABIT_WCAT"
+  "EXISTSABIT_DEF" > "HOL4Vec.word_bitop.EXISTSABIT_DEF"
+  "EXISTSABIT" > "HOL4Vec.word_bitop.EXISTSABIT"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL/word_num.imp	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,35 @@
+import
+
+import_segment "hol4"
+
+def_maps
+  "NWORD" > "NWORD_def"
+  "NVAL" > "NVAL_def"
+  "NLIST" > "NLIST_def"
+  "LVAL" > "LVAL_def"
+
+const_maps
+  "NWORD" > "HOL4Vec.word_num.NWORD"
+  "LVAL" > "HOL4Vec.word_num.LVAL"
+
+thm_maps
+  "NWORD_def" > "HOL4Vec.word_num.NWORD_def"
+  "NWORD_PWORDLEN" > "HOL4Vec.word_num.NWORD_PWORDLEN"
+  "NWORD_LENGTH" > "HOL4Vec.word_num.NWORD_LENGTH"
+  "NWORD_DEF" > "HOL4Vec.word_num.NWORD_DEF"
+  "NVAL_WORDLEN_0" > "HOL4Vec.word_num.NVAL_WORDLEN_0"
+  "NVAL_WCAT2" > "HOL4Vec.word_num.NVAL_WCAT2"
+  "NVAL_WCAT1" > "HOL4Vec.word_num.NVAL_WCAT1"
+  "NVAL_WCAT" > "HOL4Vec.word_num.NVAL_WCAT"
+  "NVAL_MAX" > "HOL4Vec.word_num.NVAL_MAX"
+  "NVAL_DEF" > "HOL4Vec.word_num.NVAL_DEF"
+  "NVAL1" > "HOL4Vec.word_num.NVAL1"
+  "NVAL0" > "HOL4Vec.word_num.NVAL0"
+  "NLIST_DEF" > "HOL4Vec.word_num.NLIST_DEF"
+  "LVAL_def" > "HOL4Vec.word_num.LVAL_def"
+  "LVAL_SNOC" > "HOL4Vec.word_num.LVAL_SNOC"
+  "LVAL_MAX" > "HOL4Vec.word_num.LVAL_MAX"
+  "LVAL_DEF" > "HOL4Vec.word_num.LVAL_DEF"
+  "LVAL" > "HOL4Vec.word_num.LVAL"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4Compat.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,533 @@
+theory HOL4Compat = HOL4Setup + Divides + Primes + Real:
+
+lemma EXISTS_UNIQUE_DEF: "(Ex1 P) = (Ex P & (ALL x y. P x & P y --> (x = y)))"
+  by auto
+
+lemma COND_DEF:"(If b t f) = (@x. ((b = True) --> (x = t)) & ((b = False) --> (x = f)))"
+  by auto
+
+constdefs
+  LET :: "['a \<Rightarrow> 'b,'a] \<Rightarrow> 'b"
+  "LET f s == f s"
+
+lemma [hol4rew]: "LET f s = Let s f"
+  by (simp add: LET_def Let_def)
+
+lemmas [hol4rew] = ONE_ONE_rew
+
+lemma bool_case_DEF: "(bool_case x y b) = (if b then x else y)"
+  by simp;
+
+lemma INR_INL_11: "(ALL y x. (Inl x = Inl y) = (x = y)) & (ALL y x. (Inr x = Inr y) = (x = y))"
+  by safe
+
+consts
+  ISL :: "'a + 'b => bool"
+  ISR :: "'a + 'b => bool"
+
+primrec ISL_def:
+  "ISL (Inl x) = True"
+  "ISL (Inr x) = False"
+
+primrec ISR_def:
+  "ISR (Inl x) = False"
+  "ISR (Inr x) = True"
+
+lemma ISL: "(ALL x. ISL (Inl x)) & (ALL y. ~ISL (Inr y))"
+  by simp
+
+lemma ISR: "(ALL x. ISR (Inr x)) & (ALL y. ~ISR (Inl y))"
+  by simp
+
+consts
+  OUTL :: "'a + 'b => 'a"
+  OUTR :: "'a + 'b => 'b"
+
+primrec OUTL_def:
+  "OUTL (Inl x) = x"
+
+primrec OUTR_def:
+  "OUTR (Inr x) = x"
+
+lemma OUTL: "OUTL (Inl x) = x"
+  by simp
+
+lemma OUTR: "OUTR (Inr x) = x"
+  by simp
+
+lemma sum_case_def: "(ALL f g x. sum_case f g (Inl x) = f x) & (ALL f g y. sum_case f g (Inr y) = g y)"
+  by simp;
+
+lemma one: "ALL v. v = ()"
+  by simp;
+
+lemma option_case_def: "(!u f. option_case u f None = u) & (!u f x. option_case u f (Some x) = f x)"
+  by simp
+
+lemma OPTION_MAP_DEF: "(!f x. option_map f (Some x) = Some (f x)) & (!f. option_map f None = None)"
+  by simp
+
+consts
+  IS_SOME :: "'a option => bool"
+  IS_NONE :: "'a option => bool"
+
+primrec IS_SOME_def:
+  "IS_SOME (Some x) = True"
+  "IS_SOME None = False"
+
+primrec IS_NONE_def:
+  "IS_NONE (Some x) = False"
+  "IS_NONE None = True"
+
+lemma IS_NONE_DEF: "(!x. IS_NONE (Some x) = False) & (IS_NONE None = True)"
+  by simp
+
+lemma IS_SOME_DEF: "(!x. IS_SOME (Some x) = True) & (IS_SOME None = False)"
+  by simp
+
+consts
+  OPTION_JOIN :: "'a option option => 'a option"
+
+primrec OPTION_JOIN_def:
+  "OPTION_JOIN None = None"
+  "OPTION_JOIN (Some x) = x"
+
+lemma OPTION_JOIN_DEF: "(OPTION_JOIN None = None) & (ALL x. OPTION_JOIN (Some x) = x)"
+  by simp;
+
+lemma PAIR: "(fst x,snd x) = x"
+  by simp
+
+lemma PAIR_MAP: "prod_fun f g p = (f (fst p),g (snd p))"
+  by (simp add: prod_fun_def split_def)
+
+lemma pair_case_def: "split = split"
+  ..;
+
+lemma LESS_OR_EQ: "m <= (n::nat) = (m < n | m = n)"
+  by auto
+
+constdefs
+  nat_gt :: "nat => nat => bool"
+  "nat_gt == %m n. n < m"
+  nat_ge :: "nat => nat => bool"
+  "nat_ge == %m n. nat_gt m n | m = n"
+
+lemma [hol4rew]: "nat_gt m n = (n < m)"
+  by (simp add: nat_gt_def)
+
+lemma [hol4rew]: "nat_ge m n = (n <= m)"
+  by (auto simp add: nat_ge_def nat_gt_def)
+
+lemma GREATER_DEF: "ALL m n. (n < m) = (n < m)"
+  by simp
+
+lemma GREATER_OR_EQ: "ALL m n. n <= (m::nat) = (n < m | m = n)"
+  by auto
+
+lemma LESS_DEF: "m < n = (? P. (!n. P (Suc n) --> P n) & P m & ~P n)"
+proof safe
+  assume "m < n"
+  def P == "%n. n <= m"
+  have "(!n. P (Suc n) \<longrightarrow> P n) & P m & ~P n"
+  proof (auto simp add: P_def)
+    assume "n <= m"
+    from prems
+    show False
+      by auto
+  qed
+  thus "? P. (!n. P (Suc n) \<longrightarrow> P n) & P m & ~P n"
+    by auto
+next
+  fix P
+  assume alln: "!n. P (Suc n) \<longrightarrow> P n"
+  assume pm: "P m"
+  assume npn: "~P n"
+  have "!k q. q + k = m \<longrightarrow> P q"
+  proof
+    fix k
+    show "!q. q + k = m \<longrightarrow> P q"
+    proof (induct k,simp_all)
+      show "P m" .
+    next
+      fix k
+      assume ind: "!q. q + k = m \<longrightarrow> P q"
+      show "!q. Suc (q + k) = m \<longrightarrow> P q"
+      proof (rule+)
+	fix q
+	assume "Suc (q + k) = m"
+	hence "(Suc q) + k = m"
+	  by simp
+	with ind
+	have psq: "P (Suc q)"
+	  by simp
+	from alln
+	have "P (Suc q) --> P q"
+	  ..
+	with psq
+	show "P q"
+	  by simp
+      qed
+    qed
+  qed
+  hence "!q. q + (m - n) = m \<longrightarrow> P q"
+    ..
+  hence hehe: "n + (m - n) = m \<longrightarrow> P n"
+    ..
+  show "m < n"
+  proof (rule classical)
+    assume "~(m<n)"
+    hence "n <= m"
+      by simp
+    with hehe
+    have "P n"
+      by simp
+    with npn
+    show "m < n"
+      ..
+  qed
+qed;
+
+constdefs
+  FUNPOW :: "('a => 'a) => nat => 'a => 'a"
+  "FUNPOW f n == f ^ n"
+
+lemma FUNPOW: "(ALL f x. (f ^ 0) x = x) &
+  (ALL f n x. (f ^ Suc n) x = (f ^ n) (f x))"
+proof auto
+  fix f n x
+  have "ALL x. f ((f ^ n) x) = (f ^ n) (f x)"
+    by (induct n,auto)
+  thus "f ((f ^ n) x) = (f ^ n) (f x)"
+    ..
+qed
+
+lemma [hol4rew]: "FUNPOW f n = f ^ n"
+  by (simp add: FUNPOW_def)
+
+lemma ADD: "(!n. (0::nat) + n = n) & (!m n. Suc m + n = Suc (m + n))"
+  by simp
+
+lemma MULT: "(!n. (0::nat) * n = 0) & (!m n. Suc m * n = m * n + n)"
+  by simp
+
+lemma SUB: "(!m. (0::nat) - m = 0) & (!m n. (Suc m) - n = (if m < n then 0 else Suc (m - n)))"
+  apply simp
+  apply arith
+  done
+
+lemma MAX_DEF: "max (m::nat) n = (if m < n then n else m)"
+  by (simp add: max_def)
+
+lemma MIN_DEF: "min (m::nat) n = (if m < n then m else n)"
+  by (simp add: min_def)
+
+lemma DIVISION: "(0::nat) < n --> (!k. (k = k div n * n + k mod n) & k mod n < n)"
+  by simp
+
+constdefs
+  ALT_ZERO :: nat
+  "ALT_ZERO == 0"
+  NUMERAL_BIT1 :: "nat \<Rightarrow> nat"
+  "NUMERAL_BIT1 n == n + (n + Suc 0)"
+  NUMERAL_BIT2 :: "nat \<Rightarrow> nat"
+  "NUMERAL_BIT2 n == n + (n + Suc (Suc 0))"
+  NUMERAL :: "nat \<Rightarrow> nat"
+  "NUMERAL x == x"
+
+lemma [hol4rew]: "NUMERAL ALT_ZERO = 0"
+  by (simp add: ALT_ZERO_def NUMERAL_def)
+
+lemma [hol4rew]: "NUMERAL (NUMERAL_BIT1 ALT_ZERO) = 1"
+  by (simp add: ALT_ZERO_def NUMERAL_BIT1_def NUMERAL_def)
+
+lemma [hol4rew]: "NUMERAL (NUMERAL_BIT2 ALT_ZERO) = 2"
+  by (simp add: ALT_ZERO_def NUMERAL_BIT2_def NUMERAL_def)
+
+lemma EXP: "(!m. m ^ 0 = (1::nat)) & (!m n. m ^ Suc n = m * (m::nat) ^ n)"
+  by auto
+
+lemma num_case_def: "(!b f. nat_case b f 0 = b) & (!b f n. nat_case b f (Suc n) = f n)"
+  by simp;
+
+lemma divides_def: "(a::nat) dvd b = (? q. b = q * a)"
+  by (auto simp add: dvd_def);
+
+lemma list_case_def: "(!v f. list_case v f [] = v) & (!v f a0 a1. list_case v f (a0#a1) = f a0 a1)"
+  by simp
+
+consts
+  list_size :: "('a \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> nat"
+
+primrec
+  "list_size f [] = 0"
+  "list_size f (a0#a1) = 1 + (f a0 + list_size f a1)"
+
+lemma list_size_def: "(!f. list_size f [] = 0) &
+         (!f a0 a1. list_size f (a0#a1) = 1 + (f a0 + list_size f a1))"
+  by simp
+
+lemma list_case_cong: "! M M' v f. M = M' & (M' = [] \<longrightarrow>  v = v') &
+           (!a0 a1. (M' = a0#a1) \<longrightarrow> (f a0 a1 = f' a0 a1)) -->
+           (list_case v f M = list_case v' f' M')"
+proof clarify
+  fix M M' v f
+  assume "M' = [] \<longrightarrow> v = v'"
+    and "!a0 a1. M' = a0 # a1 \<longrightarrow> f a0 a1 = f' a0 a1"
+  show "list_case v f M' = list_case v' f' M'"
+  proof (rule List.list.case_cong)
+    show "M' = M'"
+      ..
+  next
+    assume "M' = []"
+    with prems
+    show "v = v'"
+      by auto
+  next
+    fix a0 a1
+    assume "M' = a0 # a1"
+    with prems
+    show "f a0 a1 = f' a0 a1"
+      by auto
+  qed
+qed
+
+lemma list_Axiom: "ALL f0 f1. EX fn. (fn [] = f0) & (ALL a0 a1. fn (a0#a1) = f1 a0 a1 (fn a1))"
+proof safe
+  fix f0 f1
+  def fn == "list_rec f0 f1"
+  have "fn [] = f0 & (ALL a0 a1. fn (a0 # a1) = f1 a0 a1 (fn a1))"
+    by (simp add: fn_def)
+  thus "EX fn. fn [] = f0 & (ALL a0 a1. fn (a0 # a1) = f1 a0 a1 (fn a1))"
+    by auto
+qed
+
+lemma list_Axiom_old: "EX! fn. (fn [] = x) & (ALL h t. fn (h#t) = f (fn t) h t)"
+proof safe
+  def fn == "list_rec x (%h t r. f r h t)"
+  have "fn [] = x & (ALL h t. fn (h # t) = f (fn t) h t)"
+    by (simp add: fn_def)
+  thus "EX fn. fn [] = x & (ALL h t. fn (h # t) = f (fn t) h t)"
+    by auto
+next
+  fix fn1 fn2
+  assume "ALL h t. fn1 (h # t) = f (fn1 t) h t"
+  assume "ALL h t. fn2 (h # t) = f (fn2 t) h t"
+  assume "fn2 [] = fn1 []"
+  show "fn1 = fn2"
+  proof
+    fix xs
+    show "fn1 xs = fn2 xs"
+      by (induct xs,simp_all add: prems) 
+  qed
+qed
+
+lemma NULL_DEF: "(null [] = True) & (!h t. null (h # t) = False)"
+  by simp
+
+constdefs
+  sum :: "nat list \<Rightarrow> nat"
+  "sum l == foldr (op +) l 0"
+
+lemma SUM: "(sum [] = 0) & (!h t. sum (h#t) = h + sum t)"
+  by (simp add: sum_def)
+
+lemma APPEND: "(!l. [] @ l = l) & (!l1 l2 h. (h#l1) @ l2 = h# l1 @ l2)"
+  by simp
+
+lemma FLAT: "(concat [] = []) & (!h t. concat (h#t) = h @ (concat t))"
+  by simp
+
+lemma LENGTH: "(length [] = 0) & (!h t. length (h#t) = Suc (length t))"
+  by simp
+
+lemma MAP: "(!f. map f [] = []) & (!f h t. map f (h#t) = f h#map f t)"
+  by simp
+
+lemma MEM: "(!x. x mem [] = False) & (!x h t. x mem (h#t) = ((x = h) | x mem t))"
+  by auto
+
+lemma FILTER: "(!P. filter P [] = []) & (!P h t.
+           filter P (h#t) = (if P h then h#filter P t else filter P t))"
+  by simp
+
+lemma REPLICATE: "(ALL x. replicate 0 x = []) &
+  (ALL n x. replicate (Suc n) x = x # replicate n x)"
+  by simp
+
+constdefs
+  FOLDR :: "[['a,'b]\<Rightarrow>'b,'b,'a list] \<Rightarrow> 'b"
+  "FOLDR f e l == foldr f l e"
+
+lemma [hol4rew]: "FOLDR f e l = foldr f l e"
+  by (simp add: FOLDR_def)
+
+lemma FOLDR: "(!f e. foldr f [] e = e) & (!f e x l. foldr f (x#l) e = f x (foldr f l e))"
+  by simp
+
+lemma FOLDL: "(!f e. foldl f e [] = e) & (!f e x l. foldl f e (x#l) = foldl f (f e x) l)"
+  by simp
+
+lemma EVERY_DEF: "(!P. list_all P [] = True) & (!P h t. list_all P (h#t) = (P h & list_all P t))"
+  by simp
+
+consts
+  list_exists :: "['a \<Rightarrow> bool,'a list] \<Rightarrow> bool"
+
+primrec
+  list_exists_Nil: "list_exists P Nil = False"
+  list_exists_Cons: "list_exists P (x#xs) = (if P x then True else list_exists P xs)"
+
+lemma list_exists_DEF: "(!P. list_exists P [] = False) &
+         (!P h t. list_exists P (h#t) = (P h | list_exists P t))"
+  by simp
+
+consts
+  map2 :: "[['a,'b]\<Rightarrow>'c,'a list,'b list] \<Rightarrow> 'c list"
+
+primrec
+  map2_Nil: "map2 f [] l2 = []"
+  map2_Cons: "map2 f (x#xs) l2 = f x (hd l2) # map2 f xs (tl l2)"
+
+lemma MAP2: "(!f. map2 f [] [] = []) & (!f h1 t1 h2 t2. map2 f (h1#t1) (h2#t2) = f h1 h2#map2 f t1 t2)"
+  by simp
+
+lemma list_INDUCT: "\<lbrakk> P [] ; !t. P t \<longrightarrow> (!h. P (h#t)) \<rbrakk> \<Longrightarrow> !l. P l"
+proof
+  fix l
+  assume "P []"
+  assume allt: "!t. P t \<longrightarrow> (!h. P (h # t))"
+  show "P l"
+  proof (induct l)
+    show "P []" .
+  next
+    fix h t
+    assume "P t"
+    with allt
+    have "!h. P (h # t)"
+      by auto
+    thus "P (h # t)"
+      ..
+  qed
+qed
+
+lemma list_CASES: "(l = []) | (? t h. l = h#t)"
+  by (induct l,auto)
+
+constdefs
+  ZIP :: "'a list * 'b list \<Rightarrow> ('a * 'b) list"
+  "ZIP == %(a,b). zip a b"
+
+lemma ZIP: "(zip [] [] = []) &
+  (!x1 l1 x2 l2. zip (x1#l1) (x2#l2) = (x1,x2)#zip l1 l2)"
+  by simp
+
+lemma [hol4rew]: "ZIP (a,b) = zip a b"
+  by (simp add: ZIP_def)
+
+consts
+  unzip :: "('a * 'b) list \<Rightarrow> 'a list * 'b list"
+
+primrec
+  unzip_Nil: "unzip [] = ([],[])"
+  unzip_Cons: "unzip (xy#xys) = (let zs = unzip xys in (fst xy # fst zs,snd xy # snd zs))"
+
+lemma UNZIP: "(unzip [] = ([],[])) &
+         (!x l. unzip (x#l) = (fst x#fst (unzip l),snd x#snd (unzip l)))"
+  by (simp add: Let_def)
+
+lemma REVERSE: "(rev [] = []) & (!h t. rev (h#t) = (rev t) @ [h])"
+  by simp;
+
+lemma REAL_SUP_ALLPOS: "\<lbrakk> ALL x. P (x::real) \<longrightarrow> 0 < x ; EX x. P x; EX z. ALL x. P x \<longrightarrow> x < z \<rbrakk> \<Longrightarrow> EX s. ALL y. (EX x. P x & y < x) = (y < s)"
+proof safe
+  fix x z
+  assume allx: "ALL x. P x \<longrightarrow> 0 < x"
+  assume px: "P x"
+  assume allx': "ALL x. P x \<longrightarrow> x < z"
+  have "EX s. ALL y. (EX x : Collect P. y < x) = (y < s)"
+  proof (rule posreal_complete)
+    show "ALL x : Collect P. 0 < x"
+    proof safe
+      fix x
+      assume "P x"
+      from allx
+      have "P x \<longrightarrow> 0 < x"
+	..
+      thus "0 < x"
+	by (simp add: prems)
+    qed
+  next
+    from px
+    show "EX x. x : Collect P"
+      by auto
+  next
+    from allx'
+    show "EX y. ALL x : Collect P. x < y"
+      apply simp
+      ..
+  qed
+  thus "EX s. ALL y. (EX x. P x & y < x) = (y < s)"
+    by simp
+qed
+
+lemma REAL_10: "~((1::real) = 0)"
+  by simp
+
+lemma REAL_ADD_ASSOC: "(x::real) + (y + z) = x + y + z"
+  by simp
+
+lemma REAL_MUL_ASSOC: "(x::real) * (y * z) = x * y * z"
+  by simp
+
+lemma REAL_ADD_LINV:  "-x + x = (0::real)"
+  by simp
+
+lemma REAL_MUL_LINV: "x ~= (0::real) ==> inverse x * x = 1"
+  by simp
+
+lemma REAL_LT_TOTAL: "((x::real) = y) | x < y | y < x"
+  by auto;
+
+lemma [hol4rew]: "real (0::nat) = 0"
+  by simp
+
+lemma [hol4rew]: "real (1::nat) = 1"
+  by simp
+
+lemma [hol4rew]: "real (2::nat) = 2"
+  by simp
+
+lemma real_lte: "((x::real) <= y) = (~(y < x))"
+  by auto
+
+lemma real_of_num: "((0::real) = 0) & (!n. real (Suc n) = real n + 1)"
+  by (simp add: real_of_nat_Suc)
+
+lemma abs: "abs (x::real) = (if 0 <= x then x else -x)"
+  by (simp add: real_abs_def)
+
+lemma pow: "(!x::real. x ^ 0 = 1) & (!x::real. ALL n. x ^ (Suc n) = x * x ^ n)"
+  by simp;
+
+constdefs
+  real_gt :: "real => real => bool" 
+  "real_gt == %x y. y < x"
+
+lemma [hol4rew]: "real_gt x y = (y < x)"
+  by (simp add: real_gt_def)
+
+lemma real_gt: "ALL x (y::real). (y < x) = (y < x)"
+  by simp
+
+constdefs
+  real_ge :: "real => real => bool"
+  "real_ge x y == y <= x"
+
+lemma [hol4rew]: "real_ge x y = (y <= x)"
+  by (simp add: real_ge_def)
+
+lemma real_ge: "ALL x y. (y <= x) = (y <= x)"
+  by simp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4Setup.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,166 @@
+theory HOL4Setup = MakeEqual
+  files ("proof_kernel.ML") ("replay.ML") ("hol4rews.ML") ("import_package.ML"):
+
+section {* General Setup *}
+
+lemma eq_imp: "P = Q \<Longrightarrow> P \<longrightarrow> Q"
+  by auto
+
+lemma HOLallI: "(!! bogus. P bogus) \<Longrightarrow> (ALL bogus. P bogus)"
+proof -
+  assume "!! bogus. P bogus"
+  thus "ALL x. P x"
+    ..
+qed
+
+consts
+  ONE_ONE :: "('a => 'b) => bool"
+  ONTO    :: "('a => 'b) => bool"
+
+defs
+  ONE_ONE_DEF: "ONE_ONE f == ALL x y. f x = f y --> x = y"
+  ONTO_DEF   : "ONTO f == ALL y. EX x. y = f x"
+
+lemma ONE_ONE_rew: "ONE_ONE f = inj_on f UNIV"
+  by (simp add: ONE_ONE_DEF inj_on_def)
+
+lemma INFINITY_AX: "EX (f::ind \<Rightarrow> ind). (inj f & ~(ONTO f))"
+proof (rule exI,safe)
+  show "inj Suc_Rep"
+    by (rule inj_Suc_Rep)
+next
+  assume "ONTO Suc_Rep"
+  hence "ALL y. EX x. y = Suc_Rep x"
+    by (simp add: ONTO_DEF surj_def)
+  hence "EX x. Zero_Rep = Suc_Rep x"
+    by (rule spec)
+  thus False
+  proof (rule exE)
+    fix x
+    assume "Zero_Rep = Suc_Rep x"
+    hence "Suc_Rep x = Zero_Rep"
+      ..
+    with Suc_Rep_not_Zero_Rep
+    show False
+      ..
+  qed
+qed
+
+lemma EXISTS_DEF: "Ex P = P (Eps P)"
+proof (rule iffI)
+  assume "Ex P"
+  thus "P (Eps P)"
+    ..
+next
+  assume "P (Eps P)"
+  thus "Ex P"
+    ..
+qed
+
+consts
+  TYPE_DEFINITION :: "('a => bool) => ('b => 'a) => bool"
+
+defs
+  TYPE_DEFINITION: "TYPE_DEFINITION p rep == ((ALL x y. (rep x = rep y) --> (x = y)) & (ALL x. (p x = (EX y. x = rep y))))"
+
+lemma ex_imp_nonempty: "Ex P ==> EX x. x : (Collect P)"
+  by simp
+
+lemma light_ex_imp_nonempty: "P t ==> EX x. x : (Collect P)"
+proof -
+  assume "P t"
+  hence "EX x. P x"
+    ..
+  thus ?thesis
+    by (rule ex_imp_nonempty)
+qed
+
+lemma light_imp_as: "[| Q --> P; P --> Q |] ==> P = Q"
+  by blast
+
+lemma typedef_hol2hol4:
+  assumes a: "type_definition (Rep::'a=>'b) Abs (Collect P)"
+  shows "EX rep. TYPE_DEFINITION P (rep::'a=>'b)"
+proof -
+  from a
+  have td: "(ALL x. P (Rep x)) & (ALL x. Abs (Rep x) = x) & (ALL y. P y \<longrightarrow> Rep (Abs y) = y)"
+    by (simp add: type_definition_def)
+  have ed: "TYPE_DEFINITION P Rep"
+  proof (auto simp add: TYPE_DEFINITION)
+    fix x y
+    assume "Rep x = Rep y"
+    from td have "x = Abs (Rep x)"
+      by auto
+    also have "Abs (Rep x) = Abs (Rep y)"
+      by (simp add: prems)
+    also from td have "Abs (Rep y) = y"
+      by auto
+    finally show "x = y" .
+  next
+    fix x
+    assume "P x"
+    with td
+    have "Rep (Abs x) = x"
+      by auto
+    hence "x = Rep (Abs x)"
+      ..
+    thus "EX y. x = Rep y"
+      ..
+  next
+    fix y
+    from td
+    show "P (Rep y)"
+      by auto
+  qed
+  show ?thesis
+    apply (rule exI [of _ Rep])
+    apply (rule ed)
+    .
+qed
+
+lemma typedef_hol2hollight:
+  assumes a: "type_definition (Rep::'a=>'b) Abs (Collect P)"
+  shows "(Abs (Rep a) = a) & (P r = (Rep (Abs r) = r))"
+proof
+  from a
+  show "Abs (Rep a) = a"
+    by (rule type_definition.Rep_inverse)
+next
+  show "P r = (Rep (Abs r) = r)"
+  proof
+    assume "P r"
+    hence "r \<in> (Collect P)"
+      by simp
+    with a
+    show "Rep (Abs r) = r"
+      by (rule type_definition.Abs_inverse)
+  next
+    assume ra: "Rep (Abs r) = r"
+    from a
+    have "Rep (Abs r) \<in> (Collect P)"
+      by (rule type_definition.Rep)
+    thus "P r"
+      by (simp add: ra)
+  qed
+qed
+
+lemma termspec_help: "[| Ex P ; c == Eps P |] ==> P c"
+  apply simp
+  apply (rule someI_ex)
+  .
+
+lemma typedef_helper: "EX x. P x \<Longrightarrow> EX x. x \<in> (Collect P)"
+  by simp
+
+use "hol4rews.ML"
+
+setup hol4_setup
+parse_ast_translation smarter_trueprop_parsing
+
+use "proof_kernel.ML"
+use "replay.ML"
+use "import_package.ML"
+
+setup ImportPackage.setup
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4Syntax.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,6 @@
+theory HOL4Syntax = HOL4Setup
+  files "import_syntax.ML":
+
+ML {* HOL4ImportSyntax.setup() *}
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/MakeEqual.thy	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,68 @@
+theory MakeEqual = Main
+  files "shuffler.ML":
+
+setup Shuffler.setup
+
+lemma conj_norm [shuffle_rule]: "(A & B ==> PROP C) == ([| A ; B |] ==> PROP C)"
+proof
+  assume "A & B ==> PROP C" A B
+  thus "PROP C"
+    by auto
+next
+  assume "[| A; B |] ==> PROP C" "A & B"
+  thus "PROP C"
+    by auto
+qed
+
+lemma imp_norm [shuffle_rule]: "(Trueprop (A --> B)) == (A ==> B)"
+proof
+  assume "A --> B" A
+  thus B ..
+next
+  assume "A ==> B"
+  thus "A --> B"
+    by auto
+qed
+
+lemma all_norm [shuffle_rule]: "(Trueprop (ALL x. P x)) == (!!x. P x)"
+proof
+  fix x
+  assume "ALL x. P x"
+  thus "P x" ..
+next
+  assume "!!x. P x"
+  thus "ALL x. P x"
+    ..
+qed
+
+lemma ex_norm [shuffle_rule]: "(EX x. P x ==> PROP Q) == (!!x. P x ==> PROP Q)"
+proof
+  fix x
+  assume ex: "EX x. P x ==> PROP Q"
+  assume "P x"
+  hence "EX x. P x" ..
+  with ex show "PROP Q" .
+next
+  assume allx: "!!x. P x ==> PROP Q"
+  assume "EX x. P x"
+  hence p: "P (SOME x. P x)"
+    ..
+  from allx
+  have "P (SOME x. P x) ==> PROP Q"
+    .
+  with p
+  show "PROP Q"
+    by auto
+qed
+
+lemma eq_norm [shuffle_rule]: "Trueprop (t = u) == (t == u)"
+proof
+  assume "t = u"
+  thus "t == u" by simp
+next
+  assume "t == u"
+  thus "t = u"
+    by simp
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/ROOT.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,2 @@
+use_thy "HOL4Compat";
+use_thy "HOL4Syntax";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/hol4rews.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,773 @@
+structure StringPair = TableFun(type key = string * string val ord = prod_ord string_ord string_ord);
+
+type holthm = (term * term) list * thm
+
+datatype ImportStatus =
+	 NoImport
+       | Generating of string
+       | Replaying of string
+
+structure HOL4DefThyArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/import_status"
+type T = ImportStatus
+val empty = NoImport
+val copy = I
+val prep_ext = I
+fun merge (NoImport,NoImport) = NoImport
+  | merge _ = (warning "Import status set during merge"; NoImport)
+fun print sg import_status =
+    Pretty.writeln (Pretty.str (case import_status of NoImport => "No current import" | Generating thyname => ("Generating " ^ thyname) | Replaying thyname => ("Replaying " ^ thyname)))
+end;
+
+structure HOL4DefThy = TheoryDataFun(HOL4DefThyArgs);
+
+structure ImportSegmentArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/import"
+type T = string
+val empty = ""
+val copy = I
+val prep_ext = I
+fun merge ("",arg) = arg
+  | merge (arg,"") = arg
+  | merge (s1,s2) =
+    if s1 = s2
+    then s1
+    else error "Trying to merge two different import segments"
+fun print sg import_segment =
+    Pretty.writeln (Pretty.str ("Import segment: " ^ import_segment))
+end;
+
+structure ImportSegment = TheoryDataFun(ImportSegmentArgs);
+
+val get_import_segment = ImportSegment.get
+val set_import_segment = ImportSegment.put
+
+structure HOL4UNamesArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/used_names"
+type T = string list
+val empty = []
+val copy = I
+val prep_ext = I
+fun merge ([],[]) = []
+  | merge _ = error "Used names not empty during merge"
+fun print sg used_names =
+    Pretty.writeln (Pretty.str "Printing of HOL4/used_names Not implemented")
+end;
+
+structure HOL4UNames = TheoryDataFun(HOL4UNamesArgs);
+
+structure HOL4DumpArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/dump_data"
+type T = string * string * string list
+val empty = ("","",[])
+val copy = I
+val prep_ext = I
+fun merge (("","",[]),("","",[])) = ("","",[])
+  | merge _ = error "Dump data not empty during merge"
+fun print sg dump_data =
+    Pretty.writeln (Pretty.str "Printing of HOL4/dump_data Not implemented")
+end;
+
+structure HOL4Dump = TheoryDataFun(HOL4DumpArgs);
+
+structure HOL4MovesArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/moves"
+type T = string Symtab.table
+val empty = Symtab.empty
+val copy = I
+val prep_ext = I
+val merge = Symtab.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 moves:"
+	(Symtab.foldl (fn (xl,(bef,aft)) => (Pretty.str (bef ^ " --> " ^ aft)::xl)) ([],tab)))
+end;
+
+structure HOL4Moves = TheoryDataFun(HOL4MovesArgs);
+
+structure HOL4ImportsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/imports"
+type T = string Symtab.table
+val empty = Symtab.empty
+val copy = I
+val prep_ext = I
+val merge = Symtab.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 imports:"
+	(Symtab.foldl (fn (xl,(thyname,segname)) => (Pretty.str (thyname ^ " imported from segment " ^ segname)::xl)) ([],tab)))
+end;
+
+structure HOL4Imports = TheoryDataFun(HOL4ImportsArgs);
+
+fun get_segment2 thyname thy =
+    let
+	val imps = HOL4Imports.get thy
+    in
+	Symtab.lookup(imps,thyname)
+    end
+
+fun set_segment thyname segname thy =
+    let
+	val imps = HOL4Imports.get thy
+	val imps' = Symtab.update_new((thyname,segname),imps)
+    in
+	HOL4Imports.put imps' thy
+    end
+
+structure HOL4CMovesArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/constant_moves"
+type T = string Symtab.table
+val empty = Symtab.empty
+val copy = I
+val prep_ext = I
+val merge = Symtab.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 constant moves:"
+	(Symtab.foldl (fn (xl,(bef,aft)) => (Pretty.str (bef ^ " --> " ^ aft)::xl)) ([],tab)))
+end;
+
+structure HOL4CMoves = TheoryDataFun(HOL4CMovesArgs);
+
+structure HOL4MapsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/mappings"
+type T = (string option) StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 mappings:"
+	(StringPair.foldl (fn (xl,((bthy,bthm),isathm)) => (Pretty.str (bthy ^ "." ^ bthm ^ (case isathm of Some th => " --> " ^ th | None => "IGNORED"))::xl)) ([],tab)))
+end;
+
+structure HOL4Maps = TheoryDataFun(HOL4MapsArgs);
+
+structure HOL4ThmsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/theorems"
+type T = holthm StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 mappings:"
+	(StringPair.foldl (fn (xl,((bthy,bthm),(_,thm))) => (Pretty.str (bthy ^ "." ^ bthm ^ ":")::(Display.pretty_thm thm)::xl)) ([],tab)))
+end;
+
+structure HOL4Thms = TheoryDataFun(HOL4ThmsArgs);
+
+structure HOL4ConstMapsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/constmappings"
+type T = (bool * string * typ option) StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 constant mappings:"
+	(StringPair.foldl (fn (xl,((bthy,bconst),(internal,isaconst,_))) => (Pretty.str (bthy ^ "." ^ bconst ^ " --> " ^ isaconst ^ (if internal then " (*)" else ""))::xl)) ([],tab)))
+end;
+
+structure HOL4ConstMaps = TheoryDataFun(HOL4ConstMapsArgs);
+
+structure HOL4RenameArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/renamings"
+type T = string StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 constant renamings:"
+	(StringPair.foldl (fn (xl,((bthy,bconst),newname)) => (Pretty.str (bthy ^ "." ^ bconst ^ " --> " ^ newname)::xl)) ([],tab)))
+end;
+
+structure HOL4Rename = TheoryDataFun(HOL4RenameArgs);
+
+structure HOL4DefMapsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/def_maps"
+type T = string StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 constant definitions:"
+	(StringPair.foldl (fn (xl,((bthy,bconst),newname)) => (Pretty.str (bthy ^ "." ^ bconst ^ ": " ^ newname)::xl)) ([],tab)))
+end;
+
+structure HOL4DefMaps = TheoryDataFun(HOL4DefMapsArgs);
+
+structure HOL4TypeMapsArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/typemappings"
+type T = (bool * string) StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 type mappings:"
+	(StringPair.foldl (fn (xl,((bthy,bconst),(internal,isaconst))) => (Pretty.str (bthy ^ "." ^ bconst ^ " --> " ^ isaconst ^ (if internal then " (*)" else ""))::xl)) ([],tab)))
+end;
+
+structure HOL4TypeMaps = TheoryDataFun(HOL4TypeMapsArgs);
+
+structure HOL4PendingArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/pending"
+type T = ((term * term) list * thm) StringPair.table
+val empty = StringPair.empty
+val copy = I
+val prep_ext = I
+val merge = StringPair.merge (K true)
+fun print sg tab =
+    Pretty.writeln (Pretty.big_list "HOL4 pending theorems:"
+	(StringPair.foldl (fn (xl,((bthy,bthm),(_,th))) => (Pretty.chunks [Pretty.str (bthy ^ "." ^ bthm ^ ":"),Display.pretty_thm th]::xl)) ([],tab)))
+end;
+
+structure HOL4Pending = TheoryDataFun(HOL4PendingArgs);
+
+structure HOL4RewritesArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL4/rewrites"
+type T = thm list
+val empty = []
+val copy = I
+val prep_ext = I
+val merge = Library.gen_union Thm.eq_thm
+fun print sg thms =
+    Pretty.writeln (Pretty.big_list "HOL4 rewrite rules:"
+				    (map Display.pretty_thm thms))
+end;
+
+structure HOL4Rewrites = TheoryDataFun(HOL4RewritesArgs);
+
+val hol4_debug = ref false
+fun message s = if !hol4_debug then writeln s else ()
+
+fun add_hol4_rewrite (thy,th) =
+    let
+	val _ = message "Adding HOL4 rewrite"
+	val th1 = th RS eq_reflection
+	val current_rews = HOL4Rewrites.get thy
+	val new_rews = gen_ins Thm.eq_thm (th1,current_rews)
+	val updated_thy  = HOL4Rewrites.put new_rews thy
+    in
+	(updated_thy,th1)
+    end;
+
+fun ignore_hol4 bthy bthm thy =
+    let
+	val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
+	val curmaps = HOL4Maps.get thy
+	val newmaps = StringPair.update_new(((bthy,bthm),None),curmaps)
+	val upd_thy = HOL4Maps.put newmaps thy
+    in
+	upd_thy
+    end
+
+val opt_get_output_thy = #2 o HOL4Dump.get
+
+fun get_output_thy thy =
+    case #2 (HOL4Dump.get thy) of
+	"" => error "No theory file being output"
+      | thyname => thyname
+
+val get_output_dir = #1 o HOL4Dump.get
+
+fun add_hol4_move bef aft thy =
+    let
+	val curmoves = HOL4Moves.get thy
+	val newmoves = Symtab.update_new((bef,aft),curmoves)
+    in
+	HOL4Moves.put newmoves thy
+    end
+
+fun get_hol4_move bef thy =
+    let
+	val moves = HOL4Moves.get thy
+    in
+	Symtab.lookup(moves,bef)
+    end
+
+fun follow_name thmname thy =
+    let
+	val moves = HOL4Moves.get thy
+	fun F thmname =
+	    case Symtab.lookup(moves,thmname) of
+		Some name => F name
+	      | None => thmname
+    in
+	F thmname
+    end
+
+fun add_hol4_cmove bef aft thy =
+    let
+	val curmoves = HOL4CMoves.get thy
+	val newmoves = Symtab.update_new((bef,aft),curmoves)
+    in
+	HOL4CMoves.put newmoves thy
+    end
+
+fun get_hol4_cmove bef thy =
+    let
+	val moves = HOL4CMoves.get thy
+    in
+	Symtab.lookup(moves,bef)
+    end
+
+fun follow_cname thmname thy =
+    let
+	val moves = HOL4CMoves.get thy
+	fun F thmname =
+	    case Symtab.lookup(moves,thmname) of
+		Some name => F name
+	      | None => thmname
+    in
+	F thmname
+    end
+
+fun add_hol4_mapping bthy bthm isathm thy =
+    let
+	val isathm = follow_name isathm thy
+	val _ = message ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)
+	val curmaps = HOL4Maps.get thy
+	val newmaps = StringPair.update_new(((bthy,bthm),Some isathm),curmaps)
+	val upd_thy = HOL4Maps.put newmaps thy
+    in
+	upd_thy
+    end;
+
+fun get_hol4_type_mapping bthy tycon thy =
+    let
+	val maps = HOL4TypeMaps.get thy
+    in
+	StringPair.lookup(maps,(bthy,tycon))
+    end
+
+fun get_hol4_mapping bthy bthm thy =
+    let
+	val curmaps = HOL4Maps.get thy
+    in
+	StringPair.lookup(curmaps,(bthy,bthm))
+    end;
+
+fun add_hol4_const_mapping bthy bconst internal isaconst thy =
+    let
+	val thy = case opt_get_output_thy thy of
+		      "" => thy
+		    | output_thy => if internal
+				    then add_hol4_cmove (Sign.full_name (sign_of thy) bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+				    else thy
+	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+	val curmaps = HOL4ConstMaps.get thy
+	val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,None)),curmaps)
+	val upd_thy = HOL4ConstMaps.put newmaps thy
+    in
+	upd_thy
+    end;
+
+fun add_hol4_const_renaming bthy bconst newname thy =
+    let
+	val currens = HOL4Rename.get thy
+	val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
+	val newrens = StringPair.update_new(((bthy,bconst),newname),currens)
+	val upd_thy = HOL4Rename.put newrens thy
+    in
+	upd_thy
+    end;
+
+fun get_hol4_const_renaming bthy bconst thy =
+    let
+	val currens = HOL4Rename.get thy
+    in
+	StringPair.lookup(currens,(bthy,bconst))
+    end;
+
+fun get_hol4_const_mapping bthy bconst thy =
+    let
+	val bconst = case get_hol4_const_renaming bthy bconst thy of
+			 Some name => name
+		       | None => bconst
+	val maps = HOL4ConstMaps.get thy
+    in
+	StringPair.lookup(maps,(bthy,bconst))
+    end
+
+fun add_hol4_const_wt_mapping bthy bconst internal isaconst typ thy =
+    let
+	val thy = case opt_get_output_thy thy of
+		      "" => thy
+		    | output_thy => if internal
+				    then add_hol4_cmove (Sign.full_name (sign_of thy) bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+				    else thy
+	val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+	val curmaps = HOL4ConstMaps.get thy
+	val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,Some typ)),curmaps)
+	val upd_thy = HOL4ConstMaps.put newmaps thy
+    in
+	upd_thy
+    end;
+
+fun add_hol4_type_mapping bthy bconst internal isaconst thy =
+    let
+	val curmaps = HOL4TypeMaps.get thy
+	val _ = message ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+	val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst)),curmaps)
+	val upd_thy = HOL4TypeMaps.put newmaps thy
+    in
+	upd_thy
+    end;
+
+fun add_hol4_pending bthy bthm hth thy =
+    let
+	val thmname = Sign.full_name (sign_of thy) bthm
+	val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
+	val curpend = HOL4Pending.get thy
+	val newpend = StringPair.update_new(((bthy,bthm),hth),curpend)
+	val upd_thy = HOL4Pending.put newpend thy
+	val thy' = case opt_get_output_thy upd_thy of
+		       "" => add_hol4_mapping bthy bthm thmname upd_thy
+		     | output_thy =>
+		       let
+			   val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
+		       in
+			   upd_thy |> add_hol4_move thmname new_thmname
+				   |> add_hol4_mapping bthy bthm new_thmname
+		       end
+    in
+	thy'
+    end;
+
+fun get_hol4_theorem thyname thmname thy =
+    let
+	val isathms = HOL4Thms.get thy
+    in
+	StringPair.lookup (isathms,(thyname,thmname))
+    end
+
+fun add_hol4_theorem thyname thmname hth thy =
+    let
+	val _ = message ("Adding HOL4 theorem " ^ thyname ^ "." ^ thmname)
+	val isathms = HOL4Thms.get thy
+	val isathms' = StringPair.update_new (((thyname,thmname),hth),isathms)
+	val thy' = HOL4Thms.put isathms' thy
+    in
+	thy'
+    end
+
+fun export_hol4_pending thy =
+    let
+	val rews    = HOL4Rewrites.get thy
+	val outthy  = get_output_thy thy
+	fun process (thy,((bthy,bthm),hth as (_,thm))) =
+	    let
+		val sg      = sign_of thy
+		val thm1 = rewrite_rule (map (transfer_sg sg) rews) (transfer_sg sg thm)
+		val thm2 = standard thm1
+		val (thy2,_) = PureThy.store_thm((bthm,thm2),[]) thy
+		val thy5 = add_hol4_theorem bthy bthm hth thy2
+	    in
+		thy5
+	    end
+
+	val pending = HOL4Pending.get thy
+	val thy1    = StringPair.foldl process (thy,pending)
+	val thy2    = HOL4Pending.put (StringPair.empty) thy1
+    in
+	thy2
+    end;
+
+fun setup_dump (dir,thyname) thy =
+    HOL4Dump.put (dir,thyname,[]) thy
+
+fun add_dump str thy =
+    let
+	val (dir,thyname,curdump) = HOL4Dump.get thy
+    in
+	HOL4Dump.put (dir,thyname,str::curdump) thy
+    end
+
+fun flush_dump thy =
+    let
+	val (dir,thyname,dumpdata) = HOL4Dump.get thy
+	val os = TextIO.openOut (OS.Path.joinDirFile {dir=dir,
+						      file=thyname ^ ".thy"})
+	val _  = app (fn s => TextIO.output(os,s ^ "\n\n")) (rev dumpdata)
+	val  _ = TextIO.closeOut os
+    in
+	HOL4Dump.put ("","",[]) thy
+    end
+
+fun set_generating_thy thyname thy =
+    case HOL4DefThy.get thy of
+	NoImport => HOL4DefThy.put (Generating thyname) thy
+      | _ => error "Import already in progess"
+
+fun set_replaying_thy thyname thy =
+    case HOL4DefThy.get thy of
+	NoImport => HOL4DefThy.put (Replaying thyname) thy
+      | _ => error "Import already in progess"
+
+fun clear_import_thy thy =
+    case HOL4DefThy.get thy of
+	NoImport => error "No import in progress"
+      | _ => HOL4DefThy.put NoImport thy
+
+fun get_generating_thy thy =
+    case HOL4DefThy.get thy of
+	Generating thyname => thyname
+      | _ => error "No theory being generated"
+
+fun get_replaying_thy thy =
+    case HOL4DefThy.get thy of
+	Replaying thyname => thyname
+      | _ => error "No theory being replayed"
+
+fun get_import_thy thy =
+    case HOL4DefThy.get thy of
+	Replaying thyname => thyname
+      | Generating thyname => thyname
+      | _ => error "No theory being imported"
+
+fun should_ignore thyname thy thmname =
+    case get_hol4_mapping thyname thmname thy of
+	Some None => true 
+      | _ => false
+
+val trans_string =
+    let
+	fun quote s = "\"" ^ s ^ "\""
+	fun F [] = []
+	  | F (#"\\" :: cs) = patch #"\\" cs
+	  | F (#"\"" :: cs) = patch #"\"" cs
+	  | F (c     :: cs) = c :: F cs
+	and patch c rest = #"\\" :: c :: F rest
+    in
+	quote o String.implode o F o String.explode
+    end
+
+fun dump_import_thy thyname thy =
+    let
+	val output_dir = get_output_dir thy
+	val output_thy = get_output_thy thy
+	val import_segment = get_import_segment thy
+	val sg = sign_of thy
+	val os = TextIO.openOut (OS.Path.joinDirFile {dir=output_dir,
+						      file=thyname ^ ".imp"})
+	fun out s = TextIO.output(os,s)
+	val (ignored,mapped) =
+	    StringPair.foldl (fn ((ign,map),((bthy,bthm),v)) =>
+				 if bthy = thyname
+				 then (case v of
+					   None => (bthm::ign,map)
+					 | Some w => (ign,(bthm,w)::map))
+				 else (ign,map))
+				 (([],[]),HOL4Maps.get thy)
+	val constmaps =
+	    StringPair.foldl (fn (map,((bthy,bthm),v)) =>
+				 if bthy = thyname
+				 then (bthm,v)::map
+				 else map)
+				 ([],HOL4ConstMaps.get thy)
+
+	val constrenames =
+	    StringPair.foldl (fn (map,((bthy,bthm),v)) =>
+				 if bthy = thyname
+				 then (bthm,v)::map
+				 else map)
+				 ([],HOL4Rename.get thy)
+
+	val typemaps =
+	    StringPair.foldl (fn (map,((bthy,bthm),v)) =>
+				 if bthy = thyname
+				 then (bthm,v)::map
+				 else map)
+				 ([],HOL4TypeMaps.get thy)
+
+	val defmaps =
+	    StringPair.foldl (fn (map,((bthy,bthm),v)) =>
+				 if bthy = thyname
+				 then (bthm,v)::map
+				 else map)
+				 ([],HOL4DefMaps.get thy)
+
+	fun new_name internal isa =
+	    if internal
+	    then
+		let
+		    val paths = NameSpace.unpack isa
+		    val i = drop(length paths - 2,paths)
+		in
+		    case i of
+			[seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
+		      | _ => error "hol4rews.dump internal error"
+		end
+	    else
+		isa
+
+	val _ = out "import\n\n"
+
+	val _ = out ("import_segment " ^ trans_string import_segment ^ "\n\n")
+	val _ = if null defmaps
+		then ()
+		else out "def_maps"
+	val _ = app (fn (hol,isa) =>
+			out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string isa))) defmaps
+	val _ = if null defmaps
+		then ()
+		else out "\n\n"
+
+	val _ = if null typemaps
+		then ()
+		else out "type_maps"
+	val _ = app (fn (hol,(internal,isa)) =>
+			out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (new_name internal isa)))) typemaps
+	val _ = if null typemaps
+		then ()
+		else out "\n\n"
+
+	val _ = if null constmaps
+		then ()
+		else out "const_maps"
+	val _ = app (fn (hol,(internal,isa,opt_ty)) =>
+			(out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
+			 case opt_ty of
+			     Some ty => out (" :: \"" ^ (string_of_ctyp (ctyp_of sg ty)) ^ "\"")
+			   | None => ())) constmaps
+	val _ = if null constmaps
+		then ()
+		else out "\n\n"
+
+	val _ = if null constrenames
+		then ()
+		else out "const_renames"
+	val _ = app (fn (old,new) =>
+			out ("\n  " ^ (trans_string old) ^ " > " ^ (trans_string new))) constrenames
+	val _ = if null constrenames
+		then ()
+		else out "\n\n"
+
+	val _ = if null mapped
+		then ()
+		else out "thm_maps"
+	val _ = app (fn (hol,isa) => out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string isa))) mapped
+	val _ = if null mapped 
+		then ()
+		else out "\n\n"
+
+	val _ = if null ignored
+		then ()
+		else out "ignore_thms"
+	val _ = app (fn ign => out ("\n  " ^ (trans_string ign))) ignored
+	val _ = if null ignored
+		then ()
+		else out "\n\n"
+
+	val _ = out "end\n"
+	val _ = TextIO.closeOut os
+    in
+	thy
+    end
+
+fun set_used_names names thy =
+    let
+	val unames = HOL4UNames.get thy
+    in
+	case unames of
+	    [] => HOL4UNames.put names thy
+	  | _ => error "hol4rews.set_used_names called on initialized data!"
+    end
+
+val clear_used_names = HOL4UNames.put HOL4UNamesArgs.empty
+
+fun get_defmap thyname const thy =
+    let
+	val maps = HOL4DefMaps.get thy
+    in
+	StringPair.lookup(maps,(thyname,const))
+    end
+
+fun add_defmap thyname const defname thy =
+    let
+	val _ = message ("Adding defmap " ^ thyname ^ "." ^ const ^ " --> " ^ defname)
+	val maps = HOL4DefMaps.get thy
+	val maps' = StringPair.update_new(((thyname,const),defname),maps)
+	val thy' = HOL4DefMaps.put maps' thy
+    in
+	thy'
+    end
+
+fun get_defname thyname name thy =
+    let
+	val maps = HOL4DefMaps.get thy
+	fun F dname = (dname,add_defmap thyname name dname thy)
+    in
+	case StringPair.lookup(maps,(thyname,name)) of
+	    Some thmname => (thmname,thy)
+	  | None =>
+	    let
+		val used = HOL4UNames.get thy
+		val defname = def_name name
+		val pdefname = name ^ "_primdef"
+	    in
+		if not (defname mem used)
+		then F defname                 (* name_def *)
+		else if not (pdefname mem used)
+		then F pdefname                (* name_primdef *)
+		else F (variant used pdefname) (* last resort *)
+	    end
+    end
+
+local
+    fun handle_meta [x as Appl[Appl[Constant "_constrain", Constant "==", _],_,_]] = x
+      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant "all", _],_]] = x
+      | handle_meta [x as Appl[Appl[Constant "_constrain", Constant "==>", _],_,_]] = x
+      | handle_meta [x] = Appl[Constant "Trueprop",x]
+in
+val smarter_trueprop_parsing = [("Trueprop",handle_meta)]
+end
+
+local
+    fun initial_maps thy =
+	thy |> add_hol4_type_mapping "min" "bool" false "bool"
+	    |> add_hol4_type_mapping "min" "fun" false "fun"
+	    |> add_hol4_type_mapping "min" "ind" false "Nat.ind"
+	    |> add_hol4_const_mapping "min" "=" false "op ="
+	    |> add_hol4_const_mapping "min" "==>" false "op -->"
+	    |> add_hol4_const_mapping "min" "@" false "Hilbert_Choice.Eps"
+in
+val hol4_setup =
+    [HOL4Rewrites.init,
+     HOL4Maps.init,
+     HOL4UNames.init,
+     HOL4DefMaps.init,
+     HOL4Moves.init,
+     HOL4CMoves.init,
+     HOL4ConstMaps.init,
+     HOL4Rename.init,
+     HOL4TypeMaps.init,
+     HOL4Pending.init,
+     HOL4Thms.init,
+     HOL4Dump.init,
+     HOL4DefThy.init,
+     HOL4Imports.init,
+     ImportSegment.init,
+     initial_maps,
+     Attrib.add_attributes [("hol4rew",
+			     (Attrib.no_args add_hol4_rewrite,
+			      Attrib.no_args Attrib.undef_local_attribute),
+			     "HOL4 rewrite rule")]]
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import_package.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,61 @@
+signature IMPORT_PACKAGE =
+sig
+    val import_meth: Args.src -> Proof.context -> Proof.method
+    val setup      : (theory -> theory) list
+    val debug      : bool ref
+end
+
+structure ImportPackage :> IMPORT_PACKAGE =
+struct
+
+val debug = ref false
+fun message s = if !debug then writeln s else ()
+
+val string_of_thm = Library.setmp print_mode [] string_of_thm
+val string_of_cterm = Library.setmp print_mode [] string_of_cterm
+
+fun import_tac (thyname,thmname) =
+    if !SkipProof.quick_and_dirty
+    then SkipProof.cheat_tac
+    else
+     fn thy =>
+     fn th =>
+	let
+	    val sg = sign_of_thm th
+	    val prem = hd (prems_of th)
+	    val _ = message ("Import trying to prove " ^ (string_of_cterm (cterm_of sg prem)))
+	    val int_thms = fst (Replay.setup_int_thms thyname thy)
+	    val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
+	    val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
+	    val thm = snd (ProofKernel.to_isa_thm hol4thm)
+	    val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
+	    val thm = equal_elim rew thm
+	    val prew = ProofKernel.rewrite_hol4_term prem thy
+	    val prem' = #2 (Logic.dest_equals (prop_of prew))
+	    val _ = message ("Import proved " ^ (string_of_thm thm))  
+	in
+	    case Shuffler.set_prop thy prem' [("",thm)] of
+		Some (_,thm) =>
+		let
+		    val _ = if prem' aconv (prop_of thm)
+			    then message "import: Terms match up"
+			    else message "import: Terms DO NOT match up"
+		    val thm' = equal_elim (symmetric prew) thm
+		    val res = bicompose true (false,thm',0) 1 th
+		in
+		    res
+		end
+	      | None => (message "import: set_prop didn't succeed"; no_tac th)
+	end
+	
+val import_meth = Method.simple_args (Args.name -- Args.name)
+		  (fn arg =>
+		   fn ctxt =>
+		      let
+			  val thy = ProofContext.theory_of ctxt
+		      in
+			  Method.SIMPLE_METHOD (import_tac arg thy)
+		      end)
+
+val setup = [Method.add_method("import",import_meth,"Import HOL4 theorem")]
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import_syntax.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,227 @@
+signature HOL4_IMPORT_SYNTAX =
+sig
+    type token = OuterSyntax.token
+
+    val import_segment: token list -> (theory -> theory) * token list
+    val import_theory : token list -> (theory -> theory) * token list
+    val end_import    : token list -> (theory -> theory) * token list
+
+    val setup_theory  : token list -> (theory -> theory) * token list
+    val end_setup     : token list -> (theory -> theory) * token list
+
+    val thm_maps      : token list -> (theory -> theory) * token list
+    val ignore_thms   : token list -> (theory -> theory) * token list
+    val type_maps     : token list -> (theory -> theory) * token list
+    val def_maps      : token list -> (theory -> theory) * token list
+    val const_maps    : token list -> (theory -> theory) * token list
+    val const_moves   : token list -> (theory -> theory) * token list
+    val const_renames : token list -> (theory -> theory) * token list
+
+    val setup        : unit -> unit
+end
+
+structure HOL4ImportSyntax :> HOL4_IMPORT_SYNTAX =
+struct
+
+type token = OuterSyntax.token
+
+local structure P = OuterParse and K = OuterSyntax.Keyword in
+
+(* Parsers *)
+
+val import_segment = P.name >> set_import_segment
+
+val import_theory = P.name >> (fn thyname =>
+			       fn thy =>
+				  thy |> set_generating_thy thyname
+				      |> Theory.add_path thyname
+				      |> add_dump (";setup_theory " ^ thyname))
+
+val end_import = Scan.succeed
+		     (fn thy =>
+			 let
+			     val thyname = get_generating_thy thy
+			     val segname = get_import_segment thy
+			     val (int_thms,facts) = Replay.setup_int_thms thyname thy
+			     val thmnames = filter (not o should_ignore thyname thy) facts
+			 in
+			     thy |> clear_import_thy
+				 |> set_segment thyname segname
+				 |> set_used_names facts
+				 |> Replay.import_thms thyname int_thms thmnames
+				 |> clear_used_names
+				 |> export_hol4_pending
+				 |> Theory.parent_path
+				 |> dump_import_thy thyname
+				 |> add_dump ";end_setup"
+			 end)
+
+val ignore_thms = Scan.repeat1 P.name
+			       >> (fn ignored =>
+				   fn thy =>
+				      let
+					  val thyname = get_import_thy thy
+				      in
+					  foldl (fn (thy,thmname) => ignore_hol4 thyname thmname thy) (thy,ignored)
+				      end)
+
+val thm_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+			    >> (fn thmmaps =>
+				fn thy =>
+				   let
+				       val thyname = get_import_thy thy
+				   in
+				       foldl (fn (thy,(hol,isa)) => add_hol4_mapping thyname hol isa thy) (thy,thmmaps)
+				   end)
+
+val type_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+			     >> (fn typmaps =>
+				 fn thy =>
+				    let
+					val thyname = get_import_thy thy
+				    in
+					foldl (fn (thy,(hol,isa)) => add_hol4_type_mapping thyname hol false isa thy) (thy,typmaps)
+				    end)
+
+val def_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname)
+			    >> (fn defmaps =>
+				fn thy =>
+				   let
+				       val thyname = get_import_thy thy
+				   in
+				       foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
+				   end)
+
+val const_renames = Scan.repeat1 (P.name --| P.$$$ ">" -- P.name)
+				 >> (fn renames =>
+				     fn thy =>
+					let
+					    val thyname = get_import_thy thy
+					in
+					    foldl (fn (thy,(hol,isa)) => add_hol4_const_renaming thyname hol isa thy) (thy,renames)
+					end)
+																      
+val const_maps = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+			      >> (fn constmaps =>
+				  fn thy =>
+				     let
+					 val thyname = get_import_thy thy
+				     in
+					 foldl (fn (thy,((hol,isa),None)) => add_hol4_const_mapping thyname hol false isa thy
+						 | (thy,((hol,isa),Some ty)) => add_hol4_const_wt_mapping thyname hol false isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
+				     end)
+
+val const_moves = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+			       >> (fn constmaps =>
+				   fn thy =>
+				      let
+					  val thyname = get_import_thy thy
+				      in
+					  foldl (fn (thy,((hol,isa),None)) => add_hol4_const_mapping thyname hol true isa thy
+						  | (thy,((hol,isa),Some ty)) => add_hol4_const_wt_mapping thyname hol true isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
+				      end)
+
+fun import_thy s =
+    let
+	val is = TextIO.openIn(s ^ ".imp")
+	val inp = TextIO.inputAll is
+	val _ = TextIO.closeIn is
+	val orig_source = Source.of_string inp
+	val symb_source = Symbol.source false orig_source
+	val lexes = ref (OuterLex.make_lexicon ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"],
+			 Scan.empty_lexicon)
+	val get_lexes = fn () => !lexes
+	val token_source = OuterLex.source false get_lexes (Position.line 1) symb_source
+	val token_list = filter (not o (OuterLex.is_kind OuterLex.Space)) (Source.exhaust token_source)
+	val import_segmentP = OuterParse.$$$ "import_segment" |-- import_segment
+	val type_mapsP = OuterParse.$$$ "type_maps" |-- type_maps
+	val def_mapsP = OuterParse.$$$ "def_maps" |-- def_maps
+	val const_mapsP = OuterParse.$$$ "const_maps" |-- const_maps
+	val const_movesP = OuterParse.$$$ "const_moves" |-- const_moves
+	val const_renamesP = OuterParse.$$$ "const_renames" |-- const_renames
+	val ignore_thmsP = OuterParse.$$$ "ignore_thms" |-- ignore_thms
+	val thm_mapsP = OuterParse.$$$ "thm_maps" |-- thm_maps
+	val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
+	val importP = OuterParse.$$$ "import" |-- Scan.repeat parser --| OuterParse.$$$ "end"
+	fun apply [] thy = thy
+	  | apply (f::fs) thy = apply fs (f thy)
+    in
+	apply (set_replaying_thy s::Theory.add_path s::(fst (importP token_list)))
+    end
+
+val setup_theory = P.name
+		       >>
+		       (fn thyname =>
+			fn thy =>
+			   case HOL4DefThy.get thy of
+			       NoImport => thy |> import_thy thyname
+			     | Generating _ => error "Currently generating a theory!"
+			     | Replaying _ => thy |> clear_import_thy |> import_thy thyname)
+
+val end_setup = Scan.succeed (fn thy =>
+				 let
+				     val thyname = get_import_thy thy
+				     val segname = get_import_segment thy
+				 in
+				     thy |> set_segment thyname segname
+					 |> clear_import_thy
+					 |> Theory.parent_path
+				 end)
+
+val set_dump  = P.string -- P.string   >> setup_dump
+val fl_dump   = Scan.succeed flush_dump
+val append_dump = (P.verbatim || P.string) >> add_dump
+
+val parsers = 
+  [OuterSyntax.command "import_segment"
+		       "Set import segment name"
+		       K.thy_decl (import_segment >> Toplevel.theory),
+   OuterSyntax.command "import_theory"
+		       "Set default HOL4 theory name"
+		       K.thy_decl (import_theory >> Toplevel.theory),
+   OuterSyntax.command "end_import"
+		       "End HOL4 import"
+		       K.thy_decl (end_import >> Toplevel.theory),
+   OuterSyntax.command "setup_theory"
+		       "Setup HOL4 theory replaying"
+		       K.thy_decl (setup_theory >> Toplevel.theory),
+   OuterSyntax.command "end_setup"
+		       "End HOL4 setup"
+		       K.thy_decl (end_setup >> Toplevel.theory),
+   OuterSyntax.command "setup_dump"
+		       "Setup the dump file name"
+		       K.thy_decl (set_dump >> Toplevel.theory),
+   OuterSyntax.command "append_dump"
+		       "Add text to dump file"
+		       K.thy_decl (append_dump >> Toplevel.theory),
+   OuterSyntax.command "flush_dump"
+		       "Write the dump to file"
+		       K.thy_decl (fl_dump >> Toplevel.theory),
+   OuterSyntax.command "ignore_thms"
+		       "Theorems to ignore in next HOL4 theory import"
+		       K.thy_decl (ignore_thms >> Toplevel.theory),
+   OuterSyntax.command "type_maps"
+		       "Map HOL4 type names to existing Isabelle/HOL type names"
+		       K.thy_decl (type_maps >> Toplevel.theory),
+   OuterSyntax.command "def_maps"
+		       "Map HOL4 constant names to their primitive definitions"
+		       K.thy_decl (def_maps >> Toplevel.theory),
+   OuterSyntax.command "thm_maps"
+		       "Map HOL4 theorem names to existing Isabelle/HOL theorem names"
+		       K.thy_decl (thm_maps >> Toplevel.theory),
+   OuterSyntax.command "const_renames"
+		       "Rename HOL4 const names"
+		       K.thy_decl (const_renames >> Toplevel.theory),
+   OuterSyntax.command "const_moves"
+		       "Rename HOL4 const names to other HOL4 constants"
+		       K.thy_decl (const_moves >> Toplevel.theory),
+   OuterSyntax.command "const_maps"
+		       "Map HOL4 const names to existing Isabelle/HOL const names"
+		       K.thy_decl (const_maps >> Toplevel.theory)]
+
+fun setup () = (OuterSyntax.add_keywords[">"];
+		OuterSyntax.add_parsers parsers)
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/proof_kernel.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,2142 @@
+signature ProofKernel =
+sig
+    type hol_type
+    type tag
+    type term
+    type thm
+    type ('a,'b) subst
+	 
+    type proof_info
+    datatype proof = Proof of proof_info * proof_content
+	 and proof_content
+	   = PRefl of term
+	   | PInstT of proof * (hol_type,hol_type) subst
+	   | PSubst of proof list * term * proof
+	   | PAbs of proof * term
+	   | PDisch of proof * term
+	   | PMp of proof * proof
+	   | PHyp of term
+	   | PAxm of string * term
+	   | PDef of string * string * term
+	   | PTmSpec of string * string list * proof
+	   | PTyDef of string * string * proof
+	   | PTyIntro of string * string * string * string * term * term * proof
+	   | POracle of tag * term list * term
+	   | PDisk
+	   | PSpec of proof * term
+	   | PInst of proof * (term,term) subst
+	   | PGen of proof * term
+	   | PGenAbs of proof * term option * term list
+	   | PImpAS of proof * proof
+	   | PSym of proof
+	   | PTrans of proof * proof
+	   | PComb of proof * proof
+	   | PEqMp of proof * proof
+	   | PEqImp of proof
+	   | PExists of proof * term * term
+	   | PChoose of term * proof * proof
+	   | PConj of proof * proof
+	   | PConjunct1 of proof
+	   | PConjunct2 of proof
+	   | PDisj1 of proof * term
+	   | PDisj2 of proof * term
+	   | PDisjCases of proof * proof * proof
+	   | PNotI of proof
+	   | PNotE of proof
+	   | PContr of proof * term
+
+    exception PK of string * string
+
+    val get_proof_dir: string -> theory -> string
+    val debug : bool ref
+    val disk_info_of : proof -> (string * string) option
+    val set_disk_info_of : proof -> string -> string -> unit
+    val mk_proof : proof_content -> proof
+    val content_of : proof -> proof_content
+    val import_proof : string -> string -> theory -> (theory -> term) option * (theory -> proof)
+
+    val rewrite_hol4_term: Term.term -> theory -> Thm.thm
+
+    val type_of : term -> hol_type
+
+    val get_thm  : string -> string         -> theory -> (theory * thm option)
+    val get_def  : string -> string -> term -> theory -> (theory * thm option)
+    val get_axiom: string -> string         -> theory -> (theory * thm option)
+
+    val store_thm : string -> string -> thm -> theory -> theory * thm
+
+    val to_isa_thm : thm -> (term * term) list * Thm.thm
+    val to_isa_term: term -> Term.term
+
+    val REFL : term -> theory -> theory * thm
+    val ASSUME : term -> theory -> theory * thm
+    val INST_TYPE : (hol_type,hol_type) subst -> thm -> theory -> theory * thm
+    val INST : (term,term)subst -> thm -> theory -> theory * thm
+    val EQ_MP : thm -> thm -> theory -> theory * thm
+    val EQ_IMP_RULE : thm -> theory -> theory * thm
+    val SUBST : thm list -> term -> thm -> theory -> theory * thm
+    val DISJ_CASES : thm -> thm -> thm -> theory -> theory * thm
+    val DISJ1: thm -> term -> theory -> theory * thm
+    val DISJ2: term -> thm -> theory -> theory * thm
+    val IMP_ANTISYM: thm -> thm -> theory -> theory * thm
+    val SYM : thm -> theory -> theory * thm
+    val MP : thm -> thm -> theory -> theory * thm
+    val GEN : term -> thm -> theory -> theory * thm
+    val CHOOSE : term -> thm -> thm -> theory -> theory * thm
+    val EXISTS : term -> term -> thm -> theory -> theory * thm
+    val ABS : term -> thm -> theory -> theory * thm
+    val GEN_ABS : term option -> term list -> thm -> theory -> theory * thm
+    val TRANS : thm -> thm -> theory -> theory * thm
+    val CCONTR : term -> thm -> theory -> theory * thm
+    val CONJ : thm -> thm -> theory -> theory * thm
+    val CONJUNCT1: thm -> theory -> theory * thm
+    val CONJUNCT2: thm -> theory -> theory * thm
+    val NOT_INTRO: thm -> theory -> theory * thm
+    val NOT_ELIM : thm -> theory -> theory * thm
+    val SPEC : term -> thm -> theory -> theory * thm
+    val COMB : thm -> thm -> theory -> theory * thm
+    val DISCH: term -> thm -> theory -> theory * thm
+
+    val type_introduction: string -> string -> string -> string -> string -> term * term -> thm -> theory -> theory * thm
+
+    val new_definition : string -> string -> term -> theory -> theory * thm
+    val new_specification : string -> string -> string list -> thm -> theory -> theory * thm
+    val new_type_definition : string -> string -> string -> thm -> theory -> theory * thm
+    val new_axiom : string -> term -> theory -> theory * thm
+
+end
+
+structure ProofKernel :> ProofKernel =
+struct
+type hol_type = Term.typ
+type term = Term.term
+datatype tag = Tag of string list
+type ('a,'b) subst = ('a * 'b) list
+datatype thm = HOLThm of (Term.term * Term.term) list * Thm.thm
+
+datatype proof_info
+  = Info of {disk_info: (string * string) option ref}
+	    
+datatype proof = Proof of proof_info * proof_content
+     and proof_content
+       = PRefl of term
+       | PInstT of proof * (hol_type,hol_type) subst
+       | PSubst of proof list * term * proof
+       | PAbs of proof * term
+       | PDisch of proof * term
+       | PMp of proof * proof
+       | PHyp of term
+       | PAxm of string * term
+       | PDef of string * string * term
+       | PTmSpec of string * string list * proof
+       | PTyDef of string * string * proof
+       | PTyIntro of string * string * string * string * term * term * proof
+       | POracle of tag * term list * term
+       | PDisk
+       | PSpec of proof * term
+       | PInst of proof * (term,term) subst
+       | PGen of proof * term
+       | PGenAbs of proof * term option * term list
+       | PImpAS of proof * proof
+       | PSym of proof
+       | PTrans of proof * proof
+       | PComb of proof * proof
+       | PEqMp of proof * proof
+       | PEqImp of proof
+       | PExists of proof * term * term
+       | PChoose of term * proof * proof
+       | PConj of proof * proof
+       | PConjunct1 of proof
+       | PConjunct2 of proof
+       | PDisj1 of proof * term
+       | PDisj2 of proof * term
+       | PDisjCases of proof * proof * proof
+       | PNotI of proof
+       | PNotE of proof
+       | PContr of proof * term
+
+exception PK of string * string
+fun ERR f mesg = PK (f,mesg)
+
+fun print_exn e = 
+    case e of
+	PK (m,s) => (writeln ("PK (" ^ m ^ "): " ^ s); raise e)
+      | _ => Goals.print_exn e
+
+(* Compatibility. *)
+
+fun quote s = "\"" ^ s ^ "\""
+
+fun alphanum str = case String.explode str of
+		       first::rest => Char.isAlpha first andalso forall (fn c => Char.isAlphaNum c orelse c = #"_") rest
+		     | _ => error "ProofKernel.alphanum: Empty constname??"
+
+fun mk_syn c = if alphanum c then NoSyn else Syntax.literal c
+
+val keywords = ["open"]
+fun quotename c = if alphanum c andalso not (c mem keywords) then c else quote c
+
+fun smart_string_of_cterm ct =
+    let
+	val {sign,t,T,...} = rep_cterm ct
+        (* Hack to avoid parse errors with Trueprop *)
+	val ct  = (cterm_of sign (HOLogic.dest_Trueprop t)
+		   handle TERM _ => ct)
+	fun match cu = t aconv (term_of cu)
+	fun G 0 = I
+	  | G 1 = Library.setmp show_types true
+	  | G 2 = Library.setmp show_all_types true
+	  | G _ = error ("ProofKernel.smart_string_of_cterm internal error: " ^ (G 2 string_of_cterm ct))
+	fun F sh_br n =
+	    let
+		val str = Library.setmp show_brackets sh_br (G n string_of_cterm) ct
+		val cu  = transform_error (read_cterm sign) (str,T)
+	    in
+		if match cu
+		then quote str
+		else F false (n+1)
+	    end
+	    handle ERROR_MESSAGE mesg =>
+		   if String.isPrefix "Ambiguous" mesg andalso
+		      not sh_br
+		   then F true n
+		   else F false (n+1)
+    in
+	transform_error (Library.setmp Syntax.ambiguity_is_error true (F false)) 0
+    end
+    handle ERROR_MESSAGE mesg =>
+	   (writeln "Exception in smart_string_of_cterm!";
+	    writeln mesg;
+	    quote (string_of_cterm ct))
+
+val smart_string_of_thm = smart_string_of_cterm o cprop_of
+
+fun prth th = writeln ((Library.setmp print_mode [] string_of_thm) th)
+fun prc ct = writeln ((Library.setmp print_mode [] string_of_cterm) ct)
+val prin = Library.setmp print_mode [] prin
+fun pth (HOLThm(ren,thm)) =
+    let
+	val _ = writeln "Renaming:"
+	val _ = app (fn(v,w) => (prin v; writeln " -->"; prin w)) ren
+	val _ = prth thm
+    in
+	()
+    end
+
+fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
+fun mk_proof p = Proof(Info{disk_info = ref None},p)
+fun content_of (Proof(_,p)) = p
+
+fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
+    disk_info := Some(thyname,thmname)
+
+structure Lib =
+struct
+fun wrap b e s = String.concat[b,s,e]
+
+fun assoc x =
+    let
+	fun F [] = raise PK("Lib.assoc","Not found")
+	  | F ((x',y)::rest) = if x = x'
+			       then y
+			       else F rest
+    in
+	F
+    end
+fun i mem L = 
+    let fun itr [] = false 
+	  | itr (a::rst) = i=a orelse itr rst 
+    in itr L end;
+    
+fun insert i L = if i mem L then L else i::L
+					
+fun mk_set [] = []
+  | mk_set (a::rst) = insert a (mk_set rst)
+		      
+fun [] union S = S
+  | S union [] = S
+  | (a::rst) union S2 = rst union (insert a S2)
+			
+fun implode_subst [] = []
+  | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
+  | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
+
+fun apboth f (x,y) = (f x,f y)
+end
+open Lib
+
+structure Tag =
+struct
+val empty_tag = Tag []
+fun read name = Tag [name]
+fun merge (Tag tag1) (Tag tag2) = Tag (Lib.union(tag1,tag2))
+end
+
+(* Acutal code. *)
+
+fun get_segment thyname l = (Lib.assoc "s" l
+			     handle PK _ => thyname)
+val get_name    = Lib.assoc "n"
+
+local
+    open LazyScan
+    infix 7 |-- --|
+    infix 5 :-- -- ^^
+    infix 3 >>
+    infix 0 ||
+in
+exception XML of string
+
+datatype xml = Elem of string * (string * string) list * xml list
+datatype XMLtype = XMLty of xml | FullType of hol_type
+datatype XMLterm = XMLtm of xml | FullTerm of term
+
+fun pair x y = (x,y)
+
+fun scan_id toks =
+    let
+        val (x,toks2) = one Char.isAlpha toks
+        val (xs,toks3) = any Char.isAlphaNum toks2
+    in
+        (String.implode (x::xs),toks3)
+    end
+
+fun scan_string str c =
+    let
+	fun F [] toks = (c,toks)
+	  | F (c::cs) toks =
+	    case LazySeq.getItem toks of
+		Some(c',toks') =>
+		if c = c'
+		then F cs toks'
+		else raise SyntaxError
+	      | None => raise SyntaxError
+    in
+	F (String.explode str)
+    end
+
+local
+    val scan_entity =
+	(scan_string "amp;" #"&")
+	    || scan_string "quot;" #"\""
+	    || scan_string "gt;" #">"
+	    || scan_string "lt;" #"<"
+in
+fun scan_nonquote toks =
+    case LazySeq.getItem toks of
+	Some (c,toks') =>
+	(case c of
+	     #"\"" => raise SyntaxError
+	   | #"&" => scan_entity toks'
+	   | c => (c,toks'))
+      | None => raise SyntaxError
+end
+
+val scan_string = $$ #"\"" |-- repeat scan_nonquote --| $$ #"\"" >>
+		     String.implode
+
+val scan_attribute = scan_id -- $$ #"=" |-- scan_string
+
+val scan_start_of_tag = $$ #"<" |-- scan_id --
+			   repeat ($$ #" " |-- scan_attribute)
+
+val scan_end_of_tag = $$ #"/" |-- $$ #">" |-- succeed []
+val scan_end_tag = $$ #"<" |-- $$ #"/" |-- scan_id --| $$ #">"
+
+fun scan_children id = $$ #">" |-- repeat scan_tag -- scan_end_tag >>
+		       (fn (chldr,id') => if id = id'
+					  then chldr
+					  else raise XML "Tag mismatch")
+and scan_tag toks =
+    let
+	val ((id,atts),toks2) = scan_start_of_tag toks
+	val (chldr,toks3) = (scan_children id || scan_end_of_tag) toks2
+    in
+	(Elem (id,atts,chldr),toks3)
+    end
+end
+
+val type_of = Term.type_of
+
+val boolT = Type("bool",[])
+val propT = Type("prop",[])
+
+fun mk_defeq name rhs thy =
+    let
+	val ty = type_of rhs
+    in
+	Logic.mk_equals (Const(Sign.intern_const (sign_of thy) name,ty),rhs)
+    end
+
+fun mk_teq name rhs thy =
+    let
+	val ty = type_of rhs
+    in
+	HOLogic.mk_eq (Const(Sign.intern_const (sign_of thy) name,ty),rhs)
+    end
+
+fun intern_const_name thyname const thy =
+    case get_hol4_const_mapping thyname const thy of
+	Some (_,cname,_) => cname
+      | None => (case get_hol4_const_renaming thyname const thy of
+		     Some cname => Sign.intern_const (sign_of thy) (thyname ^ "." ^ cname)
+		   | None => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const))
+
+fun intern_type_name thyname const thy =
+    case get_hol4_type_mapping thyname const thy of
+	Some (_,cname) => cname
+      | None => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const)
+
+fun mk_vartype name = TFree(name,["HOL.type"])
+fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args)
+
+val mk_var = Free
+
+fun dom_rng (Type("fun",[dom,rng])) = (dom,rng)
+  | dom_rng _ = raise ERR "dom_rng" "Not a functional type"
+
+fun mk_thy_const thy Thy Name Ty = Const(intern_const_name Thy Name thy,Ty)
+
+local
+    fun get_type sg thyname name =
+	case Sign.const_type sg name of
+	    Some ty => ty
+	  | None => raise ERR "get_type" (name ^ ": No such constant")
+in
+fun prim_mk_const thy Thy Name =
+    let
+	val name = intern_const_name Thy Name thy
+	val cmaps = HOL4ConstMaps.get thy
+    in
+	case StringPair.lookup(cmaps,(Thy,Name)) of
+	    Some(_,_,Some ty) => Const(name,ty)
+	  | _ => Const(name,get_type (sign_of thy) Thy name)
+    end
+end
+
+fun mk_comb(f,a) = f $ a
+fun mk_abs(x,a) = Term.lambda x a
+
+(* Needed for HOL Light *)
+fun protect_tyvarname s =
+    let
+	fun no_quest s =
+	    if Char.contains s #"?"
+	    then String.translate (fn #"?" => "q_" | c => Char.toString c) s
+	    else s
+	fun beg_prime s =
+	    if String.isPrefix "'" s
+	    then s
+	    else "'" ^ s
+    in
+	s |> no_quest |> beg_prime
+    end
+fun protect_varname s =
+    let
+	fun no_beg_underscore s =
+	    if String.isPrefix "_" s
+	    then "dummy" ^ s
+	    else s
+    in
+	s |> no_beg_underscore
+    end
+
+structure TypeNet =
+struct
+fun get_type_from_index thy thyname types is =
+    case Int.fromString is of
+	SOME i => (case Array.sub(types,i) of
+		       FullType ty => ty
+		     | XMLty xty =>
+		       let
+			   val ty = get_type_from_xml thy thyname types xty
+			   val _  = Array.update(types,i,FullType ty)
+		       in
+			   ty
+		       end)
+      | NONE => raise ERR "get_type_from_index" "Bad index"
+and get_type_from_xml thy thyname types =
+    let
+	fun gtfx (Elem("tyi",[("i",iS)],[])) =
+		 get_type_from_index thy thyname types iS
+	  | gtfx (Elem("tyc",atts,[])) =
+	    mk_thy_type thy
+			(get_segment thyname atts)
+			(get_name atts)
+			[]
+	  | gtfx (Elem("tyv",[("n",s)],[])) = mk_vartype (protect_tyvarname s)
+	  | gtfx (Elem("tya",[],(Elem("tyc",atts,[]))::tys)) =
+	    mk_thy_type thy
+			(get_segment thyname atts)
+			(get_name atts)
+			(map gtfx tys)
+	  | gtfx _ = raise ERR "get_type" "Bad type"
+    in
+	gtfx
+    end
+
+fun input_types thyname (Elem("tylist",[("i",i)],xtys)) =
+    let
+	val types = Array.array(valOf (Int.fromString i),XMLty (Elem("",[],[])))
+	fun IT _ [] = ()
+	  | IT n (xty::xtys) =
+	    (Array.update(types,n,XMLty xty);
+	     IT (n+1) xtys)
+	val _ = IT 0 xtys
+    in
+	types
+    end
+  | input_types _ _ = raise ERR "input_types" "Bad type list"
+end
+
+structure TermNet =
+struct
+fun get_term_from_index thy thyname types terms is =
+    case Int.fromString is of
+	SOME i => (case Array.sub(terms,i) of
+		       FullTerm tm => tm
+		     | XMLtm xtm =>
+		       let
+			   val tm = get_term_from_xml thy thyname types terms xtm
+			   val _  = Array.update(terms,i,FullTerm tm)
+		       in
+			   tm
+		       end)
+      | NONE => raise ERR "get_term_from_index" "Bad index"
+and get_term_from_xml thy thyname types terms =
+    let
+	fun get_type [] = None
+	  | get_type [ty] = Some (TypeNet.get_type_from_xml thy thyname types ty)
+	  | get_type _ = raise ERR "get_term" "Bad type"
+
+	fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
+	    mk_var(protect_varname name,TypeNet.get_type_from_index thy thyname types tyi)
+	  | gtfx (Elem("tmc",atts,[])) =
+	    let
+		val segment = get_segment thyname atts
+		val name = get_name atts
+	    in
+		mk_thy_const thy segment name (TypeNet.get_type_from_index thy thyname types (Lib.assoc "t" atts))
+		handle PK _ => prim_mk_const thy segment name
+	    end
+	  | gtfx (Elem("tma",[("f",tmf),("a",tma)],[])) =
+	    let
+		val f = get_term_from_index thy thyname types terms tmf
+		val a = get_term_from_index thy thyname types terms tma
+	    in
+		mk_comb(f,a)
+	    end
+	  | gtfx (Elem("tml",[("x",tmx),("a",tma)],[])) =
+	    let
+		val x = get_term_from_index thy thyname types terms tmx
+		val a = get_term_from_index thy thyname types terms tma
+	    in
+		mk_abs(x,a)
+	    end
+	  | gtfx (Elem("tmi",[("i",iS)],[])) =
+	    get_term_from_index thy thyname types terms iS
+	  | gtfx (Elem(tag,_,_)) =
+	    raise ERR "get_term" ("Not a term: "^tag)
+    in
+	gtfx
+    end
+
+fun input_terms thyname types (Elem("tmlist",[("i",i)],xtms)) =
+    let
+	val terms = Array.array(valOf(Int.fromString i),XMLtm (Elem("",[],[])))
+
+	fun IT _ [] = ()
+	  | IT n (xtm::xtms) =
+	    (Array.update(terms,n,XMLtm xtm);
+	     IT (n+1) xtms)
+	val _ = IT 0 xtms
+    in
+	terms
+    end
+  | input_terms _ _ _ = raise ERR "input_terms" "Bad term list"
+end
+
+fun get_proof_dir (thyname:string) thy =
+    let
+	val import_segment =
+	    case get_segment2 thyname thy of
+		Some seg => seg
+	      | None => get_import_segment thy
+	val defpath = [(getenv "ISABELLE_HOME_USER") ^ "/proofs"]
+	val path =
+	    case getenv "PROOF_DIRS" of
+		"" => defpath
+	      | path => (space_explode ":" path) @ defpath
+	fun find [] = error ("Unable to find import segment " ^ import_segment)
+	  | find (p::ps) = (let
+				val dir = p ^ "/" ^ import_segment
+			    in
+				if OS.FileSys.isDir dir
+				then dir
+				else find ps
+			    end) handle OS.SysErr _ => find ps
+    in
+	find path ^ "/" ^ thyname
+    end
+			   
+fun proof_file_name thyname thmname thy =
+    let
+	val path = get_proof_dir thyname thy
+	val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
+    in
+	String.concat[path,"/",thmname,".prf"]
+    end
+	
+fun xml_to_proof thyname types terms prf thy =
+    let
+	val xml_to_hol_type = TypeNet.get_type_from_xml thy thyname types
+	val xml_to_term = TermNet.get_term_from_xml thy thyname types terms
+
+	fun index_to_term is =
+	    TermNet.get_term_from_index thy thyname types terms is
+
+	fun x2p (Elem("prefl",[("i",is)],[])) = mk_proof (PRefl (index_to_term is))
+	  | x2p (Elem("pinstt",[],p::lambda)) =
+	    let
+		val p = x2p p
+		val lambda = implode_subst (map xml_to_hol_type lambda)
+	    in
+		mk_proof (PInstT(p,lambda))
+	    end
+	  | x2p (Elem("psubst",[("i",is)],prf::prfs)) =
+	    let
+		val tm = index_to_term is
+		val prf = x2p prf
+		val prfs = map x2p prfs
+	    in
+		mk_proof (PSubst(prfs,tm,prf))
+	    end
+	  | x2p (Elem("pabs",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val t = index_to_term is
+	    in
+		mk_proof (PAbs (p,t))
+	    end
+	  | x2p (Elem("pdisch",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val t = index_to_term is
+	    in
+		mk_proof (PDisch (p,t))
+	    end
+	  | x2p (Elem("pmp",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PMp(p1,p2))
+	    end
+	  | x2p (Elem("phyp",[("i",is)],[])) = mk_proof (PHyp (index_to_term is))
+	  | x2p (Elem("paxiom",[("n",n),("i",is)],[])) =
+	    mk_proof (PAxm(n,index_to_term is))
+	  | x2p (Elem("pfact",atts,[])) =
+	    let
+		val thyname = get_segment thyname atts
+		val thmname = get_name atts
+		val p = mk_proof PDisk
+		val _  = set_disk_info_of p thyname thmname
+	    in
+		p
+	    end
+	  | x2p (Elem("pdef",[("s",seg),("n",name),("i",is)],[])) =
+	    mk_proof (PDef(seg,name,index_to_term is))
+	  | x2p (Elem("ptmspec",[("s",seg)],p::names)) =
+	    let
+		val names = map (fn Elem("name",[("n",name)],[]) => name
+				  | _ => raise ERR "x2p" "Bad proof (ptmspec)") names
+	    in
+		mk_proof (PTmSpec(seg,names,x2p p))
+	    end
+	  | x2p (Elem("ptyintro",[("s",seg),("n",name),("a",abs_name),("r",rep_name)],[xP,xt,p])) =
+	    let
+		val P = xml_to_term xP
+		val t = xml_to_term xt
+	    in
+		mk_proof (PTyIntro(seg,name,abs_name,rep_name,P,t,x2p p))
+	    end
+	  | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
+	    mk_proof (PTyDef(seg,name,x2p p))
+	  | x2p (xml as Elem("poracle",[],chldr)) =
+	    let
+		val (oracles,terms) = Library.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
+		val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | xml => raise ERR "x2p" "bad oracle") oracles
+		val (c,asl) = case terms of
+				  [] => raise ERR "x2p" "Bad oracle description"
+				| (hd::tl) => (hd,tl)
+		val tg = foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) (ors,Tag.empty_tag)
+	    in
+		mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c))
+	    end
+	  | x2p (Elem("pspec",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val tm = index_to_term is
+	    in
+		mk_proof (PSpec(p,tm))
+	    end
+	  | x2p (Elem("pinst",[],p::theta)) =
+	    let
+		val p = x2p p
+		val theta = implode_subst (map xml_to_term theta)
+	    in
+		mk_proof (PInst(p,theta))
+	    end
+	  | x2p (Elem("pgen",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val tm = index_to_term is
+	    in
+		mk_proof (PGen(p,tm))
+	    end
+	  | x2p (Elem("pgenabs",[],prf::tms)) =
+	    let
+		val p = x2p prf
+		val tml = map xml_to_term tms
+	    in
+		mk_proof (PGenAbs(p,None,tml))
+	    end
+	  | x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
+	    let
+		val p = x2p prf
+		val tml = map xml_to_term tms
+	    in
+		mk_proof (PGenAbs(p,Some (index_to_term is),tml))
+	    end
+	  | x2p (Elem("pimpas",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PImpAS(p1,p2))
+	    end
+	  | x2p (Elem("psym",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PSym p)
+	    end
+	  | x2p (Elem("ptrans",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PTrans(p1,p2))
+	    end
+	  | x2p (Elem("pcomb",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PComb(p1,p2))
+	    end
+	  | x2p (Elem("peqmp",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PEqMp(p1,p2))
+	    end
+	  | x2p (Elem("peqimp",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PEqImp p)
+	    end
+	  | x2p (Elem("pexists",[("e",ise),("w",isw)],[prf])) =
+	    let
+		val p = x2p prf
+		val ex = index_to_term ise
+		val w = index_to_term isw
+	    in
+		mk_proof (PExists(p,ex,w))
+	    end
+	  | x2p (Elem("pchoose",[("i",is)],[prf1,prf2])) =
+	    let
+		val v  = index_to_term is
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PChoose(v,p1,p2))
+	    end
+	  | x2p (Elem("pconj",[],[prf1,prf2])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+	    in
+		mk_proof (PConj(p1,p2))
+	    end
+	  | x2p (Elem("pconjunct1",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PConjunct1 p)
+	    end
+	  | x2p (Elem("pconjunct2",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PConjunct2 p)
+	    end
+	  | x2p (Elem("pdisj1",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val t = index_to_term is
+	    in
+		mk_proof (PDisj1 (p,t))
+	    end
+	  | x2p (Elem("pdisj2",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val t = index_to_term is
+	    in
+		mk_proof (PDisj2 (p,t))
+	    end
+	  | x2p (Elem("pdisjcases",[],[prf1,prf2,prf3])) =
+	    let
+		val p1 = x2p prf1
+		val p2 = x2p prf2
+		val p3 = x2p prf3
+	    in
+		mk_proof (PDisjCases(p1,p2,p3))
+	    end
+	  | x2p (Elem("pnoti",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PNotI p)
+	    end
+	  | x2p (Elem("pnote",[],[prf])) =
+	    let
+		val p = x2p prf
+	    in
+		mk_proof (PNotE p)
+	    end
+	  | x2p (Elem("pcontr",[("i",is)],[prf])) =
+	    let
+		val p = x2p prf
+		val t = index_to_term is
+	    in
+		mk_proof (PContr (p,t))
+	    end
+	  | x2p xml = raise ERR "x2p" "Bad proof"
+    in
+	x2p prf
+    end
+
+fun import_proof thyname thmname thy =
+    let
+	val is = TextIO.openIn(proof_file_name thyname thmname thy)
+	val (proof_xml,_) = scan_tag (LazySeq.of_instream is)
+	val _ = TextIO.closeIn is
+    in
+	case proof_xml of
+	    Elem("proof",[],xtypes::xterms::prf::rest) =>
+	    let
+		val types = TypeNet.input_types thyname xtypes
+		val terms = TermNet.input_terms thyname types xterms
+	    in
+		(case rest of
+		     [] => None
+		   | [xtm] => Some (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
+		   | _ => raise ERR "import_proof" "Bad argument list",
+		 xml_to_proof thyname types terms prf)
+	    end
+	  | _ => raise ERR "import_proof" "Bad proof"
+    end
+
+fun uniq_compose m th i st =
+    let
+	val res = bicompose false (false,th,m) i st
+    in
+	case Seq.pull res of
+	    Some (th,rest) => (case Seq.pull rest of
+				   Some _ => raise ERR "uniq_compose" "Not unique!"
+				 | None => th)
+	  | None => raise ERR "uniq_compose" "No result"
+    end
+
+val reflexivity_thm = thm "refl"
+val substitution_thm = thm "subst"
+val mp_thm = thm "mp"
+val imp_antisym_thm = thm "light_imp_as"
+val disch_thm = thm "impI"
+val ccontr_thm = thm "ccontr"
+
+val meta_eq_to_obj_eq_thm = thm "meta_eq_to_obj_eq"
+
+val gen_thm = thm "HOLallI"
+val choose_thm = thm "exE"
+val exists_thm = thm "exI"
+val conj_thm = thm "conjI"
+val conjunct1_thm = thm "conjunct1"
+val conjunct2_thm = thm "conjunct2"
+val spec_thm = thm "spec"
+val disj_cases_thm = thm "disjE"
+val disj1_thm = thm "disjI1"
+val disj2_thm = thm "disjI2"
+
+local
+    val th = thm "not_def"
+    val sg = sign_of_thm th
+    val pp = reflexive (cterm_of sg (Const("Trueprop",boolT-->propT)))
+in
+val not_elim_thm = combination pp th
+end
+
+val not_intro_thm = symmetric not_elim_thm
+val abs_thm = thm "ext"
+val trans_thm = thm "trans"
+val symmetry_thm = thm "sym"
+val transitivity_thm = thm "trans"
+val eqmp_thm = thm "iffD1"
+val eqimp_thm = thm "HOL4Setup.eq_imp"
+val comb_thm = thm "cong"
+
+(* Beta-eta normalizes a theorem (only the conclusion, not the *
+hypotheses!)  *)
+
+fun beta_eta_thm th =
+    let
+	val th1 = Thm.equal_elim (Thm.beta_conversion true (cprop_of th))  th
+	val th2 = Thm.equal_elim (Thm.eta_conversion       (cprop_of th1)) th1
+    in
+	th2
+    end
+
+fun implies_elim_all th =
+    foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th)
+
+fun norm_hyps th =
+    th |> beta_eta_thm
+       |> implies_elim_all
+       |> implies_intr_hyps
+
+fun mk_GEN v th sg =
+    let
+	val c = HOLogic.dest_Trueprop (concl_of th)
+	val cv = cterm_of sg v
+	val lc = Term.lambda v c
+	val clc = Thm.cterm_of sg lc
+	val cvty = ctyp_of_term cv
+	val th1 = implies_elim_all th
+	val th2 = beta_eta_thm (forall_intr cv th1)
+	val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [Some cvty] [Some clc] gen_thm))
+	val c = prop_of th3
+	val vname = fst(dest_Free v)
+	val (cold,cnew) = case c of
+			      tpc $ (Const("All",allT) $ Abs(oldname,ty,body)) =>
+			      (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
+			    | tpc $ (Const("All",allT) $ rest) => (tpc,tpc)
+			    | _ => raise ERR "mk_GEN" "Unknown conclusion"
+	val th4 = Thm.rename_boundvars cold cnew th3
+	val res = implies_intr_hyps th4
+    in
+	res
+    end
+
+(* rotate left k places, leaving the first j and last l premises alone
+*)
+
+fun permute_prems j k 0 th = Thm.permute_prems j k th
+  | permute_prems j k l th =
+    th |> Thm.permute_prems 0 (~l)
+       |> Thm.permute_prems (j+l) k
+       |> Thm.permute_prems 0 l
+
+fun rearrange sg tm th =
+    let
+	val tm' = Pattern.beta_eta_contract tm
+	fun find []      n = permute_prems 0 1 0 (implies_intr (Thm.cterm_of sg tm) th)
+	  | find (p::ps) n = if tm' aconv (Pattern.beta_eta_contract p)
+			     then permute_prems n 1 0 th
+			     else find ps (n+1)
+    in
+	find (prems_of th) 0
+    end
+
+fun zip (x::xs) (y::ys) = (x,y)::(zip xs ys)
+  | zip [] [] = []
+  | zip _ _ = raise ERR "zip" "arguments not of same length"
+
+fun mk_INST dom rng th =
+    th |> forall_intr_list dom
+       |> forall_elim_list rng
+
+fun apply_tyinst_typ tyinst =
+    let
+	fun G (ty as TFree _) =
+	    (case try (Lib.assoc ty) tyinst of
+		 Some ty' => ty'
+	       | None => ty)
+	  | G (Type(tyname,tys)) = Type(tyname,map G tys)
+	  | G (TVar _) = raise ERR "apply_tyinst_typ" "Scematic variable found"
+    in
+	G
+    end
+
+fun apply_tyinst_term tyinst =
+    let
+	val G = apply_tyinst_typ tyinst
+	fun F (tm as Bound _) = tm
+	  | F (tm as Free(vname,ty)) = Free(vname,G ty)
+	  | F (tm as Const(vname,ty)) = Const(vname,G ty)
+	  | F (tm1 $ tm2) = (F tm1) $ (F tm2)
+	  | F (Abs(vname,ty,body)) = Abs(vname,G ty,F body)
+	  | F (Var _) = raise ERR "apply_tyinst_term" "Schematic variable found"
+    in
+	F
+    end
+
+fun apply_inst_term tminst =
+    let
+	fun F (tm as Bound _) = tm
+	  | F (tm as Free _) =
+	    (case try (Lib.assoc tm) tminst of
+		 Some tm' => tm'
+	       | None => tm)
+	  | F (tm as Const _) = tm
+	  | F (tm1 $ tm2) = (F tm1) $ (F tm2)
+	  | F (Abs(vname,ty,body)) = Abs(vname,ty,F body)
+	  | F (Var _) = raise ERR "apply_inst_term" "Schematic variable found"
+    in
+	F
+    end
+
+val collect_vars =
+    let
+	fun F vars (Bound _) = vars
+	  | F vars (tm as Free _) =
+	    if tm mem vars
+	    then vars
+	    else (tm::vars)
+	  | F vars (Const _) = vars
+	  | F vars (tm1 $ tm2) = F (F vars tm1) tm2
+	  | F vars (Abs(_,_,body)) = F vars body
+	  | F vars (Var _) = raise ERR "collect_vars" "Schematic variable found"
+    in
+	F []
+    end
+
+(* Code for disambiguating variablenames (wrt. types) *)
+
+val disamb_info_empty = {vars=[],rens=[]}
+
+fun rens_of {vars,rens} = rens
+
+fun name_of_var (Free(vname,_)) = vname
+  | name_of_var _ = raise ERR "name_of_var" "Not a variable"
+
+fun disamb_helper {vars,rens} tm =
+    let
+	val varstm = collect_vars tm
+	fun process (v as Free(vname,ty),(vars,rens,inst)) =
+	    if v mem vars
+	    then (vars,rens,inst)
+	    else (case try (Lib.assoc v) rens of
+		      Some vnew => (vars,rens,(v,vnew)::inst)
+		    | None => if exists (fn Free(vname',_) => vname = vname' | _ => raise ERR "disamb_helper" "Bad varlist") vars
+			      then
+				  let
+				      val tmnames = map name_of_var varstm
+				      val varnames = map name_of_var vars
+				      val (dom,rng) = ListPair.unzip rens
+				      val rensnames = (map name_of_var dom) @ (map name_of_var rng)
+				      val instnames = map name_of_var (snd (ListPair.unzip inst))
+				      val allnames = tmnames @ varnames @ rensnames @ instnames
+				      val vnewname = Term.variant allnames vname
+				      val vnew = Free(vnewname,ty)
+				  in
+				      (vars,(v,vnew)::rens,(v,vnew)::inst)
+				  end
+			      else (v::vars,rens,inst))
+	  | process _ = raise ERR "disamb_helper" "Internal error"
+
+	val (vars',rens',inst) =
+	    foldr process (varstm,(vars,rens,[]))
+    in
+	({vars=vars',rens=rens'},inst)
+    end
+
+fun disamb_term_from info tm =
+    let
+	val (info',inst) = disamb_helper info tm
+    in
+	(info',apply_inst_term inst tm)
+    end
+
+fun swap (x,y) = (y,x)
+
+fun has_ren (HOLThm([],_)) = false
+  | has_ren _ = true
+
+fun prinfo {vars,rens} = (writeln "Vars:";
+			  app prin vars;
+			  writeln "Renaming:";
+			  app (fn(x,y)=>(prin x; writeln " -->"; prin y)) rens)
+
+fun disamb_thm_from info (hth as HOLThm(rens,thm)) =
+    let
+	val inv_rens = map swap rens
+	val orig_thm = apply_inst_term inv_rens (prop_of thm)
+	val (info',inst) = disamb_helper info orig_thm
+	val rens' = map (apsnd (apply_inst_term inst)) inv_rens
+	val (dom,rng) = ListPair.unzip (rens' @ inst)
+	val sg = sign_of_thm thm
+	val thm' = mk_INST (map (cterm_of sg) dom) (map (cterm_of sg) rng) thm
+    in
+	(info',thm')
+    end
+
+fun disamb_terms_from info tms =
+    foldr (fn (tm,(info,tms)) =>
+	      let
+		  val (info',tm') = disamb_term_from info tm
+	      in
+		  (info',tm'::tms)
+	      end)
+	  (tms,(info,[]))
+
+fun disamb_thms_from info hthms =
+    foldr (fn (hthm,(info,thms)) =>
+	      let
+		  val (info',tm') = disamb_thm_from info hthm
+	      in
+		  (info',tm'::thms)
+	      end)
+	  (hthms,(info,[]))
+
+fun disamb_term tm   = disamb_term_from disamb_info_empty tm
+fun disamb_terms tms = disamb_terms_from disamb_info_empty tms
+fun disamb_thm thm   = disamb_thm_from disamb_info_empty thm
+fun disamb_thms thms = disamb_thms_from disamb_info_empty thms
+
+fun norm_hthm sg (hth as HOLThm([],_)) = hth
+  | norm_hthm sg (hth as HOLThm(rens,th)) =
+    let
+	val vars = collect_vars (prop_of th)
+	val (rens',inst,_) =
+	    foldr (fn((ren as (vold as Free(vname,_),vnew)),
+		      (rens,inst,vars)) =>
+		     (case Library.find_first (fn Free(v,_) => v = vname | _ => false) vars of
+			  Some v' => if v' = vold
+				     then (rens,(vnew,vold)::inst,vold::vars)
+				     else (ren::rens,(vold,vnew)::inst,vnew::vars)
+			| None => (rens,(vnew,vold)::inst,vold::vars))
+		   | _ => raise ERR "norm_hthm" "Internal error")
+		  (rens,([],[],vars))
+	val (dom,rng) = ListPair.unzip inst
+	val th' = mk_INST (map (cterm_of sg) dom) (map (cterm_of sg) rng) th
+	val nvars = collect_vars (prop_of th')
+	val rens' = filter (fn(_,v) => v mem nvars) rens
+	val res = HOLThm(rens',th')
+    in
+	res
+    end
+
+(* End of disambiguating code *)
+
+val debug = ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+val conjE_helper = Thm.permute_prems 0 1 conjE
+
+fun get_hol4_thm thyname thmname thy =
+    case get_hol4_theorem thyname thmname thy of
+	Some hth => Some (HOLThm hth)
+      | None => 
+	let
+	    val pending = HOL4Pending.get thy
+	in
+	    case StringPair.lookup (pending,(thyname,thmname)) of
+		Some hth => Some (HOLThm hth)
+	      | None => None
+	end
+
+fun non_trivial_term_consts tm =
+    filter (fn c => not (c = "Trueprop" orelse
+			 c = "All" orelse
+			 c = "op -->" orelse
+			 c = "op &" orelse
+			 c = "op =")) (Term.term_consts tm) 
+
+fun match_consts t (* th *) =
+    let
+	fun add_consts (Const (c, _), cs) =
+	    (case c of
+		 "op =" => "==" ins_string cs
+	       | "op -->" => "==>" ins_string cs
+	       | "All" => cs
+	       | "all" => cs
+	       | "op &" => cs
+	       | "Trueprop" => cs
+	       | _ => c ins_string cs)
+	  | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
+	  | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
+	  | add_consts (_, cs) = cs
+	val t_consts = add_consts(t,[])
+    in
+	fn th => eq_set(t_consts,add_consts(prop_of th,[]))
+    end
+
+fun split_name str =
+    let
+	val sub = Substring.all str
+	val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
+	val (newstr,u) = apboth Substring.string (Substring.splitr (fn c => c = #"_") f)
+    in
+	if not (idx = "") andalso u = "_"
+	then Some (newstr,valOf(Int.fromString idx))
+	else None
+    end
+    handle _ => None
+
+fun rewrite_hol4_term t thy =
+    let
+	val sg = sign_of thy
+
+	val hol4rews1 = map (transfer_sg sg) (HOL4Rewrites.get thy)
+	val hol4ss = empty_ss setmksimps single addsimps hol4rews1
+    in
+	transfer_sg sg (Simplifier.full_rewrite hol4ss (cterm_of sg t))
+    end
+
+
+fun get_isabelle_thm thyname thmname hol4conc thy =
+    let
+	val sg = sign_of thy
+
+	val (info,hol4conc') = disamb_term hol4conc
+	val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
+	val isaconc =
+	    case concl_of i2h_conc of
+		Const("==",_) $ lhs $ _ => lhs
+	      | _ => error "get_isabelle_thm" "Bad rewrite rule"
+	val _ = (message "Original conclusion:";
+		 if_debug prin hol4conc';
+		 message "Modified conclusion:";
+		 if_debug prin isaconc)
+
+	fun mk_res th = HOLThm(rens_of info,equal_elim i2h_conc th)
+    in
+	case get_hol4_mapping thyname thmname thy of
+	    Some (Some thmname) =>
+	    let
+		val _ = message ("Looking for " ^ thmname)
+		val th1 = (Some (transform_error (PureThy.get_thm thy) thmname)
+			   handle ERROR_MESSAGE _ =>
+				  (case split_name thmname of
+				       Some (listname,idx) => (Some (nth_elem(idx-1,PureThy.get_thms thy listname))
+							       handle _ => None)
+				     | None => None))
+	    in
+		case th1 of
+		    Some th2 =>
+		    (case Shuffler.set_prop thy isaconc [(thmname,th2)] of
+			 Some (_,th) => (message "YES";(thy, Some (mk_res th)))
+		       | None => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
+		  | None => (message "NO1";error "get_isabelle_thm" "Bad mapping")
+	    end
+	  | Some None => error ("Trying to access ignored theorem " ^ thmname)
+	  | None =>
+	    let
+		val _ = (message "Looking for conclusion:";
+			 if_debug prin isaconc)
+		val cs = non_trivial_term_consts isaconc
+		val _ = (message "Looking for consts:";
+			 message (String.concat cs))
+		val pot_thms = Shuffler.find_potential thy isaconc
+		val _ = message ((Int.toString (length pot_thms)) ^ " potential theorems")
+	    in
+		case Shuffler.set_prop thy isaconc pot_thms of
+		    Some (isaname,th) =>
+		    let
+			val hth as HOLThm args = mk_res th
+			val thy' =  thy |> add_hol4_theorem thyname thmname args
+					|> add_hol4_mapping thyname thmname isaname
+		    in
+			(thy',Some hth)
+		    end
+		  | None => (thy,None)
+	    end
+    end
+    handle _ => (thy,None)
+
+fun get_thm thyname thmname thy =
+    case get_hol4_thm thyname thmname thy of
+	Some hth => (thy,Some hth)
+      | None => ((case fst (import_proof thyname thmname thy) of
+		      Some f => get_isabelle_thm thyname thmname (f thy) thy
+		    | None => (thy,None))
+		 handle e as IO.Io _ => (thy,None)
+		      | e as PK _ => (thy,None))
+
+fun rename_const thyname thy name =
+    case get_hol4_const_renaming thyname name thy of
+	Some cname => cname
+      | None => name
+
+fun get_def thyname constname rhs thy =
+    let
+	val constname = rename_const thyname thy constname
+	val (thmname,thy') = get_defname thyname constname thy
+	val _ = message ("Looking for definition " ^ thyname ^ "." ^ thmname)
+    in
+	get_isabelle_thm thyname thmname (mk_teq (thyname ^ "." ^ constname) rhs thy') thy'
+    end
+
+fun get_axiom thyname axname thy =
+    case get_thm thyname axname thy of
+	arg as (_,Some _) => arg
+      | _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
+
+fun intern_store_thm gen_output thyname thmname hth thy =
+    let
+	val sg = sign_of thy
+	val (hth' as HOLThm (args as (_,th))) = norm_hthm sg hth
+	val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
+		else ()
+	val rew = rewrite_hol4_term (concl_of th) thy
+	val th  = equal_elim rew th
+	val thy' = add_hol4_pending thyname thmname args thy
+	val thy2 = if gen_output
+		   then add_dump ("lemma " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^ "\n  by (import " ^ thyname ^ " " ^ thmname ^ ")") thy'
+		   else thy'
+    in
+	(thy2,hth')
+    end
+
+val store_thm = intern_store_thm true
+
+fun mk_REFL ctm =
+    let
+	val cty = Thm.ctyp_of_term ctm
+    in
+	Drule.instantiate' [Some cty] [Some ctm] reflexivity_thm
+    end
+
+fun REFL tm thy =
+    let
+	val _ = message "REFL:"
+	val (info,tm') = disamb_term tm
+	val sg = sign_of thy
+	val ctm = Thm.cterm_of sg tm'
+	val res = HOLThm(rens_of info,mk_REFL ctm)
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun ASSUME tm thy =
+    let
+	val _ = message "ASSUME:"
+	val (info,tm') = disamb_term tm
+	val sg = sign_of thy
+	val ctm = Thm.cterm_of sg (HOLogic.mk_Trueprop tm')
+	val th = Thm.trivial ctm
+	val res = HOLThm(rens_of info,th)
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun INST_TYPE lambda (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "INST_TYPE:"
+	val _ = if_debug pth hth
+	val sg = sign_of thy
+	val tys_before = add_term_tfrees (prop_of th,[])
+	val th1 = varifyT th
+	val tys_after = add_term_tvars (prop_of th1,[])
+	val tyinst = map (fn (bef,(i,_)) =>
+			     (case try (Lib.assoc (TFree bef)) lambda of
+				  Some ty => (i,ctyp_of sg ty)
+				| None => (i,ctyp_of sg (TFree bef))
+			     ))
+			 (zip tys_before tys_after)
+	val res = Drule.instantiate (tyinst,[]) th1
+	val appty = apboth (apply_tyinst_term lambda)
+	val hth = HOLThm(map appty rens,res)
+	val res = norm_hthm sg hth
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun INST sigma hth thy =
+    let
+	val _ = message "INST:"
+	val _ = if_debug (app (fn (x,y) => (prin x; prin y))) sigma
+	val _ = if_debug pth hth
+	val sg = sign_of thy
+	val (sdom,srng) = ListPair.unzip sigma
+	val (info,th) = disamb_thm hth
+	val (info',srng') = disamb_terms_from info srng
+	val rens = rens_of info'
+	val sdom' = map (apply_inst_term rens) sdom
+	val th1 = mk_INST (map (cterm_of sg) sdom') (map (cterm_of sg) srng') th
+	val res = HOLThm(rens,th1)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun EQ_IMP_RULE (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "EQ_IMP_RULE:"
+	val _ = if_debug pth hth
+	val res = HOLThm(rens,th RS eqimp_thm)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun mk_EQ_MP th1 th2 = [beta_eta_thm th1,beta_eta_thm th2] MRS eqmp_thm
+
+fun EQ_MP hth1 hth2 thy =
+    let
+	val _ = message "EQ_MP:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val res = HOLThm(rens_of info,mk_EQ_MP th1 th2)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun mk_COMB th1 th2 sg =
+    let
+	val (f,g) = case concl_of th1 of
+			_ $ (Const("op =",_) $ f $ g) => (f,g)
+		      | _ => raise ERR "mk_COMB" "First theorem not an equality"
+	val (x,y) = case concl_of th2 of
+			_ $ (Const("op =",_) $ x $ y) => (x,y)
+		      | _ => raise ERR "mk_COMB" "Second theorem not an equality"
+	val fty = type_of f
+	val (fd,fr) = dom_rng fty
+	val comb_thm' = Drule.instantiate'
+			    [Some (ctyp_of sg fd),Some (ctyp_of sg fr)]
+			    [Some (cterm_of sg f),Some (cterm_of sg g),
+			     Some (cterm_of sg x),Some (cterm_of sg y)] comb_thm
+    in
+	[th1,th2] MRS comb_thm'
+    end
+
+fun SUBST rews ctxt hth thy =
+    let
+	val _ = message "SUBST:"
+	val _ = if_debug (app pth) rews
+	val _ = if_debug prin ctxt
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info1,ctxt') = disamb_term_from info ctxt
+	val (info2,rews') = disamb_thms_from info1 rews
+
+	val sg = sign_of thy
+	val cctxt = cterm_of sg ctxt'
+	fun subst th [] = th
+	  | subst th (rew::rews) = subst (mk_COMB th rew sg) rews
+	val res = HOLThm(rens_of info2,mk_EQ_MP (subst (mk_REFL cctxt) rews') th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun DISJ_CASES hth hth1 hth2 thy =
+    let
+	val _ = message "DISJ_CASES:"
+	val _ = if_debug (app pth) [hth,hth1,hth2]
+	val (info,th) = disamb_thm hth
+	val (info1,th1) = disamb_thm_from info hth1
+	val (info2,th2) = disamb_thm_from info1 hth2
+	val sg = sign_of thy
+	val th1 = norm_hyps th1
+	val th2 = norm_hyps th2
+	val (l,r) = case concl_of th of
+			_ $ (Const("op |",_) $ l $ r) => (l,r)
+		      | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
+	val th1' = rearrange sg (HOLogic.mk_Trueprop l) th1
+	val th2' = rearrange sg (HOLogic.mk_Trueprop r) th2
+	val res1 = th RS disj_cases_thm
+	val res2 = uniq_compose ((nprems_of th1')-1) th1' ((nprems_of th)+1) res1
+	val res3 = uniq_compose ((nprems_of th2')-1) th2' (nprems_of res2) res2
+	val res  = HOLThm(rens_of info2,res3)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun DISJ1 hth tm thy =
+    let
+	val _ = message "DISJ1:"
+	val _ = if_debug pth hth
+	val _ = if_debug prin tm
+	val (info,th) = disamb_thm hth
+	val (info',tm') = disamb_term_from info tm
+	val sg = sign_of thy
+	val ct = Thm.cterm_of sg tm'
+	val disj1_thm' = Drule.instantiate' [] [None,Some ct] disj1_thm
+	val res = HOLThm(rens_of info',th RS disj1_thm')
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun DISJ2 tm hth thy =
+    let
+	val _ = message "DISJ1:"
+	val _ = if_debug prin tm
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',tm') = disamb_term_from info tm
+	val sg = sign_of thy
+	val ct = Thm.cterm_of sg tm'
+	val disj2_thm' = Drule.instantiate' [] [None,Some ct] disj2_thm
+	val res = HOLThm(rens_of info',th RS disj2_thm')
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun IMP_ANTISYM hth1 hth2 thy =
+    let
+	val _ = message "IMP_ANTISYM:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val th = [beta_eta_thm th1,beta_eta_thm th2] MRS imp_antisym_thm
+	val res = HOLThm(rens_of info,th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun SYM (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "SYM:"
+	val _ = if_debug pth hth
+	val th = th RS symmetry_thm
+	val res = HOLThm(rens,th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun MP hth1 hth2 thy =
+    let
+	val _ = message "MP:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val th = [beta_eta_thm th1,beta_eta_thm th2] MRS mp_thm
+	val res = HOLThm(rens_of info,th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun CONJ hth1 hth2 thy =
+    let
+	val _ = message "CONJ:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val th = [th1,th2] MRS conj_thm
+	val res = HOLThm(rens_of info,th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun CONJUNCT1 (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "CONJUNCT1:"
+	val _ = if_debug pth hth
+	val res = HOLThm(rens,th RS conjunct1_thm)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun CONJUNCT2 (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "CONJUNCT1:"
+	val _ = if_debug pth hth
+	val res = HOLThm(rens,th RS conjunct2_thm)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun EXISTS ex wit hth thy =
+    let
+	val _ = message "EXISTS:"
+	val _ = if_debug prin ex
+	val _ = if_debug prin wit
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',[ex',wit']) = disamb_terms_from info [ex,wit]
+	val sg = sign_of thy
+	val cwit = cterm_of sg wit'
+	val cty = ctyp_of_term cwit
+	val a = case ex' of
+		    (Const("Ex",_) $ a) => a
+		  | _ => raise ERR "EXISTS" "Argument not existential"
+	val ca = cterm_of sg a
+	val exists_thm' = beta_eta_thm (Drule.instantiate' [Some cty] [Some ca,Some cwit] exists_thm)
+	val th1 = beta_eta_thm th
+	val th2 = implies_elim_all th1
+	val th3 = th2 COMP exists_thm'
+	val th  = implies_intr_hyps th3
+	val res = HOLThm(rens_of info',th)
+	val _   = message "RESULT:"
+	val _   = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun CHOOSE v hth1 hth2 thy =
+    let
+	val _ = message "CHOOSE:"
+	val _ = if_debug prin v
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val (info',v') = disamb_term_from info v
+	fun strip 0 _ th = th
+	  | strip n (p::ps) th =
+	    strip (n-1) ps (implies_elim th (assume p))
+	  | strip _ _ _ = raise ERR "CHOOSE" "strip error"
+	val sg = sign_of thy
+	val cv = cterm_of sg v'
+	val th2 = norm_hyps th2
+	val cvty = ctyp_of_term cv
+	val _$c = concl_of th2
+	val cc = cterm_of sg c
+	val a = case concl_of th1 of
+		    _ $ (Const("Ex",_) $ a) => a
+		  | _ => raise ERR "CHOOSE" "Conclusion not existential"
+	val ca = cterm_of (sign_of_thm th1) a
+	val choose_thm' = beta_eta_thm (Drule.instantiate' [Some cvty] [Some ca,Some cc] choose_thm)
+	val th21 = rearrange sg (HOLogic.mk_Trueprop (a $ v')) th2
+	val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
+	val th23 = beta_eta_thm (forall_intr cv th22)
+	val th11 = implies_elim_all (beta_eta_thm th1)
+	val th' = th23 COMP (th11 COMP choose_thm')
+	val th = implies_intr_hyps th'
+	val res = HOLThm(rens_of info',th)
+	val _   = message "RESULT:"
+	val _   = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun GEN v hth thy =
+    let
+	val _ = message "GEN:"
+	val _ = if_debug prin v
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',v') = disamb_term_from info v
+	val res = HOLThm(rens_of info',mk_GEN v' th (sign_of thy))
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun SPEC tm hth thy =
+    let
+	val _ = message "SPEC:"
+	val _ = if_debug prin tm
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',tm') = disamb_term_from info tm
+	val sg = sign_of thy
+	val ctm = Thm.cterm_of sg tm'
+	val cty = Thm.ctyp_of_term ctm
+	val spec' = Drule.instantiate' [Some cty] [None,Some ctm] spec_thm
+	val th = th RS spec'
+	val res = HOLThm(rens_of info',th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun COMB hth1 hth2 thy =
+    let
+	val _ = message "COMB:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val sg = sign_of thy
+	val res = HOLThm(rens_of info,mk_COMB th1 th2 sg)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun TRANS hth1 hth2 thy =
+    let
+	val _ = message "TRANS:"
+	val _ = if_debug pth hth1
+	val _ = if_debug pth hth2
+	val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+	val th = [th1,th2] MRS trans_thm
+	val res = HOLThm(rens_of info,th)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+	
+
+fun CCONTR tm hth thy =
+    let
+	val _ = message "SPEC:"
+	val _ = if_debug prin tm
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',tm') = disamb_term_from info tm
+	val th = norm_hyps th
+	val sg = sign_of thy
+	val ct = cterm_of sg tm'
+	val th1 = rearrange sg (HOLogic.mk_Trueprop (Const("Not",boolT-->boolT) $ tm')) th
+	val ccontr_thm' = Drule.instantiate' [] [Some ct] ccontr_thm
+	val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
+	val res = HOLThm(rens_of info',res1)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun mk_ABS v th sg =
+    let
+	val cv = cterm_of sg v
+	val th1 = implies_elim_all (beta_eta_thm th)
+	val (f,g) = case concl_of th1 of
+			_ $ (Const("op =",_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+		      | _ => raise ERR "mk_ABS" "Bad conclusion"
+	val (fd,fr) = dom_rng (type_of f)
+	val abs_thm' = Drule.instantiate' [Some (ctyp_of sg fd), Some (ctyp_of sg fr)] [Some (cterm_of sg f), Some (cterm_of sg g)] abs_thm
+	val th2 = forall_intr cv th1
+	val th3 = th2 COMP abs_thm'
+	val res = implies_intr_hyps th3
+    in
+	res
+    end
+
+fun ABS v hth thy =
+    let
+	val _ = message "ABS:"
+	val _ = if_debug prin v
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',v') = disamb_term_from info v
+	val sg = sign_of thy
+	val res = HOLThm(rens_of info',mk_ABS v' th sg)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun GEN_ABS copt vlist hth thy =
+    let
+	val _ = message "GEN_ABS:"
+	val _ = case copt of
+		    Some c => if_debug prin c
+		  | None => ()
+	val _ = if_debug (app prin) vlist
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',vlist') = disamb_terms_from info vlist
+	val sg = sign_of thy
+	val th1 =
+	    case copt of
+		Some (c as Const(cname,cty)) =>
+		let
+		    fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!"
+		      | inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty
+							    then ty2
+							    else ty
+		      | inst_type ty1 ty2 (ty as Type(name,tys)) =
+			Type(name,map (inst_type ty1 ty2) tys)
+		in
+		    foldr (fn (v,th) =>
+			      let
+				  val cdom = fst (dom_rng (fst (dom_rng cty)))
+				  val vty  = type_of v
+				  val newcty = inst_type cdom vty cty
+				  val cc = cterm_of sg (Const(cname,newcty))
+			      in
+				  mk_COMB (mk_REFL cc) (mk_ABS v th sg) sg
+			      end) (vlist',th)
+		end
+	      | Some _ => raise ERR "GEN_ABS" "Bad constant"
+	      | None => 
+		foldr (fn (v,th) => mk_ABS v th sg) (vlist',th)
+	val res = HOLThm(rens_of info',th1)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun NOT_INTRO (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "NOT_INTRO:"
+	val _ = if_debug pth hth
+	val sg = sign_of thy
+	val th1 = implies_elim_all (beta_eta_thm th)
+	val a = case concl_of th1 of
+		    _ $ (Const("op -->",_) $ a $ Const("False",_)) => a
+		  | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
+	val ca = cterm_of sg a
+	val th2 = equal_elim (Drule.instantiate' [] [Some ca] not_intro_thm) th1
+	val res = HOLThm(rens,implies_intr_hyps th2)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun NOT_ELIM (hth as HOLThm(rens,th)) thy =
+    let
+	val _ = message "NOT_INTRO:"
+	val _ = if_debug pth hth
+	val sg = sign_of thy
+	val th1 = implies_elim_all (beta_eta_thm th)
+	val a = case concl_of th1 of
+		    _ $ (Const("Not",_) $ a) => a
+		  | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
+	val ca = cterm_of sg a
+	val th2 = equal_elim (Drule.instantiate' [] [Some ca] not_elim_thm) th1
+	val res = HOLThm(rens,implies_intr_hyps th2)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+fun DISCH tm hth thy =
+    let
+	val _ = message "DISCH:"
+	val _ = if_debug prin tm
+	val _ = if_debug pth hth
+	val (info,th) = disamb_thm hth
+	val (info',tm') = disamb_term_from info tm
+	val prems = prems_of th
+	val sg = sign_of thy
+	val th1 = beta_eta_thm th
+	val th2 = implies_elim_all th1
+	val th3 = implies_intr (cterm_of sg (HOLogic.mk_Trueprop tm')) th2
+	val th4 = th3 COMP disch_thm
+	val res = HOLThm(rens_of info',implies_intr_hyps th4)
+	val _ = message "RESULT:"
+	val _ = if_debug pth res
+    in
+	(thy,res)
+    end
+
+val spaces = String.concat o separate " "
+
+fun new_definition thyname constname rhs thy =
+    let
+	val constname = rename_const thyname thy constname
+	val _ = warning ("Introducing constant " ^ constname)
+	val (thmname,thy) = get_defname thyname constname thy
+	val (info,rhs') = disamb_term rhs
+	val ctype = type_of rhs'
+	val csyn = mk_syn constname
+	val thy1 = case HOL4DefThy.get thy of
+		       Replaying _ => thy
+		     | _ => Theory.add_consts_i [(constname,ctype,csyn)] thy
+	val eq = mk_defeq constname rhs' thy1
+	val (thy2,thms) = PureThy.add_defs_i false [((thmname,eq),[])] thy1
+	val def_thm = hd thms
+	val thm' = def_thm RS meta_eq_to_obj_eq_thm
+	val (thy',th) = (thy2, thm')
+	val fullcname = Sign.intern_const (sign_of thy') constname
+	val thy'' = add_hol4_const_mapping thyname constname true fullcname thy'
+	val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
+	val sg = sign_of thy''
+	val rew = rewrite_hol4_term eq thy''
+	val crhs = cterm_of sg (#2 (Logic.dest_equals (prop_of rew)))
+	val thy22 = if (def_name constname) = thmname
+		    then
+			add_dump ("constdefs\n  " ^ (quotename constname) ^ " :: \"" ^ (string_of_ctyp (ctyp_of sg ctype)) ^ "\" " ^ (Syntax.string_of_mixfix csyn) ^ "\n  " ^ (smart_string_of_cterm crhs)) thy''
+		    else
+			add_dump ("consts\n  " ^ (quotename constname) ^ " :: \"" ^ (string_of_ctyp (ctyp_of sg ctype)) ^
+				  "\" " ^ (Syntax.string_of_mixfix csyn) ^ "\n\ndefs\n  " ^ (quotename thmname) ^ ": " ^ (smart_string_of_cterm crhs))
+				 thy''
+
+	val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
+		     Some (_,res) => HOLThm(rens_of linfo,res)
+		   | None => raise ERR "new_definition" "Bad conclusion"
+	val fullname = Sign.full_name sg thmname
+	val thy22' = case opt_get_output_thy thy22 of
+			 "" => add_hol4_mapping thyname thmname fullname thy22
+		       | output_thy =>
+			 let
+			     val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
+			 in
+			     thy22 |> add_hol4_move fullname moved_thmname
+				   |> add_hol4_mapping thyname thmname moved_thmname
+			 end
+	val _ = message "new_definition:"
+	val _ = if_debug pth hth
+    in
+	(thy22',hth)
+    end
+    handle e => (message "exception in new_definition"; print_exn e)
+
+val commafy = String.concat o separate ", "
+
+local
+    val helper = thm "termspec_help"
+in
+fun new_specification thyname thmname names hth thy =
+    case HOL4DefThy.get thy of
+	Replaying _ => (thy,hth)
+      | _ => 
+	let
+	    val _ = message "NEW_SPEC:"
+	    val _ = if_debug pth hth
+	    val names = map (rename_const thyname thy) names
+	    val _ = warning ("Introducing constants " ^ (commafy names))
+	    val (HOLThm(rens,th)) = norm_hthm (sign_of thy) hth
+	    val thy1 = case HOL4DefThy.get thy of
+			   Replaying _ => thy
+			 | _ =>
+			   let
+			       fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
+				 | dest_eta_abs body =
+				   let
+				       val (dT,rT) = dom_rng (type_of body)
+				   in
+				       ("x",dT,body $ Bound 0)
+				   end
+				   handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
+			       fun dest_exists (Const("Ex",_) $ abody) =
+				   dest_eta_abs abody
+				 | dest_exists tm =
+				   raise ERR "new_specification" "Bad existential formula"
+					 
+			       val (consts,_) = foldl (fn ((cs,ex),cname) =>
+							  let
+							      val (_,cT,p) = dest_exists ex
+							  in
+							      ((cname,cT,mk_syn cname)::cs,p)
+							  end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
+			       val sg = sign_of thy
+			       val str = foldl (fn (acc,(c,T,csyn)) =>
+						   acc ^ "\n  " ^ (quotename c) ^ " :: \"" ^ (string_of_ctyp (ctyp_of sg T)) ^ "\" " ^ (Syntax.string_of_mixfix csyn)) ("consts",consts)
+			       val thy' = add_dump str thy
+			   in
+			       Theory.add_consts_i consts thy'
+			   end
+
+	    val thy1 = foldr (fn(name,thy)=>
+				snd (get_defname thyname name thy)) (names,thy1)
+	    fun new_name name = fst (get_defname thyname name thy1)
+	    val (thy',res) = SpecificationPackage.add_specification_i None
+				 (map (fn name => (new_name name,name,false)) names)
+				 (thy1,th)
+	    val res' = Drule.freeze_all res
+	    val hth = HOLThm(rens,res')
+	    val rew = rewrite_hol4_term (concl_of res') thy'
+	    val th  = equal_elim rew res'
+	    fun handle_const (name,thy) =
+		let
+		    val defname = def_name name
+		    val (newname,thy') = get_defname thyname name thy
+		in
+		    (if defname = newname
+		     then quotename name
+		     else (quotename newname) ^ ": " ^ (quotename name),thy')
+		end
+	    val (new_names,thy') = foldr (fn(name,(names,thy)) =>
+					    let
+						val (name',thy') = handle_const (name,thy)
+					    in
+						(name'::names,thy')
+					    end) (names,([],thy'))
+	    val thy'' = add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^
+				  "\n  by (import " ^ thyname ^ " " ^ thmname ^ ")")
+				 thy'
+	    val _ = message "RESULT:"
+	    val _ = if_debug pth hth
+	in
+	    intern_store_thm false thyname thmname hth thy''
+	end
+	handle e => (message "exception in new_specification"; print_exn e)
+		    
+end
+			   
+fun new_axiom name tm thy = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
+				      
+fun to_isa_thm (hth as HOLThm(_,th)) =
+    let
+	val (HOLThm args) = norm_hthm (sign_of_thm th) hth
+    in
+	apsnd strip_shyps args
+    end
+
+fun to_isa_term tm = tm
+
+local
+    val light_nonempty = thm "light_ex_imp_nonempty"
+    val ex_imp_nonempty = thm "ex_imp_nonempty"
+    val typedef_hol2hol4 = thm "typedef_hol2hol4"
+    val typedef_hol2hollight = thm "typedef_hol2hollight"
+in
+fun new_type_definition thyname thmname tycname hth thy =
+    case HOL4DefThy.get thy of
+	Replaying _ => (thy,hth)
+      | _ => 
+	let
+	    val _ = message "TYPE_DEF:"
+	    val _ = if_debug pth hth
+	    val _ = warning ("Introducing type " ^ tycname)
+	    val (HOLThm(rens,td_th)) = norm_hthm (sign_of thy) hth
+	    val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
+	    val c = case concl_of th2 of
+			_ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
+		      | _ => raise ERR "new_type_definition" "Bad type definition theorem"
+	    val tfrees = term_tfrees c
+	    val tnames = map fst tfrees
+	    val tsyn = mk_syn tycname
+	    val typ = (tycname,tnames,tsyn)
+	    val (thy',typedef_info) = TypedefPackage.add_typedef_i false (Some thmname) typ c None (rtac th2 1) thy
+				      
+	    val th3 = (#type_definition typedef_info) RS typedef_hol2hol4
+
+	    val fulltyname = Sign.intern_tycon (sign_of thy') tycname
+	    val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
+
+	    val sg = sign_of thy''
+	    val (hth' as HOLThm args) = norm_hthm sg (HOLThm(rens,th3))
+	    val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
+		    else ()
+	    val thy4 = add_hol4_pending thyname thmname args thy''
+
+	    val sg = sign_of thy4
+	    val rew = rewrite_hol4_term (concl_of td_th) thy4
+	    val th  = equal_elim rew (transfer_sg sg td_th)
+	    val c   = case HOLogic.dest_Trueprop (prop_of th) of
+			  Const("Ex",exT) $ P =>
+			  let
+			      val PT = domain_type exT
+			  in
+			      Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P
+			  end
+			| _ => error "Internal error in ProofKernel.new_typedefinition"
+	    val tnames_string = if null tnames
+				then ""
+				else "(" ^ (commafy tnames) ^ ") "
+	    val proc_prop = if null tnames
+			    then smart_string_of_cterm
+			    else Library.setmp show_all_types true smart_string_of_cterm
+	    val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of sg c)) ^ " " ^ (Syntax.string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
+	    val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hol4 [OF type_definition_" ^ tycname ^ "]")
+		       thy5
+	    val _ = message "RESULT:"
+	    val _ = if_debug pth hth'
+	in
+	    (thy6,hth')
+	end
+	handle e => (message "exception in new_type_definition"; print_exn e)
+
+fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
+    case HOL4DefThy.get thy of
+	Replaying _ => (thy,hth)
+      | _ => 
+	let
+	    val _ = message "TYPE_INTRO:"
+	    val _ = if_debug pth hth
+	    val _ = warning ("Introducing type " ^ tycname ^ " (with morphisms " ^ abs_name ^ " and " ^ rep_name ^ ")")
+	    val (HOLThm(rens,td_th)) = norm_hthm (sign_of thy) hth
+	    val sg = sign_of thy
+	    val tT = type_of t
+	    val light_nonempty' =
+		Drule.instantiate' [Some (ctyp_of sg tT)]
+				   [Some (cterm_of sg P),
+				    Some (cterm_of sg t)] light_nonempty
+	    val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
+	    val c = case concl_of th2 of
+			_ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
+		      | _ => raise ERR "type_introduction" "Bad type definition theorem"
+	    val tfrees = term_tfrees c
+	    val tnames = map fst tfrees
+	    val tsyn = mk_syn tycname
+	    val typ = (tycname,tnames,tsyn)
+	    val (thy',typedef_info) = TypedefPackage.add_typedef_i false (Some thmname) typ c (Some(rep_name,abs_name)) (rtac th2 1) thy
+				      
+	    val th3 = (#type_definition typedef_info) RS typedef_hol2hollight
+
+	    val th4 = Drule.freeze_all th3
+	    val fulltyname = Sign.intern_tycon (sign_of thy') tycname
+	    val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
+
+	    val sg = sign_of thy''
+	    val (hth' as HOLThm args) = norm_hthm sg (HOLThm(rens,th4))
+	    val _ = if #maxidx (rep_thm th4) <> ~1
+		    then (Library.setmp show_types true pth hth' ; error "SCHEME!")
+		    else ()
+	    val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
+		    else ()
+	    val thy4 = add_hol4_pending thyname thmname args thy''
+
+	    val sg = sign_of thy4
+	    val P' = #2 (Logic.dest_equals (concl_of (rewrite_hol4_term P thy4)))
+	    val c   =
+		let
+		    val PT = type_of P'
+		in
+		    Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P'
+		end
+
+	    val tnames_string = if null tnames
+				then ""
+				else "(" ^ (commafy tnames) ^ ") "
+	    val proc_prop = if null tnames
+			    then smart_string_of_cterm
+			    else Library.setmp show_all_types true smart_string_of_cterm
+	    val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of sg c)) ^ " " ^ (Syntax.string_of_mixfix tsyn) ^ "\n  by (rule light_ex_imp_nonempty,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
+	    val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hollight [OF type_definition_" ^ tycname ^ "]")
+		       thy5
+	    val _ = message "RESULT:"
+	    val _ = if_debug pth hth'
+	in
+	    (thy6,hth')
+	end
+	handle e => (message "exception in type_introduction"; print_exn e)
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/replay.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,333 @@
+structure Replay =
+struct
+
+structure P = ProofKernel
+
+open ProofKernel
+
+exception REPLAY of string * string
+fun ERR f mesg = REPLAY (f,mesg)
+fun NY f = raise ERR f "NOT YET!"
+
+fun replay_proof int_thms thyname thmname prf thy =
+    let
+	fun rp (PRefl tm) thy = P.REFL tm thy
+	  | rp (PInstT(p,lambda)) thy =
+	    let
+		val (thy',th) = rp' p thy
+	    in
+		P.INST_TYPE lambda th thy'
+	    end
+	  | rp (PSubst(prfs,ctxt,prf)) thy =
+	    let
+		val (thy',ths) = foldr (fn (p,(thy,ths)) =>
+					   let
+					       val (thy',th) = rp' p thy
+					   in
+					       (thy',th::ths)
+					   end) (prfs,(thy,[]))
+		val (thy'',th) = rp' prf thy'
+	    in
+		P.SUBST ths ctxt th thy''
+	    end
+	  | rp (PAbs(prf,v)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.ABS v th thy'
+	    end
+	  | rp (PDisch(prf,tm)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.DISCH tm th thy'
+	    end
+	  | rp (PMp(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.MP th1 th2 thy2
+	    end
+	  | rp (PHyp tm) thy = P.ASSUME tm thy
+	  | rp (PDef(seg,name,rhs)) thy =
+	    (case P.get_def seg name rhs thy of
+		 (thy',Some res) => (thy',res)
+	       | (thy',None) => 
+		 if seg = thyname
+		 then P.new_definition seg name rhs thy'
+		 else raise ERR "replay_proof" "Too late for term definition")
+	  | rp (POracle(tg,asl,c)) thy = (*P.mk_oracle_thm tg (map be_contract asl,c)*) NY "ORACLE"
+	  | rp (PSpec(prf,tm)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.SPEC tm th thy'
+	    end
+	  | rp (PInst(prf,theta)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.INST theta th thy'
+	    end
+	  | rp (PGen(prf,v)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.GEN v th thy'
+	    end
+	  | rp (PGenAbs(prf,opt,vl)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.GEN_ABS opt vl th thy'
+	    end
+	  | rp (PImpAS(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.IMP_ANTISYM th1 th2 thy2
+	    end
+	  | rp (PSym prf) thy =
+	    let
+		val (thy1,th) = rp' prf thy
+	    in
+		P.SYM th thy1
+	    end
+	  | rp (PTrans(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.TRANS th1 th2 thy2
+	    end
+	  | rp (PComb(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.COMB th1 th2 thy2
+	    end
+	  | rp (PEqMp(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.EQ_MP th1 th2 thy2
+	    end
+	  | rp (PEqImp prf) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.EQ_IMP_RULE th thy'
+	    end
+	  | rp (PExists(prf,ex,wit)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.EXISTS ex wit th thy'
+	    end
+	  | rp (PChoose(v,prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.CHOOSE v th1 th2 thy2
+	    end
+	  | rp (PConj(prf1,prf2)) thy =
+	    let
+		val (thy1,th1) = rp' prf1 thy
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.CONJ th1 th2 thy2
+	    end
+	  | rp (PConjunct1 prf) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.CONJUNCT1 th thy'
+	    end
+	  | rp (PConjunct2 prf) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.CONJUNCT2 th thy'
+	    end
+	  | rp (PDisj1(prf,tm)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.DISJ1 th tm thy'
+	    end
+	  | rp (PDisj2(prf,tm)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.DISJ2 tm th thy'
+	    end
+	  | rp (PDisjCases(prf,prf1,prf2)) thy =
+	    let
+		val (thy',th)  = rp' prf  thy
+		val (thy1,th1) = rp' prf1 thy'
+		val (thy2,th2) = rp' prf2 thy1
+	    in
+		P.DISJ_CASES th th1 th2 thy2
+	    end
+	  | rp (PNotI prf) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.NOT_INTRO th thy'
+	    end
+	  | rp (PNotE prf) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.NOT_ELIM th thy'
+	    end
+	  | rp (PContr(prf,tm)) thy =
+	    let
+		val (thy',th) = rp' prf thy
+	    in
+		P.CCONTR tm th thy'
+	    end
+	  | rp (PTmSpec _) _ = raise ERR "rp" "Shouldn't reach here (PTmSpec)"
+	  | rp (PTyDef _) _ = raise ERR "rp" "Shouldn't reach here (PTyDef)"
+	  | rp (PTyIntro _) _ = raise ERR "rp" "Shouldn't reach here (PTyIntro)"
+	  | rp PDisk _ = raise ERR "rp" "Shouldn't reach here (PDisk)"
+	  | rp _ _ = raise ERR "rp" "What the hell is this? Which case did I forget?"
+	and rp' p thy =
+	    let
+		val pc = content_of p
+	    in
+		case pc of
+		    PDisk => (case disk_info_of p of
+				  Some(thyname',thmname) =>
+				  (case Int.fromString thmname of
+				       SOME i =>
+				       if thyname' = thyname
+				       then
+					   (case Array.sub(int_thms,i-1) of
+						None =>
+						let
+						    val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy
+						    val _ = Array.update(int_thms,i-1,Some th)
+						in
+						    (thy',th)
+						end
+					      | Some th => (thy,th))
+				       else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
+				     | NONE => 
+				       (case P.get_thm thyname' thmname thy of
+					    (thy',Some res) => (thy',res)
+					  | (thy',None) => 
+					    if thyname' = thyname
+					    then
+						let
+						    val _ = writeln ("Found no " ^ thmname ^ " theorem, replaying...")
+						    val (f_opt,prf) = import_proof thyname' thmname thy'
+						    val prf = prf thy'
+						    val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
+						in
+						    case content_of prf of
+							PTmSpec _ => (thy',th)
+						      | PTyDef  _ => (thy',th)
+						      | PTyIntro _ => (thy',th)
+						      | _ => P.store_thm thyname' thmname th thy'
+						end
+					    else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
+				| None => raise ERR "rp'.PDisk" "Not enough information")
+		  | PAxm(name,c) =>
+		    (case P.get_axiom thyname name thy of
+			    (thy',Some res) => (thy',res)
+			  | (thy',None) => P.new_axiom name c thy')
+		  | PTmSpec(seg,names,prf') =>
+		    let
+			val (thy',th) = rp' prf' thy
+		    in
+			P.new_specification seg thmname names th thy'
+		    end
+		  | PTyDef(seg,name,prf') =>
+		    let
+			val (thy',th) = rp' prf' thy
+		    in
+			P.new_type_definition seg thmname name th thy'
+		    end
+		  | PTyIntro(seg,name,abs_name,rep_name,P,t,prf') =>
+		    let
+			val (thy',th) = rp' prf' thy
+		    in
+			P.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
+		    end
+		  | _ => rp pc thy
+	    end
+    in
+	rp' prf thy handle e => (writeln "Exception in replay_proof"; print_exn e)
+    end
+
+fun setup_int_thms thyname thy =
+    let
+	val is = TextIO.openIn(P.get_proof_dir thyname thy ^"/facts.lst")
+	val (num_int_thms,facts) =
+	    let
+		fun get_facts facts =
+		    case TextIO.inputLine is of
+			"" => (case facts of
+				   i::facts => (valOf (Int.fromString i),rev facts)
+				 | _ => raise ERR "replay_thm" "Bad facts.lst file")
+		      | fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
+	    in
+		get_facts []
+	    end
+	val _ = TextIO.closeIn is
+	val int_thms = Array.array(num_int_thms,None:thm option)
+    in
+	(int_thms,facts)
+    end
+
+fun import_single_thm thyname int_thms thmname thy =
+    let
+	fun replay_fact (thmname,thy) =
+	    let
+		val _ = writeln ("Replaying " ^ thmname)
+		val prf = mk_proof PDisk
+		val _ = set_disk_info_of prf thyname thmname
+	    in
+		fst (replay_proof int_thms thyname thmname prf thy)
+	    end
+    in
+	replay_fact (thmname,thy)
+    end
+
+fun import_thms thyname int_thms thmnames thy =
+    let
+	fun replay_fact (thy,thmname) =
+	    let
+		val _ = writeln ("Replaying " ^ thmname)
+		val prf = mk_proof PDisk
+		val _ = set_disk_info_of prf thyname thmname
+	    in
+		fst (replay_proof int_thms thyname thmname prf thy)
+	    end
+	val res_thy = foldl replay_fact (thy,thmnames)
+    in
+	res_thy
+    end
+
+fun import_thm thyname thmname thy =
+    let
+	val int_thms = fst (setup_int_thms thyname thy)
+	fun replay_fact (thmname,thy) =
+	    let
+		val _ = writeln ("Replaying " ^ thmname)
+		val prf = mk_proof PDisk
+		val _ = set_disk_info_of prf thyname thmname
+	    in
+		fst (replay_proof int_thms thyname thmname prf thy)
+	    end
+    in
+	replay_fact (thmname,thy)
+    end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/shuffler.ML	Fri Apr 02 17:37:45 2004 +0200
@@ -0,0 +1,685 @@
+(*  Title:      Provers/shuffler.ML
+    ID:         $Id$
+    Author:     Sebastian Skalberg, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+Package for proving two terms equal by normalizing (hence the
+"shuffler" name).  Uses the simplifier for the normalization.
+*)
+
+signature Shuffler =
+sig
+    val debug      : bool ref
+
+    val norm_term  : theory -> term -> thm
+    val make_equal : theory -> term -> term -> thm option
+    val set_prop   : theory -> term -> (string * thm) list -> (string * thm) option
+
+    val find_potential: theory -> term -> (string * thm) list
+
+    val gen_shuffle_tac: theory -> bool -> (string * thm) list -> int -> tactic
+
+    val shuffle_tac: (string * thm) list -> int -> tactic
+    val search_tac : (string * thm) list -> int -> tactic
+
+    val print_shuffles: theory -> unit
+
+    val add_shuffle_rule: thm -> theory -> theory
+    val shuffle_attr: theory attribute
+
+    val setup      : (theory -> theory) list
+end
+
+structure Shuffler :> Shuffler =
+struct
+
+val debug = ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+(*Prints exceptions readably to users*)
+fun print_sign_exn_unit sign e = 
+  case e of
+     THM (msg,i,thms) =>
+	 (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
+	  seq print_thm thms)
+   | THEORY (msg,thys) =>
+	 (writeln ("Exception THEORY raised:\n" ^ msg);
+	  seq (Pretty.writeln o Display.pretty_theory) thys)
+   | TERM (msg,ts) =>
+	 (writeln ("Exception TERM raised:\n" ^ msg);
+	  seq (writeln o Sign.string_of_term sign) ts)
+   | TYPE (msg,Ts,ts) =>
+	 (writeln ("Exception TYPE raised:\n" ^ msg);
+	  seq (writeln o Sign.string_of_typ sign) Ts;
+	  seq (writeln o Sign.string_of_term sign) ts)
+   | e => raise e
+
+(*Prints an exception, then fails*)
+fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise e)
+
+val string_of_thm = Library.setmp print_mode [] string_of_thm
+val string_of_cterm = Library.setmp print_mode [] string_of_cterm
+
+val commafy = String.concat o separate ", "
+
+fun mk_meta_eq th =
+    (case concl_of th of
+	 Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection
+       | Const("==",_) $ _ $ _ => th
+       | _ => raise THM("Not an equality",0,[th]))
+    handle _ => raise THM("Couldn't make meta equality",0,[th])
+				   
+fun mk_obj_eq th =
+    (case concl_of th of
+	 Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th
+       | Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq
+       | _ => raise THM("Not an equality",0,[th]))
+    handle _ => raise THM("Couldn't make object equality",0,[th])
+
+structure ShuffleDataArgs: THEORY_DATA_ARGS =
+struct
+val name = "HOL/shuffles"
+type T = thm list
+val empty = []
+val copy = I
+val prep_ext = I
+val merge = Library.gen_union Thm.eq_thm
+fun print sg thms =
+    Pretty.writeln (Pretty.big_list "Shuffle theorems:"
+				    (map Display.pretty_thm thms))
+end
+
+structure ShuffleData = TheoryDataFun(ShuffleDataArgs)
+
+val weaken =
+    let
+	val cert = cterm_of (sign_of ProtoPure.thy)
+	val P = Free("P",propT)
+	val Q = Free("Q",propT)
+	val PQ = Logic.mk_implies(P,Q)
+	val PPQ = Logic.mk_implies(P,PQ)
+	val cP = cert P
+	val cQ = cert Q
+	val cPQ = cert PQ
+	val cPPQ = cert PPQ
+	val th1 = assume cPQ |> implies_intr_list [cPQ,cP]
+	val th3 = assume cP
+	val th4 = implies_elim_list (assume cPPQ) [th3,th3]
+				    |> implies_intr_list [cPPQ,cP]
+    in
+	equal_intr th4 th1 |> standard
+    end
+
+val imp_comm =
+    let
+	val cert = cterm_of (sign_of ProtoPure.thy)
+	val P = Free("P",propT)
+	val Q = Free("Q",propT)
+	val R = Free("R",propT)
+	val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R))
+	val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R))
+	val cP = cert P
+	val cQ = cert Q
+	val cPQR = cert PQR
+	val cQPR = cert QPR
+	val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ]
+				    |> implies_intr_list [cPQR,cQ,cP]
+	val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP]
+				    |> implies_intr_list [cQPR,cP,cQ]
+    in
+	equal_intr th1 th2 |> standard
+    end
+
+val def_norm =
+    let
+	val cert = cterm_of (sign_of ProtoPure.thy)
+	val aT = TFree("'a",logicS)
+	val bT = TFree("'b",logicS)
+	val v = Free("v",aT)
+	val P = Free("P",aT-->bT)
+	val Q = Free("Q",aT-->bT)
+	val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0)))
+	val cPQ = cert (Logic.mk_equals(P,Q))
+	val cv = cert v
+	val rew = assume cvPQ
+			 |> forall_elim cv
+			 |> abstract_rule "v" cv
+	val (lhs,rhs) = Logic.dest_equals(concl_of rew)
+	val th1 = transitive (transitive
+				  (eta_conversion (cert lhs) |> symmetric)
+				  rew)
+			     (eta_conversion (cert rhs))
+			     |> implies_intr cvPQ
+	val th2 = combination (assume cPQ) (reflexive cv)
+			      |> forall_intr cv
+			      |> implies_intr cPQ
+    in
+	equal_intr th1 th2 |> standard
+    end
+
+val all_comm =
+    let
+	val cert = cterm_of (sign_of ProtoPure.thy)
+	val xT = TFree("'a",logicS)
+	val yT = TFree("'b",logicS)
+	val P = Free("P",xT-->yT-->propT)
+	val lhs = all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))
+	val rhs = all yT $ (Abs("y",yT,all xT $ (Abs("x",xT,P $ Bound 0 $ Bound 1))))
+	val cl = cert lhs
+	val cr = cert rhs
+	val cx = cert (Free("x",xT))
+	val cy = cert (Free("y",yT))
+	val th1 = assume cr
+			 |> forall_elim_list [cy,cx]
+			 |> forall_intr_list [cx,cy]
+			 |> implies_intr cr
+	val th2 = assume cl
+			 |> forall_elim_list [cx,cy]
+			 |> forall_intr_list [cy,cx]
+			 |> implies_intr cl
+    in
+	equal_intr th1 th2 |> standard
+    end
+
+val equiv_comm =
+    let
+	val cert = cterm_of (sign_of ProtoPure.thy)
+	val T    = TFree("'a",[])
+	val t    = Free("t",T)
+	val u    = Free("u",T)
+	val ctu  = cert (Logic.mk_equals(t,u))
+	val cut  = cert (Logic.mk_equals(u,t))
+	val th1  = assume ctu |> symmetric |> implies_intr ctu
+	val th2  = assume cut |> symmetric |> implies_intr cut
+    in
+	equal_intr th1 th2 |> standard
+    end
+
+(* This simplification procedure rewrites !!x y. P x y
+deterministicly, in order for the normalization function, defined
+below, to handle nested quantifiers robustly *)
+
+local
+
+exception RESULT of int
+
+fun find_bound n (Bound i) = if i = n then raise RESULT 0
+			     else if i = n+1 then raise RESULT 1
+			     else ()
+  | find_bound n (t $ u) = (find_bound n t; find_bound n u)
+  | find_bound n (Abs(_,_,t)) = find_bound (n+1) t
+  | find_bound _ _ = ()
+
+fun swap_bound n (Bound i) = if i = n then Bound (n+1)
+			     else if i = n+1 then Bound n
+			     else Bound i
+  | swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u)
+  | swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t)
+  | swap_bound n t = t
+
+fun rew_th sg (xv as (x,xT)) (yv as (y,yT)) t =
+    let
+	val lhs = list_all ([xv,yv],t)
+	val rhs = list_all ([yv,xv],swap_bound 0 t)
+	val rew = Logic.mk_equals (lhs,rhs)
+	val init = trivial (cterm_of sg rew)
+    in
+	(all_comm RS init handle e => (message "rew_th"; print_exn e))
+    end
+
+fun quant_rewrite sg assumes (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) =
+    let
+	val res = (find_bound 0 body;2) handle RESULT i => i
+    in
+	case res of
+	    0 => Some (rew_th sg (x,xT) (y,yT) body)
+	  | 1 => if string_ord(y,x) = LESS
+		 then
+		     let
+			 val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
+			 val t_th    = reflexive (cterm_of sg t)
+			 val newt_th = reflexive (cterm_of sg newt)
+		     in
+			 Some (transitive t_th newt_th)
+		     end
+		 else None
+	  | _ => error "norm_term (quant_rewrite) internal error"
+     end
+  | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; None)
+
+fun freeze_thaw_term t =
+    let
+	val tvars = term_tvars t
+	val tfree_names = add_term_tfree_names(t,[])
+	val (type_inst,_) =
+	    foldl (fn ((inst,used),(w as (v,_),S)) =>
+		      let
+			  val v' = variant used v
+		      in
+			  ((w,TFree(v',S))::inst,v'::used)
+		      end)
+		  (([],tfree_names),tvars)
+	val t' = subst_TVars type_inst t
+    in
+	(t',map (fn (w,TFree(v,S)) => (v,TVar(w,S))) type_inst)
+    end
+
+fun inst_tfrees sg [] thm = thm
+  | inst_tfrees sg ((name,U)::rest) thm = 
+    let
+	val cU = ctyp_of sg U
+	val tfree_names = add_term_tfree_names (prop_of thm,[])
+	val (thm',rens) = varifyT' (tfree_names \ name) thm
+	val mid = 
+	    case rens of
+		[] => thm'
+	      | [(_,idx)] => instantiate ([(idx,cU)],[]) thm'
+	      | _ => error "Shuffler.inst_tfrees internal error"
+    in
+	inst_tfrees sg rest mid
+    end
+
+fun is_Abs (Abs _) = true
+  | is_Abs _ = false
+
+fun eta_redex (t $ Bound 0) =
+    let
+	fun free n (Bound i) = i = n
+	  | free n (t $ u) = free n t orelse free n u
+	  | free n (Abs(_,_,t)) = free (n+1) t
+	  | free n _ = false
+    in
+	not (free 0 t)
+    end
+  | eta_redex _ = false
+
+fun eta_contract sg assumes origt =
+    let
+	val (typet,Tinst) = freeze_thaw_term origt
+	val (init,thaw) = freeze_thaw (reflexive (cterm_of sg typet))
+	val final = inst_tfrees sg Tinst o thaw
+	val t = #1 (Logic.dest_equals (prop_of init))
+	val _ =
+	    let
+		val lhs = #1 (Logic.dest_equals (prop_of (final init)))
+	    in
+		if not (lhs aconv origt)
+		then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
+		      writeln (string_of_cterm (cterm_of sg origt));
+		      writeln (string_of_cterm (cterm_of sg lhs));
+		      writeln (string_of_cterm (cterm_of sg typet));
+		      writeln (string_of_cterm (cterm_of sg t));
+		      app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of sg T)))) Tinst;
+		      writeln "done")
+		else ()
+	    end
+    in
+	case t of
+	    Const("all",_) $ (Abs(x,xT,Const("==",eqT) $ P $ Q)) =>
+	    ((if eta_redex P andalso eta_redex Q
+	      then
+		  let
+		      val cert = cterm_of sg
+		      val v = Free(variant (add_term_free_names(t,[])) "v",xT)
+		      val cv = cert v
+		      val ct = cert t
+		      val th = (assume ct)
+				   |> forall_elim cv
+				   |> abstract_rule x cv
+		      val ext_th = eta_conversion (cert (Abs(x,xT,P)))
+		      val th' = transitive (symmetric ext_th) th
+		      val cu = cert (prop_of th')
+		      val uth = combination (assume cu) (reflexive cv)
+		      val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v)))
+				     |> transitive uth
+				     |> forall_intr cv
+				     |> implies_intr cu
+		      val rew_th = equal_intr (th' |> implies_intr ct) uth'
+		      val res = final rew_th
+		      val lhs = (#1 (Logic.dest_equals (prop_of res)))
+		  in
+		       Some res
+		  end
+	      else None)
+	     handle e => (writeln "eta_contract:";print_exn e))
+	  | _ => (error ("Bad eta_contract argument" ^ (string_of_cterm (cterm_of sg t))); None)
+    end
+
+fun beta_fun sg assume t =
+    Some (beta_conversion true (cterm_of sg t))
+
+fun eta_expand sg assumes origt =
+    let
+	val (typet,Tinst) = freeze_thaw_term origt
+	val (init,thaw) = freeze_thaw (reflexive (cterm_of sg typet))
+	val final = inst_tfrees sg Tinst o thaw
+	val t = #1 (Logic.dest_equals (prop_of init))
+	val _ =
+	    let
+		val lhs = #1 (Logic.dest_equals (prop_of (final init)))
+	    in
+		if not (lhs aconv origt)
+		then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
+		      writeln (string_of_cterm (cterm_of sg origt));
+		      writeln (string_of_cterm (cterm_of sg lhs));
+		      writeln (string_of_cterm (cterm_of sg typet));
+		      writeln (string_of_cterm (cterm_of sg t));
+		      app (fn (n,T) => writeln (n ^ ": " ^ (string_of_ctyp (ctyp_of sg T)))) Tinst;
+		      writeln "done")
+		else ()
+	    end
+    in
+	case t of
+	    Const("==",T) $ P $ Q =>
+	    if is_Abs P orelse is_Abs Q
+	    then (case domain_type T of
+		      Type("fun",[aT,bT]) =>
+		      let
+			  val cert = cterm_of sg
+			  val vname = variant (add_term_free_names(t,[])) "v"
+			  val v = Free(vname,aT)
+			  val cv = cert v
+			  val ct = cert t
+			  val th1 = (combination (assume ct) (reflexive cv))
+					|> forall_intr cv
+					|> implies_intr ct
+			  val concl = cert (concl_of th1)
+			  val th2 = (assume concl)
+					|> forall_elim cv
+					|> abstract_rule vname cv
+			  val (lhs,rhs) = Logic.dest_equals (prop_of th2)
+			  val elhs = eta_conversion (cert lhs)
+			  val erhs = eta_conversion (cert rhs)
+			  val th2' = transitive
+					 (transitive (symmetric elhs) th2)
+					 erhs
+			  val res = equal_intr th1 (th2' |> implies_intr concl)
+			  val res' = final res
+		      in
+			  Some res'
+		      end
+		    | _ => None)
+	    else None
+	  | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of sg t))); None)
+    end
+    handle e => (writeln "eta_expand internal error";print_exn e)
+
+fun mk_tfree s = TFree("'"^s,logicS)
+val xT = mk_tfree "a"
+val yT = mk_tfree "b"
+val P  = Var(("P",0),xT-->yT-->propT)
+val Q  = Var(("Q",0),xT-->yT)
+val R  = Var(("R",0),xT-->yT)
+val S  = Var(("S",0),xT)
+in
+fun beta_simproc sg = Simplifier.simproc_i
+		      sg
+		      "Beta-contraction"
+		      [Abs("x",xT,Q) $ S]
+		      beta_fun
+
+fun quant_simproc sg = Simplifier.simproc_i
+			   sg
+			   "Ordered rewriting of nested quantifiers"
+			   [all xT $ (Abs("x",xT,all yT $ (Abs("y",yT,P $ Bound 1 $ Bound 0))))]
+			   quant_rewrite
+fun eta_expand_simproc sg = Simplifier.simproc_i
+			 sg
+			 "Smart eta-expansion by equivalences"
+			 [Logic.mk_equals(Q,R)]
+			 eta_expand
+fun eta_contract_simproc sg = Simplifier.simproc_i
+			 sg
+			 "Smart handling of eta-contractions"
+			 [all xT $ (Abs("x",xT,Logic.mk_equals(Q $ Bound 0,R $ Bound 0)))]
+			 eta_contract
+end
+
+(* Disambiguates the names of bound variables in a term, returning t
+== t' where all the names of bound variables in t' are unique *)
+
+fun disamb_bound sg t =
+    let
+	
+	fun F (t $ u,idx) =
+	    let
+		val (t',idx') = F (t,idx)
+		val (u',idx'') = F (u,idx')
+	    in
+		(t' $ u',idx'')
+	    end
+	  | F (Abs(x,xT,t),idx) =
+	    let
+		val x' = "x" ^ (LargeInt.toString idx) (* amazing *)
+		val (t',idx') = F (t,idx+1)
+	    in
+		(Abs(x',xT,t'),idx')
+	    end
+	  | F arg = arg
+	val (t',_) = F (t,0)
+	val ct = cterm_of sg t
+	val ct' = cterm_of sg t'
+	val res = transitive (reflexive ct) (reflexive ct')
+	val _ = message ("disamb_term: " ^ (string_of_thm res))
+    in
+	res
+    end
+
+(* Transforms a term t to some normal form t', returning the theorem t
+== t'.  This is originally a help function for make_equal, but might
+be handy in its own right, for example for indexing terms. *)
+
+fun norm_term thy t =
+    let
+	val sg = sign_of thy
+
+	val norms = ShuffleData.get thy
+	val ss = empty_ss setmksimps single
+			  addsimps (map (transfer_sg sg) norms)
+	fun chain f th =
+	    let
+		val rhs = snd (dest_equals (cprop_of th))
+	    in
+		transitive th (f rhs)
+	    end
+
+	val th =
+	    t |> disamb_bound sg
+	      |> chain (Simplifier.full_rewrite
+			    (ss addsimprocs [quant_simproc sg,eta_expand_simproc sg,eta_contract_simproc sg]))
+	      |> chain eta_conversion
+	      |> strip_shyps
+	val _ = message ("norm_term: " ^ (string_of_thm th))
+    in
+	th
+    end
+    handle e => (writeln "norm_term internal error"; print_sign_exn (sign_of thy) e)
+
+fun is_logic_var sg v =
+    Type.of_sort (Sign.tsig_of sg) (type_of v,logicS)
+
+(* Closes a theorem with respect to free and schematic variables (does
+not touch type variables, though). *)
+
+fun close_thm th =
+    let
+	val sg = sign_of_thm th
+	val c = prop_of th
+	val all_vars = add_term_frees (c,add_term_vars(c,[]))
+	val all_rel_vars = filter (is_logic_var sg) all_vars
+    in
+	Drule.forall_intr_list (map (cterm_of sg) all_rel_vars) th
+    end
+    handle e => (writeln "close_thm internal error"; print_exn e)
+
+(* Normalizes a theorem's conclusion using norm_term. *)
+
+fun norm_thm thy th =
+    let
+	val c = prop_of th
+    in
+	equal_elim (norm_term thy c) th
+    end
+
+(* make_equal sg t u tries to construct the theorem t == u under the
+signature sg.  If it succeeds, Some (t == u) is returned, otherwise
+None is returned. *)
+
+fun make_equal sg t u =
+    let
+	val t_is_t' = norm_term sg t
+	val u_is_u' = norm_term sg u
+	val th = transitive t_is_t' (symmetric u_is_u')
+	val _ = message ("make_equal: Some " ^ (string_of_thm th))
+    in
+	Some th
+    end
+    handle e as THM _ => (message "make_equal: None";None)
+			 
+fun match_consts ignore t (* th *) =
+    let
+	fun add_consts (Const (c, _), cs) =
+	    if c mem_string ignore
+	    then cs
+	    else c ins_string cs
+	  | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
+	  | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
+	  | add_consts (_, cs) = cs
+	val t_consts = add_consts(t,[])
+    in
+     fn (name,th) =>
+	let
+	    val th_consts = add_consts(prop_of th,[])
+	in
+	    eq_set(t_consts,th_consts)
+	end
+    end
+    
+val collect_ignored =
+    foldr (fn (thm,cs) =>
+	      let
+		  val (lhs,rhs) = Logic.dest_equals (prop_of thm)
+		  val ignore_lhs = term_consts lhs \\ term_consts rhs
+		  val ignore_rhs = term_consts rhs \\ term_consts lhs
+	      in
+		  foldr (op ins_string) (ignore_lhs @ ignore_rhs,cs)
+	      end)
+
+(* set_prop t thms tries to make a theorem with the proposition t from
+one of the theorems thms, by shuffling the propositions around.  If it
+succeeds, Some theorem is returned, otherwise None.  *)
+
+fun set_prop thy t =
+    let
+	val sg = sign_of thy
+	val all_vars = add_term_frees (t,add_term_vars (t,[]))
+	val all_rel_vars = filter (is_logic_var sg) all_vars
+	val closed_t = foldr (fn (v,body) => let val vT = type_of v
+					     in all vT $ (Abs("x",vT,abstract_over(v,body))) end) (all_rel_vars,t)
+	val rew_th = norm_term thy closed_t
+	val rhs = snd (dest_equals (cprop_of rew_th))
+
+	val shuffles = ShuffleData.get thy
+	fun process [] = None
+	  | process ((name,th)::thms) =
+	    let
+		val norm_th = varifyT (norm_thm thy (close_thm (transfer_sg sg th)))
+		val triv_th = trivial rhs
+		val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
+		val mod_th = case Seq.pull (bicompose true (false,norm_th,0) 1 triv_th) of
+				 Some(th,_) => Some th
+			       | None => None
+	    in
+		case mod_th of
+		    Some mod_th =>
+		    let
+			val closed_th = equal_elim (symmetric rew_th) mod_th
+		    in
+			message ("Shuffler.set_prop succeeded by " ^ name);
+			Some (name,forall_elim_list (map (cterm_of sg) all_rel_vars) closed_th)
+		    end
+		  | None => process thms
+	    end
+	    handle e as THM _ => process thms
+    in
+	fn thms =>
+	   case process thms of
+	       res as Some (name,th) => if (prop_of th) aconv t
+					then res
+					else error "Internal error in set_prop"
+	     | None => None
+    end
+    handle e => (writeln "set_prop internal error"; print_exn e)
+
+fun find_potential thy t =
+    let
+	val shuffles = ShuffleData.get thy
+	val ignored = collect_ignored(shuffles,[])
+	val rel_consts = term_consts t \\ ignored
+	val pot_thms = PureThy.thms_containing_consts thy rel_consts
+    in
+	filter (match_consts ignored t) pot_thms
+    end
+
+fun gen_shuffle_tac thy search thms i st =
+    let
+	val _ = message ("Shuffling " ^ (string_of_thm st))
+	val t = nth_elem(i-1,prems_of st)
+	val set = set_prop thy t
+	fun process_tac thms st =
+	    case set thms of
+		Some (_,th) => Seq.of_list (compose (th,i,st))
+	      | None => Seq.empty
+    in
+	(process_tac thms APPEND (if search
+				  then process_tac (find_potential thy t)
+				  else no_tac)) st
+    end
+
+fun shuffle_tac thms i st =
+    gen_shuffle_tac (the_context()) false thms i st
+
+fun search_tac thms i st =
+    gen_shuffle_tac (the_context()) true thms i st
+
+fun shuffle_meth (thms:thm list) ctxt =
+    let
+	val thy = ProofContext.theory_of ctxt
+    in
+	Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy false (map (pair "") thms))
+    end
+
+fun search_meth ctxt =
+    let
+	val thy = ProofContext.theory_of ctxt
+	val prems = ProofContext.prems_of ctxt
+    in
+	Method.SIMPLE_METHOD' HEADGOAL (gen_shuffle_tac thy true (map (pair "premise") prems))
+    end
+
+val print_shuffles = ShuffleData.print
+
+fun add_shuffle_rule thm thy =
+    let
+	val shuffles = ShuffleData.get thy
+    in
+	if exists (curry Thm.eq_thm thm) shuffles
+	then (warning ((string_of_thm thm) ^ " already known to the shuffler");
+	      thy)
+	else ShuffleData.put (thm::shuffles) thy
+    end
+
+fun shuffle_attr (thy,thm) = (add_shuffle_rule thm thy,thm)
+
+val setup = [Method.add_method ("shuffle_tac",Method.thms_ctxt_args shuffle_meth,"solve goal by shuffling terms around"),
+	     Method.add_method ("search_tac",Method.ctxt_args search_meth,"search for suitable theorems"),
+	     ShuffleData.init,
+	     add_shuffle_rule weaken,
+	     add_shuffle_rule equiv_comm,
+	     add_shuffle_rule imp_comm,
+	     add_shuffle_rule Drule.norm_hhf_eq,
+	     add_shuffle_rule Drule.triv_forall_equality,
+	     Attrib.add_attributes [("shuffle_rule",(Attrib.no_args shuffle_attr,K Attrib.undef_local_attribute),"tell the shuffler about the theorem")]]
+end
--- a/src/HOL/IsaMakefile	Fri Apr 02 17:36:01 2004 +0200
+++ b/src/HOL/IsaMakefile	Fri Apr 02 17:37:45 2004 +0200
@@ -19,6 +19,7 @@
   HOL-CTL \
   HOL-Extraction \
       HOL-Complex-HahnBanach \
+      HOL-Complex-Import \
   HOL-Hoare \
   HOL-HoareParallel \
   HOL-IMP \
@@ -236,6 +237,51 @@
 	@$(ISATOOL) usedir $(OUT)/HOL IMPP
 
 
+## HOL-Complex-Import
+
+IMPORTER_FILES = Import/proof_kernel.ML Import/replay.ML \
+  Import/shuffler.ML Import/MakeEqual.thy Import/HOL4Setup.thy \
+  Import/HOL4Syntax.thy Import/HOL4Compat.thy Import/import_syntax.ML \
+  Import/hol4rews.ML Import/import_package.ML Import/ROOT.ML
+
+HOL-Complex-Import: HOL-Complex $(LOG)/HOL-Complex-Import.gz
+
+$(LOG)/HOL-Complex-Import.gz: $(OUT)/HOL-Complex $(IMPORTER_FILES)
+	@$(ISATOOL) usedir $(OUT)/HOL-Complex Import
+
+
+## HOL-Complex-Generate-HOL
+
+HOL-Complex-Generate-HOL: HOL-Complex $(LOG)/HOL-Complex-Generate-HOL.gz
+
+$(LOG)/HOL-Complex-Generate-HOL.gz: $(OUT)/HOL-Complex $(IMPORTER_FILES)  \
+  Import/Generate-HOL/GenHOL4Base.thy Import/Generate-HOL/GenHOL4Prob.thy \
+  Import/Generate-HOL/GenHOL4Real.thy Import/Generate-HOL/GenHOL4Vec.thy  \
+  Import/Generate-HOL/GenHOL4Word32.thy Import/Generate-HOL/ROOT.ML
+	@cd Import; $(ISATOOL) usedir $(OUT)/HOL-Complex Generate-HOL
+
+
+## HOL-Import-HOL
+
+HOL-Import-HOL: HOL-Complex $(LOG)/HOL-Import-HOL.gz
+
+HOL_IMPORT_FILES = arithmetic.imp bits.imp boolean_sequence.imp bool.imp \
+  bword_arith.imp bword_bitop.imp bword_num.imp combin.imp divides.imp \
+  hrat.imp hreal.imp ind_type.imp lim.imp list.imp marker.imp nets.imp \
+  numeral.imp num.imp one.imp operator.imp option.imp pair.imp poly.imp \
+  powser.imp pred_set.imp prime.imp prim_rec.imp prob_algebra.imp \
+  prob_canon.imp prob_extra.imp prob.imp prob_indep.imp prob_pseudo.imp \
+  prob_uniform.imp realax.imp real.imp relation.imp res_quan.imp rich_list.imp \
+  seq.imp state_transformer.imp sum.imp topology.imp transc.imp word32.imp \
+  word_base.imp word_bitop.imp word_num.imp
+
+$(LOG)/HOL-Import-HOL.gz: $(OUT)/HOL-Complex $(IMPORTER_FILES) \
+  $(HOL_IMPORT_FILES:%=Import/HOL/%) Import/HOL/HOL4Base.thy \
+  Import/HOL/HOL4Prob.thy Import/HOL/HOL4Real.thy Import/HOL/HOL4Vec.thy \
+  Import/HOL/HOL4Word32.thy Import/HOL/HOL4.thy Import/HOL/ROOT.ML
+	@cd Import/HOL; $(ISATOOL) usedir -b $(OUT)/HOL-Complex HOL-Import-HOL
+
+
 ## HOL-NumberTheory
 
 HOL-NumberTheory: HOL $(LOG)/HOL-NumberTheory.gz