Modernized HOL-Import for HOL Light
authorCezary Kaliszyk <cezarykaliszyk@gmail.com>
Sun, 01 Apr 2012 14:50:47 +0200
changeset 47258 880e587eee9f
parent 47244 a7f85074c169
child 47259 2d4ea84278da
Modernized HOL-Import for HOL Light
src/HOL/Import/HOL4/Importer.thy
src/HOL/Import/HOL4/README
src/HOL/Import/HOL4/import.ML
src/HOL/Import/HOL4/import_rews.ML
src/HOL/Import/HOL4/proof_kernel.ML
src/HOL/Import/HOL4/replay.ML
src/HOL/Import/HOL4/shuffler.ML
src/HOL/Import/HOL_Light/Compatibility.thy
src/HOL/Import/HOL_Light/Generate.thy
src/HOL/Import/HOL_Light/Generated/HOLLight.thy
src/HOL/Import/HOL_Light/Generated/hollight.imp
src/HOL/Import/HOL_Light/HOLLightInt.thy
src/HOL/Import/HOL_Light/HOLLightList.thy
src/HOL/Import/HOL_Light/HOLLightReal.thy
src/HOL/Import/HOL_Light/Imported.thy
src/HOL/Import/HOL_Light/ROOT.ML
src/HOL/Import/HOL_Light/Template/GenHOLLight.thy
src/HOL/Import/HOL_Light/generate.ML
src/HOL/Import/HOL_Light/imported.ML
src/HOL/Import/HOL_Light_Import.thy
src/HOL/Import/HOL_Light_Maps.thy
src/HOL/Import/Import_Setup.thy
src/HOL/Import/Importer.thy
src/HOL/Import/README
src/HOL/Import/ROOT.ML
src/HOL/Import/import.ML
src/HOL/Import/import_data.ML
src/HOL/Import/import_rews.ML
src/HOL/Import/import_rule.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/HOL4/Importer.thy	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,236 @@
+(*  Title:      HOL/Import/Importer.thy
+    Author:     Sebastian Skalberg, TU Muenchen
+*)
+
+theory Importer
+imports Main
+keywords
+  "import_segment" "import_theory" "end_import" "setup_theory" "end_setup"
+  "setup_dump" "append_dump" "flush_dump" "ignore_thms" "type_maps"
+  "def_maps" "thm_maps" "const_renames" "const_moves" "const_maps" :: thy_decl and ">"
+uses "shuffler.ML" "import_rews.ML" ("proof_kernel.ML") ("replay.ML") ("import.ML")
+begin
+
+setup {* Shuffler.setup #> importer_setup *}
+
+parse_ast_translation smarter_trueprop_parsing
+
+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
+
+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"
+
+defs
+  ONE_ONE_DEF: "ONE_ONE f == ALL x y. f x = f y --> x = y"
+
+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 & ~(surj f))"
+proof (rule exI,safe)
+  show "inj Suc_Rep"
+    by (rule injI) (rule Suc_Rep_inject)
+next
+  assume "surj Suc_Rep"
+  hence "ALL y. EX x. y = Suc_Rep x"
+    by (simp add: 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 b: "Rep x = Rep y"
+    from td have "x = Abs (Rep x)"
+      by auto
+    also have "Abs (Rep x) = Abs (Rep y)"
+      by (simp add: b)
+    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 "proof_kernel.ML"
+use "replay.ML"
+use "import.ML"
+
+setup Import.setup
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/README	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,19 @@
+
+ATTENTION!  This is largely outdated.  The hint to PROOF_DIRS might be
+everything which is still relevant.
+
+All the files in this directory (except this README, Importer.thy, and
+ROOT.ML) are automatically generated.  Edit the files in
+../Generate-HOL and run "isabelle make HOL-Complex-Generate-HOL" in
+~~/src/HOL, if something needs to be changed.
+
+To build the logic in this directory, simply do a "isabelle make
+HOL-Import-HOL" in ~~/src/HOL.
+
+Note that the quick_and_dirty flag is on as default for this
+directory, which means that the original external proofs are not consulted
+at all.  If a real replay of the external proofs is desired, get and
+unpack the external proof objects to somewhere on your harddisk, and set
+the variable PROOF_DIRS to the directory where the directory "hol4"
+is.  Now edit the ROOT.ML file to unset the quick_and_dirty flag and
+do "isabelle make HOL-Import-HOL" in ~~/src/HOL.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/import.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,280 @@
+(*  Title:      HOL/Import/import.ML
+    Author:     Sebastian Skalberg (TU Muenchen)
+*)
+
+signature IMPORT =
+sig
+    val debug      : bool Unsynchronized.ref
+    val import_tac : Proof.context -> string * string -> tactic
+    val setup      : theory -> theory
+end
+
+structure ImportData = Theory_Data
+(
+  type T = ProofKernel.thm option array option  (* FIXME mutable array !??!!! *)
+  val empty = NONE
+  val extend = I
+  fun merge _ = NONE
+)
+
+structure Import : IMPORT =
+struct
+
+val debug = Unsynchronized.ref false
+fun message s = if !debug then writeln s else ()
+
+fun import_tac ctxt (thyname, thmname) =
+    if ! quick_and_dirty
+    then Skip_Proof.cheat_tac (Proof_Context.theory_of ctxt)
+    else
+     fn th =>
+        let
+            val thy = Proof_Context.theory_of ctxt
+            val prem = hd (prems_of th)
+            val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
+            val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
+            val int_thms = case ImportData.get thy of
+                               NONE => fst (Replay.setup_int_thms thyname thy)
+                             | SOME a => a
+            val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
+            val importerthm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
+            val thm = snd (ProofKernel.to_isa_thm importerthm)
+            val rew = ProofKernel.rewrite_importer_term (concl_of thm) thy
+            val thm = Thm.equal_elim rew thm
+            val prew = ProofKernel.rewrite_importer_term prem thy
+            val prem' = #2 (Logic.dest_equals (prop_of prew))
+            val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm)
+            val thm = ProofKernel.disambiguate_frees thm
+            val _ = message ("Disambiguate: " ^ Display.string_of_thm ctxt 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' = Thm.equal_elim (Thm.symmetric prew) thm
+                    val res = Thm.bicompose true (false,thm',0) 1 th
+                in
+                    res
+                end
+              | NONE => (message "import: set_prop didn't succeed"; no_tac th)
+        end
+
+val setup = Method.setup @{binding import}
+  (Scan.lift (Args.name -- Args.name) >>
+    (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
+  "import external theorem"
+
+(* Parsers *)
+
+val import_segment = Parse.name >> set_import_segment
+
+
+val import_theory = (Parse.name -- Parse.name) >> (fn (location, thyname) =>
+                               fn thy =>
+                                  thy |> set_generating_thy thyname
+                                      |> Sign.add_path thyname
+                                      |> add_dump ("setup_theory " ^ quote location ^ " " ^ thyname))
+
+fun end_import toks =
+    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_out (should_ignore thyname thy) facts
+                fun replay thy = Replay.import_thms thyname int_thms thmnames thy
+            in
+                thy |> clear_import_thy
+                    |> set_segment thyname segname
+                    |> set_used_names facts
+                    |> replay 
+                    |> clear_used_names
+                    |> export_importer_pending
+                    |> Sign.parent_path
+                    |> dump_import_thy thyname
+                    |> add_dump ";end_setup"
+            end) toks
+
+val ignore_thms = Scan.repeat1 Parse.name
+                               >> (fn ignored =>
+                                   fn thy =>
+                                      let
+                                          val thyname = get_import_thy thy
+                                      in
+                                          Library.foldl (fn (thy,thmname) => ignore_importer thyname thmname thy) (thy,ignored)
+                                      end)
+
+val thm_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+                            >> (fn thmmaps =>
+                                fn thy =>
+                                   let
+                                       val thyname = get_import_thy thy
+                                   in
+                                       Library.foldl (fn (thy,(hol,isa)) => add_importer_mapping thyname hol isa thy) (thy,thmmaps)
+                                   end)
+
+val type_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+                             >> (fn typmaps =>
+                                 fn thy =>
+                                    let
+                                        val thyname = get_import_thy thy
+                                    in
+                                        Library.foldl (fn (thy,(hol,isa)) => add_importer_type_mapping thyname hol false isa thy) (thy,typmaps)
+                                    end)
+
+val def_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+                            >> (fn defmaps =>
+                                fn thy =>
+                                   let
+                                       val thyname = get_import_thy thy
+                                   in
+                                       Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
+                                   end)
+
+val const_renames = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.name)
+                                 >> (fn renames =>
+                                     fn thy =>
+                                        let
+                                            val thyname = get_import_thy thy
+                                        in
+                                            Library.foldl (fn (thy,(hol,isa)) => add_importer_const_renaming thyname hol isa thy) (thy,renames)
+                                        end)
+                                                                                                                                      
+val const_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
+                              >> (fn constmaps =>
+                                  fn thy =>
+                                     let
+                                         val thyname = get_import_thy thy
+                                     in
+                                         Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol false isa thy
+                                                 | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+                                     end)
+
+val const_moves = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
+                               >> (fn constmaps =>
+                                   fn thy =>
+                                      let
+                                          val thyname = get_import_thy thy
+                                      in
+                                          Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol true isa thy
+                                                  | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol true isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+                                      end)
+
+fun import_thy location s =
+    let
+        val src_location = Path.append (Path.explode location) (Path.basic (s ^ ".imp"))
+        val is = TextIO.openIn (File.platform_path src_location)
+        val inp = TextIO.inputAll is
+        val _ = TextIO.closeIn is
+        val orig_source = Source.of_string inp
+        val symb_source = Symbol.source orig_source
+        val lexes =
+          (Scan.make_lexicon
+            (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
+                  Scan.empty_lexicon)
+        val token_source = Token.source {do_recover = NONE} (K lexes) Position.start symb_source
+        val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
+        val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
+        val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
+        val def_mapsP = Parse.$$$ "def_maps" |-- def_maps
+        val const_mapsP = Parse.$$$ "const_maps" |-- const_maps
+        val const_movesP = Parse.$$$ "const_moves" |-- const_moves
+        val const_renamesP = Parse.$$$ "const_renames" |-- const_renames
+        val ignore_thmsP = Parse.$$$ "ignore_thms" |-- ignore_thms
+        val thm_mapsP = Parse.$$$ "thm_maps" |-- thm_maps
+        val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
+        val importP = Parse.$$$ "import" |-- Scan.repeat parser --| Parse.$$$ "end"
+        fun apply [] thy = thy
+          | apply (f::fs) thy = apply fs (f thy)
+    in
+        apply (set_replaying_thy s :: Sign.add_path s :: fst (importP token_list))
+    end
+
+fun create_int_thms thyname thy =
+    if ! quick_and_dirty
+    then thy
+    else
+        case ImportData.get thy of
+            NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
+          | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
+
+fun clear_int_thms thy =
+    if ! quick_and_dirty
+    then thy
+    else
+        case ImportData.get thy of
+            NONE => error "Import data already cleared - forgotten a setup_theory?"
+          | SOME _ => ImportData.put NONE thy
+
+val setup_theory = (Parse.name -- Parse.name)
+                       >>
+                       (fn (location, thyname) =>
+                        fn thy =>
+                           (case Importer_DefThy.get thy of
+                                NoImport => thy
+                              | Generating _ => error "Currently generating a theory!"
+                              | Replaying _ => thy |> clear_import_thy)
+                               |> import_thy location thyname
+                               |> create_int_thms thyname)
+
+fun end_setup toks =
+    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
+                    |> clear_int_thms
+                    |> Sign.parent_path
+            end) toks
+
+val set_dump  = Parse.string -- Parse.string   >> setup_dump
+fun fl_dump toks  = Scan.succeed flush_dump toks
+val append_dump = (Parse.verbatim || Parse.string) >> add_dump
+
+val _ =
+  (Outer_Syntax.command @{command_spec "import_segment"} "set import segment name"
+     (import_segment >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "import_theory"} "set default external theory name"
+     (import_theory >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "end_import"} "end external import"
+     (end_import >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "setup_theory"} "setup external theory replaying"
+     (setup_theory >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "end_setup"} "end external setup"
+     (end_setup >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "setup_dump"} "setup the dump file name"
+     (set_dump >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "append_dump"} "add text to dump file"
+     (append_dump >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "flush_dump"} "write the dump to file"
+     (fl_dump >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "ignore_thms"}
+     "theorems to ignore in next external theory import"
+     (ignore_thms >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "type_maps"}
+     "map external type names to existing Isabelle/HOL type names"
+     (type_maps >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "def_maps"}
+     "map external constant names to their primitive definitions"
+     (def_maps >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "thm_maps"}
+     "map external theorem names to existing Isabelle/HOL theorem names"
+     (thm_maps >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "const_renames"}
+     "rename external const names"
+     (const_renames >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "const_moves"}
+     "rename external const names to other external constants"
+     (const_moves >> Toplevel.theory);
+   Outer_Syntax.command @{command_spec "const_maps"}
+     "map external const names to existing Isabelle/HOL const names"
+     (const_maps >> Toplevel.theory));
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/import_rews.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,633 @@
+(*  Title:      HOL/Import/import_rews.ML
+    Author:     Sebastian Skalberg (TU Muenchen)
+*)
+
+structure StringPair = Table(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 Importer_DefThy = Theory_Data
+(
+  type T = ImportStatus
+  val empty = NoImport
+  val extend = I
+  fun merge (NoImport, NoImport) = NoImport
+    | merge _ = (warning "Import status set during merge"; NoImport)
+)
+
+structure ImportSegment = Theory_Data
+(
+  type T = string
+  val empty = ""
+  val extend = 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"
+)
+
+val get_import_segment = ImportSegment.get
+val set_import_segment = ImportSegment.put
+
+structure Importer_UNames = Theory_Data
+(
+  type T = string list
+  val empty = []
+  val extend = I
+  fun merge ([], []) = []
+    | merge _ = error "Used names not empty during merge"
+)
+
+structure Importer_Dump = Theory_Data
+(
+  type T = string * string * string list
+  val empty = ("","",[])
+  val extend = I
+  fun merge (("","",[]),("","",[])) = ("","",[])
+    | merge _ = error "Dump data not empty during merge"
+)
+
+structure Importer_Moves = Theory_Data
+(
+  type T = string Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
+
+structure Importer_Imports = Theory_Data
+(
+  type T = string Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
+
+fun get_segment2 thyname thy =
+  Symtab.lookup (Importer_Imports.get thy) thyname
+
+fun set_segment thyname segname thy =
+    let
+        val imps = Importer_Imports.get thy
+        val imps' = Symtab.update_new (thyname,segname) imps
+    in
+        Importer_Imports.put imps' thy
+    end
+
+structure Importer_CMoves = Theory_Data
+(
+  type T = string Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
+
+structure Importer_Maps = Theory_Data
+(
+  type T = string option StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Thms = Theory_Data
+(
+  type T = holthm StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_ConstMaps = Theory_Data
+(
+  type T = (bool * string * typ option) StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Rename = Theory_Data
+(
+  type T = string StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_DefMaps = Theory_Data
+(
+  type T = string StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_TypeMaps = Theory_Data
+(
+  type T = (bool * string) StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Pending = Theory_Data
+(
+  type T = ((term * term) list * thm) StringPair.table
+  val empty = StringPair.empty
+  val extend = I
+  fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Rewrites = Theory_Data
+(
+  type T = thm list
+  val empty = []
+  val extend = I
+  val merge = Thm.merge_thms
+)
+
+val importer_debug = Unsynchronized.ref false
+fun message s = if !importer_debug then writeln s else ()
+
+fun add_importer_rewrite (Context.Theory thy, th) =
+    let
+        val _ = message "Adding external rewrite"
+        val th1 = th RS @{thm eq_reflection}
+                  handle THM _ => th
+        val current_rews = Importer_Rewrites.get thy
+        val new_rews = insert Thm.eq_thm th1 current_rews
+        val updated_thy  = Importer_Rewrites.put new_rews thy
+    in
+        (Context.Theory updated_thy,th1)
+    end
+  | add_importer_rewrite (context, th) = (context,
+      (th RS @{thm eq_reflection} handle THM _ => th)
+    );
+
+fun ignore_importer bthy bthm thy =
+    let
+        val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
+        val curmaps = Importer_Maps.get thy
+        val newmaps = StringPair.update_new ((bthy,bthm),NONE) curmaps
+        val upd_thy = Importer_Maps.put newmaps thy
+    in
+        upd_thy
+    end
+
+val opt_get_output_thy = #2 o Importer_Dump.get
+
+fun get_output_thy thy =
+    case #2 (Importer_Dump.get thy) of
+        "" => error "No theory file being output"
+      | thyname => thyname
+
+val get_output_dir = #1 o Importer_Dump.get
+
+fun add_importer_move bef aft thy =
+    let
+        val curmoves = Importer_Moves.get thy
+        val newmoves = Symtab.update_new (bef, aft) curmoves
+    in
+        Importer_Moves.put newmoves thy
+    end
+
+fun get_importer_move bef thy =
+  Symtab.lookup (Importer_Moves.get thy) bef
+
+fun follow_name thmname thy =
+    let
+        val moves = Importer_Moves.get thy
+        fun F thmname =
+            case Symtab.lookup moves thmname of
+                SOME name => F name
+              | NONE => thmname
+    in
+        F thmname
+    end
+
+fun add_importer_cmove bef aft thy =
+    let
+        val curmoves = Importer_CMoves.get thy
+        val newmoves = Symtab.update_new (bef, aft) curmoves
+    in
+        Importer_CMoves.put newmoves thy
+    end
+
+fun get_importer_cmove bef thy =
+  Symtab.lookup (Importer_CMoves.get thy) bef
+
+fun follow_cname thmname thy =
+    let
+        val moves = Importer_CMoves.get thy
+        fun F thmname =
+            case Symtab.lookup moves thmname of
+                SOME name => F name
+              | NONE => thmname
+    in
+        F thmname
+    end
+
+fun add_importer_mapping bthy bthm isathm thy =
+    let 
+       (* val _ = writeln ("Before follow_name: "^isathm)  *)
+      val isathm = follow_name isathm thy
+       (* val _ = writeln ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)*)
+        val curmaps = Importer_Maps.get thy
+        val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps
+        val upd_thy = Importer_Maps.put newmaps thy
+    in
+        upd_thy
+    end;
+
+fun get_importer_type_mapping bthy tycon thy =
+    let
+        val maps = Importer_TypeMaps.get thy
+    in
+        StringPair.lookup maps (bthy,tycon)
+    end
+
+fun get_importer_mapping bthy bthm thy =
+    let
+        val curmaps = Importer_Maps.get thy
+    in
+        StringPair.lookup curmaps (bthy,bthm)
+    end;
+
+fun add_importer_const_mapping bthy bconst internal isaconst thy =
+    let
+        val thy = case opt_get_output_thy thy of
+                      "" => thy
+                    | output_thy => if internal
+                                    then add_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+                                    else thy
+        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val curmaps = Importer_ConstMaps.get thy
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,NONE)) curmaps
+        val upd_thy = Importer_ConstMaps.put newmaps thy
+    in
+        upd_thy
+    end;
+
+fun add_importer_const_renaming bthy bconst newname thy =
+    let
+        val currens = Importer_Rename.get thy
+        val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
+        val newrens = StringPair.update_new ((bthy,bconst),newname) currens
+        val upd_thy = Importer_Rename.put newrens thy
+    in
+        upd_thy
+    end;
+
+fun get_importer_const_renaming bthy bconst thy =
+    let
+        val currens = Importer_Rename.get thy
+    in
+        StringPair.lookup currens (bthy,bconst)
+    end;
+
+fun get_importer_const_mapping bthy bconst thy =
+    let
+        val bconst = case get_importer_const_renaming bthy bconst thy of
+                         SOME name => name
+                       | NONE => bconst
+        val maps = Importer_ConstMaps.get thy
+    in
+        StringPair.lookup maps (bthy,bconst)
+    end
+
+fun add_importer_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_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+                                    else thy
+        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val curmaps = Importer_ConstMaps.get thy
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,SOME typ)) curmaps
+        val upd_thy = Importer_ConstMaps.put newmaps thy
+    in
+        upd_thy
+    end;
+
+fun add_importer_type_mapping bthy bconst internal isaconst thy =
+    let
+        val curmaps = Importer_TypeMaps.get thy
+        val _ = writeln ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst)) curmaps
+               (* FIXME avoid handle x *)
+               handle x => let val (_, isaconst') = the (StringPair.lookup curmaps (bthy, bconst)) in
+                      warning ("couldn't map type "^bthy^"."^bconst^" to "^isaconst^": already mapped to "^isaconst'); raise x end
+        val upd_thy = Importer_TypeMaps.put newmaps thy
+    in
+        upd_thy
+    end;
+
+fun add_importer_pending bthy bthm hth thy =
+    let
+        val thmname = Sign.full_bname thy bthm
+        val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
+        val curpend = Importer_Pending.get thy
+        val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
+        val upd_thy = Importer_Pending.put newpend thy
+        val thy' = case opt_get_output_thy upd_thy of
+                       "" => add_importer_mapping bthy bthm thmname upd_thy
+                     | output_thy =>
+                       let
+                           val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
+                       in
+                           upd_thy |> add_importer_move thmname new_thmname
+                                   |> add_importer_mapping bthy bthm new_thmname
+                       end
+    in
+        thy'
+    end;
+
+fun get_importer_theorem thyname thmname thy =
+  let
+    val isathms = Importer_Thms.get thy
+  in
+    StringPair.lookup isathms (thyname,thmname)
+  end;
+
+fun add_importer_theorem thyname thmname hth =
+  let
+    val _ = message ("Adding external theorem " ^ thyname ^ "." ^ thmname)
+  in
+    Importer_Thms.map (StringPair.update_new ((thyname, thmname), hth))
+  end;
+
+fun export_importer_pending thy =
+  let
+    val rews = Importer_Rewrites.get thy;
+    val pending = Importer_Pending.get thy;
+    fun process ((bthy,bthm), hth as (_,thm)) thy =
+      let
+        val thm1 = rewrite_rule (map (Thm.transfer thy) rews) (Thm.transfer thy thm);
+        val thm2 = Drule.export_without_context thm1;
+      in
+        thy
+        |> Global_Theory.store_thm (Binding.name bthm, thm2)
+        |> snd
+        |> add_importer_theorem bthy bthm hth
+      end;
+  in
+    thy
+    |> StringPair.fold process pending
+    |> Importer_Pending.put StringPair.empty
+  end;
+
+fun setup_dump (dir,thyname) thy =
+    Importer_Dump.put (dir,thyname,["(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)"]) thy
+
+fun add_dump str thy =
+    let
+        val (dir,thyname,curdump) = Importer_Dump.get thy
+    in
+        Importer_Dump.put (dir,thyname,str::curdump) thy
+    end
+
+fun flush_dump thy =
+    let
+        val (dir,thyname,dumpdata) = Importer_Dump.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
+        Importer_Dump.put ("","",[]) thy
+    end
+
+fun set_generating_thy thyname thy =
+    case Importer_DefThy.get thy of
+        NoImport => Importer_DefThy.put (Generating thyname) thy
+      | _ => error "Import already in progess"
+
+fun set_replaying_thy thyname thy =
+    case Importer_DefThy.get thy of
+        NoImport => Importer_DefThy.put (Replaying thyname) thy
+      | _ => error "Import already in progess"
+
+fun clear_import_thy thy =
+    case Importer_DefThy.get thy of
+        NoImport => error "No import in progress"
+      | _ => Importer_DefThy.put NoImport thy
+
+fun get_generating_thy thy =
+    case Importer_DefThy.get thy of
+        Generating thyname => thyname
+      | _ => error "No theory being generated"
+
+fun get_replaying_thy thy =
+    case Importer_DefThy.get thy of
+        Replaying thyname => thyname
+      | _ => error "No theory being replayed"
+
+fun get_import_thy thy =
+    case Importer_DefThy.get thy of
+        Replaying thyname => thyname
+      | Generating thyname => thyname
+      | _ => error "No theory being imported"
+
+fun should_ignore thyname thy thmname =
+    case get_importer_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 input_thy = Context.theory_name thy
+        val import_segment = get_import_segment 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.fold
+      (fn ((bthy, bthm), v) => fn (ign, map) =>
+        if bthy = thyname
+        then case v
+         of NONE => (bthm :: ign, map)
+          | SOME w => (ign, (bthm, w) :: map)
+        else (ign, map)) (Importer_Maps.get thy) ([],[]);
+    fun mk init = StringPair.fold
+      (fn ((bthy, bthm), v) => if bthy = thyname then cons (bthm, v) else I) init [];
+    val constmaps = mk (Importer_ConstMaps.get thy);
+    val constrenames = mk (Importer_Rename.get thy);
+    val typemaps = mk (Importer_TypeMaps.get thy);
+    val defmaps = mk (Importer_DefMaps.get thy);
+        fun new_name internal isa =
+            if internal
+            then
+                let
+                    val paths = Long_Name.explode isa
+                    val i = drop (length paths - 2) paths
+                in
+                    case i of
+                        [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
+                      | _ => error "import_rews.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,(_,isa,opt_ty)) =>
+                        (out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
+                         case opt_ty of
+                             SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy 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"
+
+        fun gen2replay in_thy out_thy s = 
+            let
+                val ss = Long_Name.explode s
+            in
+                if (hd ss = in_thy) then 
+                    Long_Name.implode (out_thy::(tl ss))
+                else
+                    s
+            end 
+
+        val _ = if null mapped
+                then ()
+                else out "thm_maps"
+        val _ = app (fn (hol,isa) => out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (gen2replay input_thy output_thy 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 = Importer_UNames.get thy
+    in
+        case unames of
+            [] => Importer_UNames.put names thy
+          | _ => error "import_rews.set_used_names called on initialized data!"
+    end
+
+val clear_used_names = Importer_UNames.put [];
+
+fun get_defmap thyname const thy =
+    let
+        val maps = Importer_DefMaps.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 = Importer_DefMaps.get thy
+        val maps' = StringPair.update_new ((thyname,const),defname) maps
+        val thy' = Importer_DefMaps.put maps' thy
+    in
+        thy'
+    end
+
+fun get_defname thyname name thy =
+    let
+        val maps = Importer_DefMaps.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 = Importer_UNames.get thy
+                val defname = Thm.def_name name
+                val pdefname = name ^ "_primdef"
+            in
+                if not (member (op =) used defname)
+                then F defname                 (* name_def *)
+                else if not (member (op =) used pdefname)
+                then F pdefname                (* name_primdef *)
+                else F (singleton (Name.variant_list used) pdefname) (* last resort *)
+            end
+    end
+
+local
+    fun handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "=="}, _],_,_]] = x
+      | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax all}, _],_]] = x
+      | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "==>"}, _],_,_]] = x
+      | handle_meta [x] = Ast.Appl [Ast.Constant @{const_syntax Trueprop}, x]
+      | handle_meta _ = error "import_rews error: Trueprop not applied to single argument"
+in
+val smarter_trueprop_parsing = [(@{const_syntax Trueprop},handle_meta)]
+end
+
+val importer_setup =
+  add_importer_type_mapping "min" "bool" false @{type_name bool}
+  #> add_importer_type_mapping "min" "fun" false "fun"
+  #> add_importer_type_mapping "min" "ind" false @{type_name ind}
+  #> add_importer_const_mapping "min" "=" false @{const_name HOL.eq}
+  #> add_importer_const_mapping "min" "==>" false @{const_name HOL.implies}
+  #> add_importer_const_mapping "min" "@" false @{const_name "Eps"}
+  #> Attrib.setup @{binding import_rew}
+    (Scan.succeed (Thm.mixed_attribute add_importer_rewrite)) "external rewrite rule";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/proof_kernel.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,2084 @@
+(*  Title:      HOL/Import/proof_kernel.ML
+    Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
+*)
+
+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 option
+    val disambiguate_frees : Thm.thm -> Thm.thm
+    val debug : bool Unsynchronized.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_importer_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 to_hol_thm : Thm.thm -> thm
+
+    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
+
+    val prin : term -> unit
+    val protect_factname : string -> string
+    val replay_protect_varname : string -> string -> unit
+    val replay_add_dump : string -> theory -> theory
+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
+
+fun hthm2thm (HOLThm (_, th)) = th
+
+fun to_hol_thm th = HOLThm ([], th)
+
+val replay_add_dump = add_dump
+fun add_dump s thy = replay_add_dump s thy
+
+datatype proof_info
+  = Info of {disk_info: (string * string) option Unsynchronized.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)
+
+
+(* Compatibility. *)
+
+val string_of_mixfix = Pretty.string_of o Mixfix.pretty_mixfix;
+
+fun mk_syn thy c =
+  if Lexicon.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
+  else Delimfix (Syntax_Ext.escape c)
+
+fun quotename c =
+  if Lexicon.is_identifier c andalso not (Keyword.is_keyword c) then c else quote c
+
+exception SMART_STRING
+
+fun no_vars context tm =
+  let
+    val ctxt = Variable.set_body false context;
+    val ([tm'], _) = Variable.import_terms true [tm] ctxt;
+  in tm' end
+
+fun smart_string_of_cterm ctxt0 ct =
+    let
+        val ctxt = ctxt0
+          |> Config.put show_brackets false
+          |> Config.put show_all_types false
+          |> Config.put show_types false
+          |> Config.put show_sorts false;
+        val {t,T,...} = rep_cterm ct
+        (* Hack to avoid parse errors with Trueprop *)
+        val t' = HOLogic.dest_Trueprop t
+                 handle TERM _ => t
+        val tn = no_vars ctxt t'
+        fun match u =
+            let val _ = Thm.match (ct, cterm_of (Proof_Context.theory_of ctxt) u) in true end
+            handle Pattern.MATCH => false
+        fun G 0 f ctxt x = f ctxt x
+          | G 1 f ctxt x = f (Config.put show_types true ctxt) x
+          | G 2 f ctxt x = G 1 f (Config.put show_sorts true ctxt) x
+          | G 3 f ctxt x = G 2 f (Config.put show_all_types true ctxt) x
+          | G 4 f ctxt x = G 3 f (Config.put show_brackets true ctxt) x
+          | G _ _ _ _ = raise SMART_STRING
+        fun F n =
+            let
+                val str = G n Syntax.string_of_term ctxt tn
+                val _ = warning (@{make_string} (n, str))
+                val u = Syntax.parse_term ctxt str
+                val u = if t = t' then u else HOLogic.mk_Trueprop u
+                val u = Syntax.check_term ctxt (Type.constraint T u)
+            in
+                if match u
+                then quote str
+                else F (n+1)
+            end
+            handle ERROR _ => F (n + 1)
+              | SMART_STRING =>
+                  let val _ =
+                    warning ("smart_string failed for: "^ G 0 Syntax.string_of_term ctxt (term_of ct))
+                  in quote (G 2 Syntax.string_of_term ctxt tn) end
+    in
+      Print_Mode.setmp [] F 0
+    end
+
+fun smart_string_of_thm ctxt = smart_string_of_cterm ctxt o cprop_of
+
+fun prth th = writeln (Print_Mode.setmp [] Display.string_of_thm_without_context th);
+val topctxt = ML_Context.the_local_context ();
+fun prin t = writeln (Print_Mode.setmp []
+  (fn () => Syntax.string_of_term topctxt t) ());
+fun pth (HOLThm(_,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 = Unsynchronized.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 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
+infix mem;
+fun i mem L =
+    let fun itr [] = false
+          | itr (a::rst) = i=a orelse itr rst
+    in itr L end;
+
+infix union;
+fun [] union S = S
+  | S union [] = S
+  | (a::rst) union S2 = rst union (insert (op =) a S2);
+
+fun implode_subst [] = []
+  | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
+  | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
+
+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
+
+(* Actual code. *)
+
+fun get_segment thyname l = (Lib.assoc "s" l
+                             handle PK _ => thyname)
+val get_name : (string * string) list -> string = Lib.assoc "n"
+
+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 xml_to_import_xml (XML.Elem ((n, l), ts)) = Elem (n, l, map xml_to_import_xml ts)
+  | xml_to_import_xml (XML.Text _) = raise XML "Incorrect proof file: text";
+
+val type_of = Term.type_of
+
+val propT = Type("prop",[])
+
+fun mk_defeq name rhs thy =
+    let
+        val ty = type_of rhs
+    in
+        Logic.mk_equals (Const(Sign.intern_const 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 thy name,ty),rhs)
+    end
+
+fun intern_const_name thyname const thy =
+    case get_importer_const_mapping thyname const thy of
+        SOME (_,cname,_) => cname
+      | NONE => (case get_importer_const_renaming thyname const thy of
+                     SOME cname => Sign.intern_const thy (thyname ^ "." ^ cname)
+                   | NONE => Sign.intern_const thy (thyname ^ "." ^ const))
+
+fun intern_type_name thyname const thy =
+    case get_importer_type_mapping thyname const thy of
+        SOME (_,cname) => cname
+      | NONE => Sign.intern_const 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 mk_thy_const thy Thy Nam Ty = Const(intern_const_name Thy Nam thy,Ty)
+
+local
+  fun get_const sg _ name =
+    (case Sign.const_type sg name of
+      SOME ty => Const (name, ty)
+    | NONE => raise ERR "get_type" (name ^ ": No such constant"))
+in
+fun prim_mk_const thy Thy Nam =
+    let
+      val name = intern_const_name Thy Nam thy
+      val cmaps = Importer_ConstMaps.get thy
+    in
+      case StringPair.lookup cmaps (Thy,Nam) of
+        SOME(_,_,SOME ty) => Const(name,ty)
+      | _ => get_const thy Thy name
+    end
+end
+
+fun mk_comb(f,a) = f $ 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
+
+val protected_varnames = Unsynchronized.ref (Symtab.empty:string Symtab.table)
+val invented_isavar = Unsynchronized.ref 0
+
+fun innocent_varname s = Lexicon.is_identifier s andalso not (String.isPrefix "u_" s)
+
+fun valid_boundvarname s =
+  can (fn () => Syntax.read_term_global @{theory Main} ("SOME "^s^". True")) ();
+
+fun valid_varname s =
+  can (fn () => Syntax.read_term_global @{theory Main} s) ();
+
+fun protect_varname s =
+    if innocent_varname s andalso valid_varname s then s else
+    case Symtab.lookup (!protected_varnames) s of
+      SOME t => t
+    | NONE =>
+      let
+          val _ = Unsynchronized.inc invented_isavar
+          val t = "u_" ^ string_of_int (!invented_isavar)
+          val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
+      in
+          t
+      end
+
+exception REPLAY_PROTECT_VARNAME of string*string*string
+
+fun replay_protect_varname s t =
+        case Symtab.lookup (!protected_varnames) s of
+          SOME t' => raise REPLAY_PROTECT_VARNAME (s, t, t')
+        | NONE =>
+          let
+              val _ = Unsynchronized.inc invented_isavar
+              val t = "u_" ^ string_of_int (!invented_isavar)
+              val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
+          in
+              ()
+          end
+
+fun protect_boundvarname s = if innocent_varname s andalso valid_boundvarname s then s else "u"
+
+fun mk_lambda (v as Free (x, T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
+  | mk_lambda (v as Var ((x, _), T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
+  | mk_lambda v t = raise TERM ("lambda", [v, t]);
+
+fun replacestr x y s =
+let
+  val xl = raw_explode x
+  val yl = raw_explode y
+  fun isprefix [] _ = true
+    | isprefix (x::xs) (y::ys) = if x = y then isprefix xs ys else false
+    | isprefix _ _ = false
+  fun isp s = isprefix xl s
+  fun chg s = yl@(List.drop (s, List.length xl))
+  fun r [] = []
+    | r (S as (s::ss)) = if isp S then r (chg S) else s::(r ss)
+in
+  implode(r (raw_explode s))
+end
+
+fun protect_factname s = replacestr "." "_dot_" s
+fun unprotect_factname s = replacestr "_dot_" "." s
+
+val ty_num_prefix = "N_"
+
+fun startsWithDigit s = Char.isDigit (hd (String.explode s))
+
+fun protect_tyname tyn =
+  let
+    val tyn' =
+      if String.isPrefix ty_num_prefix tyn then raise (ERR "protect_ty_name" ("type name '"^tyn^"' is reserved")) else
+      (if startsWithDigit tyn then ty_num_prefix^tyn else tyn)
+  in
+    tyn'
+  end
+
+fun protect_constname tcn = tcn
+ (* if tcn = ".." then "dotdot"
+  else if tcn = "==" then "eqeq"
+  else tcn*)
+
+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)
+                        (protect_tyname (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)
+                        (protect_tyname (get_name atts))
+                        (map gtfx tys)
+          | gtfx _ = raise ERR "get_type" "Bad type"
+    in
+        gtfx
+    end
+
+fun input_types _ (Elem("tylist",[("i",i)],xtys)) =
+    let
+        val types = Array.array(the (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 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 = protect_constname(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_lambda 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 _ _ (Elem("tmlist",[("i",i)],xtms)) =
+    let
+        val terms = Array.array(the (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 path = space_explode ":" (getenv "IMPORTER_PROOFS")
+        fun find [] = NONE
+          | find (p::ps) =
+            (let
+                 val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
+             in
+                 if OS.FileSys.isDir dir
+                 then SOME dir
+                 else find ps
+             end) handle OS.SysErr _ => find ps
+    in
+        Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
+    end
+
+fun proof_file_name thyname thmname thy =
+    let
+        val path = case get_proof_dir thyname thy of
+                       SOME p => p
+                     | NONE => error "Cannot find proof files"
+        val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
+    in
+        OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = (unprotect_factname thmname), ext = SOME "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 = protect_factname (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,protect_constname 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,protect_tyname name,protect_constname abs_name,protect_constname rep_name,P,t,x2p p))
+            end
+          | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
+            mk_proof (PTyDef(seg,protect_tyname name,x2p p))
+          | x2p (Elem("poracle",[],chldr)) =
+            let
+                val (oracles,terms) = List.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
+                val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | _ => raise ERR "x2p" "bad oracle") oracles
+                val (c,asl) = case terms of
+                                  [] => raise ERR "x2p" "Bad oracle description"
+                                | (hd::tl) => (hd,tl)
+                val tg = fold_rev (Tag.merge o Tag.read) 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 _ = raise ERR "x2p" "Bad proof"
+    in
+        x2p prf
+    end
+
+fun import_proof_concl thyname thmname thy =
+    let
+        val is = TextIO.openIn(proof_file_name thyname thmname thy)
+        val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll is))
+        val _ = TextIO.closeIn is
+    in
+        case proof_xml of
+            Elem("proof",[],xtypes::xterms::_::rest) =>
+            let
+                val types = TypeNet.input_types thyname xtypes
+                val terms = TermNet.input_terms thyname types xterms
+                fun f xtm thy = TermNet.get_term_from_xml thy thyname types terms xtm
+            in
+                case rest of
+                    [] => NONE
+                  | [xtm] => SOME (f xtm)
+                  | _ => raise ERR "import_proof" "Bad argument list"
+            end
+          | _ => raise ERR "import_proof" "Bad proof"
+    end
+
+fun import_proof thyname thmname thy =
+    let
+        val is = TextIO.openIn(proof_file_name thyname thmname thy)
+        val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll 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 = Thm.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 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 thy = theory_of_thm th
+    val pp = Thm.reflexive (cterm_of thy (Const(@{const_name Trueprop},HOLogic.boolT-->propT)))
+in
+val not_elim_thm = Thm.combination pp th
+end
+
+val not_intro_thm = Thm.symmetric not_elim_thm
+val abs_thm = @{thm ext}
+val trans_thm = @{thm trans}
+val symmetry_thm = @{thm sym}
+val eqmp_thm = @{thm iffD1}
+val eqimp_thm = @{thm Importer.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 =
+    Library.foldl (fn (th,p) => Thm.implies_elim th (Thm.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 (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(@{const_name All},_) $ Abs(oldname,_,_)) =>
+                              (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
+                            | tpc $ (Const(@{const_name All},_) $ _) => (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
+
+fun rearrange sg tm th =
+    let
+        val tm' = Envir.beta_eta_contract tm
+        fun find []      n = Thm.permute_prems 0 1 (Thm.implies_intr (Thm.cterm_of sg tm) th)
+          | find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
+                             then Thm.permute_prems n 1 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
+
+(* Code for disambiguating variablenames (wrt. types) *)
+
+val disamb_info_empty = {vars=[],rens=[]}
+
+fun rens_of { vars = _, rens = rens } = rens
+
+fun disamb_term_from info tm = (info, tm)
+
+fun has_ren (HOLThm _) = false
+
+fun disamb_thm_from info (HOLThm (_,thm)) = (info, thm)
+
+fun disamb_terms_from info tms = (info, tms)
+
+fun disamb_thms_from info hthms = (info, map hthm2thm hthms)
+
+fun disamb_term tm   = disamb_term_from disamb_info_empty tm
+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 _ (hth as HOLThm _) = hth
+
+(* End of disambiguating code *)
+
+fun disambiguate_frees thm =
+    let
+      fun ERR s = error ("Drule.disambiguate_frees: "^s)
+      val ct = cprop_of thm
+      val t = term_of ct
+      val thy = theory_of_cterm ct
+      val frees = Misc_Legacy.term_frees t
+      val freenames = Term.add_free_names t []
+      val is_old_name = member (op =) freenames
+      fun name_of (Free (n, _)) = n
+        | name_of _ = ERR "name_of"
+      fun new_name' bump map n =
+          let val n' = n^bump in
+            if is_old_name n' orelse Symtab.lookup map n' <> NONE then
+              new_name' (Symbol.bump_string bump) map n
+            else
+              n'
+          end
+      val new_name = new_name' "a"
+      fun replace_name n' (Free (_, t)) = Free (n', t)
+        | replace_name _ _ = ERR "replace_name"
+      (* map: old or fresh name -> old free,
+         invmap: old free which has fresh name assigned to it -> fresh name *)
+      fun dis v (mapping as (map,invmap)) =
+          let val n = name_of v in
+            case Symtab.lookup map n of
+              NONE => (Symtab.update (n, v) map, invmap)
+            | SOME v' =>
+              if v=v' then
+                mapping
+              else
+                let val n' = new_name map n in
+                  (Symtab.update (n', v) map,
+                   Termtab.update (v, n') invmap)
+                end
+          end
+    in
+      if (length freenames = length frees) then
+        thm
+      else
+        let
+          val (_, invmap) =
+              fold dis frees (Symtab.empty, Termtab.empty)
+          fun make_subst (oldfree, newname) (intros, elims) =
+              (cterm_of thy oldfree :: intros,
+               cterm_of thy (replace_name newname oldfree) :: elims)
+          val (intros, elims) = fold make_subst (Termtab.dest invmap) ([], [])
+        in
+          forall_elim_list elims (forall_intr_list intros thm)
+        end
+    end
+
+val debug = Unsynchronized.ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+fun get_importer_thm thyname thmname thy =
+    case get_importer_theorem thyname thmname thy of
+        SOME hth => SOME (HOLThm hth)
+      | NONE =>
+        let
+            val pending = Importer_Pending.get thy
+        in
+            case StringPair.lookup pending (thyname,thmname) of
+                SOME hth => SOME (HOLThm hth)
+              | NONE => NONE
+        end
+
+fun non_trivial_term_consts t = fold_aterms
+  (fn Const (c, _) =>
+      if c = @{const_name Trueprop} orelse c = @{const_name All}
+        orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
+      then I else insert (op =) c
+    | _ => I) t [];
+
+fun split_name str =
+    let
+        val sub = Substring.full str
+        val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
+        val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
+    in
+        if not (idx = "") andalso u = "_"
+        then SOME (newstr, the (Int.fromString idx))
+        else NONE
+    end
+    handle _ => NONE  (* FIXME avoid handle _ *)
+
+fun rewrite_importer_term t thy =
+    let
+        val import_rews1 = map (Thm.transfer thy) (Importer_Rewrites.get thy)
+        val importerss = Simplifier.global_context thy empty_ss addsimps import_rews1
+    in
+        Thm.transfer thy (Simplifier.full_rewrite importerss (cterm_of thy t))
+    end
+
+fun get_isabelle_thm thyname thmname importerconc thy =
+    let
+        val (info,importerconc') = disamb_term importerconc
+        val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') 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 importerconc';
+                 message "Modified conclusion:";
+                 if_debug prin isaconc)
+
+        fun mk_res th = HOLThm (rens_of info, Thm.equal_elim i2h_conc th)
+    in
+        case get_importer_mapping thyname thmname thy of
+            SOME (SOME thmname) =>
+            let
+                val th1 = (SOME (Global_Theory.get_thm thy thmname)
+                           handle ERROR _ =>
+                                  (case split_name thmname of
+                                       SOME (listname,idx) => (SOME (nth (Global_Theory.get_thms thy listname) (idx - 1))
+                                                               handle _ => NONE)  (* FIXME avoid handle _ *)
+                                     | 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 (commas cs))
+                val pot_thms = Shuffler.find_potential thy isaconc
+                val _ = message (string_of_int (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_importer_theorem thyname thmname args
+                                        |> add_importer_mapping thyname thmname isaname
+                    in
+                        (thy',SOME hth)
+                    end
+                  | NONE => (thy,NONE)
+            end
+    end
+    handle e =>
+      if Exn.is_interrupt e then reraise e
+      else
+        (if_debug (fn () =>
+            writeln ("Exception in get_isabelle_thm:\n" ^ ML_Compiler.exn_message e)) ();
+          (thy,NONE))
+
+fun get_isabelle_thm_and_warn thyname thmname importerconc thy =
+  let
+    val (a, b) = get_isabelle_thm thyname thmname importerconc thy
+    fun warn () =
+        let
+            val (_,importerconc') = disamb_term importerconc
+            val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') thy)
+        in
+            case concl_of i2h_conc of
+                Const("==",_) $ lhs $ _ =>
+                (warning ("Failed lookup of theorem '"^thmname^"':");
+                 writeln "Original conclusion:";
+                 prin importerconc';
+                 writeln "Modified conclusion:";
+                 prin lhs)
+              | _ => ()
+        end
+  in
+      case b of
+          NONE => (warn () handle _ => (); (a,b)) (* FIXME avoid handle _ *)
+        | _ => (a, b)
+  end
+
+fun get_thm thyname thmname thy =
+    case get_importer_thm thyname thmname thy of
+        SOME hth => (thy,SOME hth)
+      | NONE => ((case import_proof_concl thyname thmname thy of
+                      SOME f => get_isabelle_thm_and_warn thyname thmname (f thy) thy
+                    | NONE => (message "No conclusion"; (thy,NONE)))
+                 handle IO.Io _ => (message "IO exception"; (thy,NONE))
+                      | PK _ => (message "PK exception"; (thy,NONE)))
+
+fun rename_const thyname thy name =
+    case get_importer_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_and_warn 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 (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
+        val rew = rewrite_importer_term (concl_of th) thy
+        val th  = Thm.equal_elim rew th
+        val thy' = add_importer_pending thyname thmname args thy
+        val th = disambiguate_frees th
+        val th = Object_Logic.rulify th
+        val thy2 =
+          if gen_output
+          then
+            add_dump ("lemma " ^ (quotename thmname) ^ ": " ^
+                (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^ "\n  by (import " ^
+                thyname ^ " " ^ (quotename 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 ctm = Thm.cterm_of thy 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 ctm = Thm.cterm_of thy (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(_,th)) thy =
+    let
+        val _ = message "INST_TYPE:"
+        val _ = if_debug pth hth
+        val tys_before = Misc_Legacy.add_term_tfrees (prop_of th,[])
+        val th1 = Thm.varifyT_global th
+        val tys_after = Misc_Legacy.add_term_tvars (prop_of th1,[])
+        val tyinst = map (fn (bef, iS) =>
+                             (case try (Lib.assoc (TFree bef)) lambda of
+                                  SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
+                                | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
+                             ))
+                         (zip tys_before tys_after)
+        val res = Drule.instantiate_normalize (tyinst,[]) th1
+        val hth = HOLThm([],res)
+        val res = norm_hthm thy 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 (sdom,srng) = ListPair.unzip (rev sigma)
+        val th = hthm2thm hth
+        val th1 = mk_INST (map (cterm_of thy) sdom) (map (cterm_of thy) srng) th
+        val res = HOLThm([],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 thy =
+    let
+        val (f,g) = case concl_of th1 of
+                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
+                      | _ => raise ERR "mk_COMB" "First theorem not an equality"
+        val (x,y) = case concl_of th2 of
+                        _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
+                      | _ => raise ERR "mk_COMB" "Second theorem not an equality"
+        val fty = type_of f
+        val (fd,fr) = Term.dest_funT fty
+        val comb_thm' = Drule.instantiate'
+                            [SOME (ctyp_of thy fd),SOME (ctyp_of thy fr)]
+                            [SOME (cterm_of thy f),SOME (cterm_of thy g),
+                             SOME (cterm_of thy x),SOME (cterm_of thy 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 cctxt = cterm_of thy ctxt'
+        fun subst th [] = th
+          | subst th (rew::rews) = subst (mk_COMB th rew thy) 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 th1 = norm_hyps th1
+        val th2 = norm_hyps th2
+        val (l,r) = case concl_of th of
+                        _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
+                      | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
+        val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
+        val th2' = rearrange thy (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 ct = Thm.cterm_of thy 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 ct = Thm.cterm_of thy 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 cwit = cterm_of thy wit'
+        val cty = ctyp_of_term cwit
+        val a = case ex' of
+                    (Const(@{const_name Ex},_) $ a) => a
+                  | _ => raise ERR "EXISTS" "Argument not existential"
+        val ca = cterm_of thy 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 (Thm.implies_elim th (Thm.assume p))
+          | strip _ _ _ = raise ERR "CHOOSE" "strip error"
+        val cv = cterm_of thy v'
+        val th2 = norm_hyps th2
+        val cvty = ctyp_of_term cv
+        val c = HOLogic.dest_Trueprop (concl_of th2)
+        val cc = cterm_of thy c
+        val a = case concl_of th1 of
+                    _ $ (Const(@{const_name Ex},_) $ a) => a
+                  | _ => raise ERR "CHOOSE" "Conclusion not existential"
+        val ca = cterm_of (theory_of_thm th1) a
+        val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
+        val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
+        val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
+        val th23 = beta_eta_thm (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 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 ctm = Thm.cterm_of thy 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 res = HOLThm(rens_of info,mk_COMB th1 th2 thy)
+        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 ct = cterm_of thy tm'
+        val th1 = rearrange thy (HOLogic.mk_Trueprop (Const(@{const_name Not},HOLogic.boolT-->HOLogic.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 thy =
+    let
+        val cv = cterm_of thy v
+        val th1 = implies_elim_all (beta_eta_thm th)
+        val (f,g) = case concl_of th1 of
+                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+                      | _ => raise ERR "mk_ABS" "Bad conclusion"
+        val (fd,fr) = Term.dest_funT (type_of f)
+        val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
+        val th2 = Thm.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 res = HOLThm(rens_of info',mk_ABS v' th thy)
+        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 th1 =
+            case copt of
+                SOME (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 (Type(name,tys)) =
+                        Type(name,map (inst_type ty1 ty2) tys)
+                in
+                    fold_rev (fn v => fn th =>
+                              let
+                                  val cdom = fst (Term.dest_funT (fst (Term.dest_funT cty)))
+                                  val vty  = type_of v
+                                  val newcty = inst_type cdom vty cty
+                                  val cc = cterm_of thy (Const(cname,newcty))
+                              in
+                                  mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy
+                              end) vlist' th
+                end
+              | SOME _ => raise ERR "GEN_ABS" "Bad constant"
+              | NONE =>
+                fold_rev (fn v => fn th => mk_ABS v th thy) 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 th1 = implies_elim_all (beta_eta_thm th)
+        val a = case concl_of th1 of
+                    _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
+                  | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
+        val ca = cterm_of thy a
+        val th2 = Thm.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 th1 = implies_elim_all (beta_eta_thm th)
+        val a = case concl_of th1 of
+                    _ $ (Const(@{const_name Not},_) $ a) => a
+                  | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
+        val ca = cterm_of thy a
+        val th2 = Thm.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 th1 = beta_eta_thm th
+        val th2 = implies_elim_all th1
+        val th3 = Thm.implies_intr (cterm_of thy (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 = space_implode " "
+
+fun new_definition thyname constname rhs thy =
+    let
+        val constname = rename_const thyname thy constname
+        val redeclared = is_some (Sign.const_type thy (Sign.intern_const thy constname));
+        val _ = warning ("Introducing constant " ^ constname)
+        val (thmname,thy) = get_defname thyname constname thy
+        val (_,rhs') = disamb_term rhs
+        val ctype = type_of rhs'
+        val csyn = mk_syn thy constname
+        val thy1 = case Importer_DefThy.get thy of
+                       Replaying _ => thy
+                     | _ => Sign.add_consts_i [(Binding.name constname,ctype,csyn)] thy
+        val eq = mk_defeq constname rhs' thy1
+        val (thms, thy2) = Global_Theory.add_defs false [((Binding.name 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 thy' constname
+        val thy'' = add_importer_const_mapping thyname constname true fullcname thy'
+        val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
+        val rew = rewrite_importer_term eq thy''
+        val crhs = cterm_of thy'' (#2 (Logic.dest_equals (prop_of rew)))
+        val thy22 =
+          if Thm.def_name constname = thmname andalso not redeclared andalso csyn = NoSyn
+          then
+              let
+                  val ctxt = Syntax.init_pretty_global thy''
+                  val p1 = quotename constname
+                  val p2 = Syntax.string_of_typ ctxt ctype
+                  val p3 = string_of_mixfix csyn
+                  val p4 = smart_string_of_cterm ctxt crhs
+              in
+                add_dump ("definition\n  " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ " where\n  " ^ p4) thy''
+              end
+          else
+              let val ctxt = Syntax.init_pretty_global thy'' in
+                add_dump ("consts\n  " ^ quotename constname ^ " :: \"" ^
+                  Syntax.string_of_typ ctxt ctype ^
+                  "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n  " ^
+                  quotename thmname ^ ": " ^ smart_string_of_cterm ctxt crhs) thy''
+              end
+        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_bname thy22 thmname
+        val thy22' = case opt_get_output_thy thy22 of
+                         "" => add_importer_mapping thyname thmname fullname thy22
+                       | output_thy =>
+                         let
+                             val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
+                         in
+                             thy22 |> add_importer_move fullname moved_thmname
+                                   |> add_importer_mapping thyname thmname moved_thmname
+                         end
+        val _ = message "new_definition:"
+        val _ = if_debug pth hth
+    in
+        (thy22',hth)
+    end
+
+fun new_specification thyname thmname names hth thy =
+    case Importer_DefThy.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 " ^ commas names)
+            val (HOLThm(rens,th)) = norm_hthm thy hth
+            val thy1 = case Importer_DefThy.get thy of
+                           Replaying _ => thy
+                         | _ =>
+                           let
+                               fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
+                                 | dest_eta_abs body =
+                                   let
+                                       val (dT,_) = Term.dest_funT (type_of body)
+                                   in
+                                       ("x",dT,body $ Bound 0)
+                                   end
+                                   handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
+                               fun dest_exists (Const(@{const_name Ex},_) $ abody) =
+                                   dest_eta_abs abody
+                                 | dest_exists _ =
+                                   raise ERR "new_specification" "Bad existential formula"
+
+                               val (consts,_) = Library.foldl (fn ((cs,ex),cname) =>
+                                                          let
+                                                              val (_,cT,p) = dest_exists ex
+                                                          in
+                                                              ((cname,cT,mk_syn thy cname)::cs,p)
+                                                          end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
+                               val str = Library.foldl (fn (acc, (c, T, csyn)) =>
+                                   acc ^ "\n  " ^ quotename c ^ " :: \"" ^
+                                   Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
+                               val thy' = add_dump str thy
+                           in
+                               Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy'
+                           end
+
+            val thy1 = fold_rev (fn name => fn thy =>
+                                snd (get_defname thyname name thy)) names thy1
+            fun new_name name = fst (get_defname thyname name thy1)
+            val names' = map (fn name => (new_name name,name,false)) names
+            val (thy',res) = Choice_Specification.add_specification NONE
+                                 names'
+                                 (thy1,th)
+            val res' = Thm.unvarify_global res
+            val hth = HOLThm(rens,res')
+            val rew = rewrite_importer_term (concl_of res') thy'
+            val th  = Thm.equal_elim rew res'
+            fun handle_const (name,thy) =
+                let
+                    val defname = Thm.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') = fold_rev (fn name => fn (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 (Syntax.init_pretty_global thy') 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
+
+fun new_axiom name _ _ = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
+
+fun to_isa_thm (hth as HOLThm(_,th)) =
+    let
+        val (HOLThm args) = norm_hthm (theory_of_thm th) hth
+    in
+        apsnd Thm.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 Importer_DefThy.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 thy hth
+            val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
+            val c = case concl_of th2 of
+                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+                      | _ => raise ERR "new_type_definition" "Bad type definition theorem"
+            val tfrees = Misc_Legacy.term_tfrees c
+            val tnames = map fst tfrees
+            val tsyn = mk_syn thy tycname
+            val ((_, typedef_info), thy') =
+              Typedef.add_typedef_global false (SOME (Binding.name thmname))
+                (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c NONE (rtac th2 1) thy
+
+            val th3 = (#type_definition (#2 typedef_info)) RS typedef_hol2hol4
+
+            val fulltyname = Sign.intern_type thy' tycname
+            val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
+
+            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th3))
+            val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
+                    else ()
+            val thy4 = add_importer_pending thyname thmname args thy''
+
+            val rew = rewrite_importer_term (concl_of td_th) thy4
+            val th  = Thm.equal_elim rew (Thm.transfer thy4 td_th)
+            val c   = case HOLogic.dest_Trueprop (prop_of th) of
+                          Const(@{const_name Ex},exT) $ P =>
+                          let
+                              val PT = domain_type exT
+                          in
+                              Const (@{const_name 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 "(" ^ commas tnames ^ ") "
+            val proc_prop =
+              smart_string_of_cterm
+                (Syntax.init_pretty_global thy4
+                  |> not (null tnames) ? Config.put show_all_types true)
+            val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " "
+                                 ^ (string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
+
+            val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2importer [OF type_definition_" ^ tycname ^ "]") thy5
+
+            val _ = message "RESULT:"
+            val _ = if_debug pth hth'
+        in
+            (thy6,hth')
+        end
+
+fun add_dump_syntax thy name =
+    let
+      val n = quotename name
+      val syn = string_of_mixfix (mk_syn thy name)
+    in
+      add_dump ("syntax\n  "^n^" :: _ "^syn) thy
+    end
+
+fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
+    case Importer_DefThy.get thy of
+        Replaying _ => (thy,
+          HOLThm([], Global_Theory.get_thm thy (thmname^"_@intern")) handle ERROR _ => 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 thy hth
+            val tT = type_of t
+            val light_nonempty' =
+                Drule.instantiate' [SOME (ctyp_of thy tT)]
+                                   [SOME (cterm_of thy P),
+                                    SOME (cterm_of thy t)] light_nonempty
+            val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
+            val c = case concl_of th2 of
+                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+                      | _ => raise ERR "type_introduction" "Bad type definition theorem"
+            val tfrees = Misc_Legacy.term_tfrees c
+            val tnames = sort_strings (map fst tfrees)
+            val tsyn = mk_syn thy tycname
+            val ((_, typedef_info), thy') =
+              Typedef.add_typedef_global false NONE
+                (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c
+                (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
+            val fulltyname = Sign.intern_type thy' tycname
+            val aty = Type (fulltyname, map mk_vartype tnames)
+            val typedef_hol2hollight' =
+                Drule.instantiate'
+                    [SOME (ctyp_of thy' aty), SOME (ctyp_of thy' tT)]
+                    [NONE, NONE, NONE, SOME (cterm_of thy' (Free ("a", aty))), SOME (cterm_of thy' (Free ("r", tT)))]
+                    typedef_hol2hollight
+            val th4 = (#type_definition (#2 typedef_info)) RS typedef_hol2hollight'
+            val _ = null (Thm.fold_terms Term.add_tvars th4 []) orelse
+              raise ERR "type_introduction" "no type variables expected any more"
+            val _ = null (Thm.fold_terms Term.add_vars th4 []) orelse
+              raise ERR "type_introduction" "no term variables expected any more"
+            val _ = message ("step 3: thyname="^thyname^", tycname="^tycname^", fulltyname="^fulltyname)
+            val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
+            val _ = message "step 4"
+            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th4))
+            val thy4 = add_importer_pending thyname thmname args thy''
+
+            val P' = P (* why !? #2 (Logic.dest_equals (concl_of (rewrite_importer_term P thy4))) *)
+            val c   =
+                let
+                    val PT = type_of P'
+                in
+                    Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P'
+                end
+
+            val tnames_string = if null tnames
+                                then ""
+                                else "(" ^ commas tnames ^ ") "
+            val proc_prop =
+              smart_string_of_cterm
+                (Syntax.init_pretty_global thy4
+                  |> not (null tnames) ? Config.put show_all_types true)
+            val thy = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^
+              " = " ^ (proc_prop (cterm_of thy4 c)) ^ " " ^
+              (string_of_mixfix tsyn) ^ " morphisms "^
+              (quote rep_name)^" "^(quote abs_name)^"\n"^
+              ("  apply (rule light_ex_imp_nonempty[where t="^
+              (proc_prop (cterm_of thy4 t))^"])\n"^
+              ("  by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
+            val str_aty = Syntax.string_of_typ_global thy aty
+            val thy = add_dump_syntax thy rep_name
+            val thy = add_dump_syntax thy abs_name
+            val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
+              " = typedef_hol2hollight \n"^
+              "  [where a=\"a :: "^str_aty^"\" and r=r" ^
+              " ,\n   OF "^(quotename ("type_definition_" ^ tycname)) ^ "]") thy
+            val _ = message "RESULT:"
+            val _ = if_debug pth hth'
+        in
+            (thy,hth')
+        end
+end
+
+val prin = prin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/replay.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,344 @@
+(*  Title:      HOL/Import/replay.ML
+    Author:     Sebastian Skalberg (TU Muenchen)
+*)
+
+structure Replay =  (* FIXME proper signature *)
+struct
+
+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 = ProofKernel.REFL tm thy
+          | rp (PInstT(p,lambda)) thy =
+            let
+                val (thy',th) = rp' p thy
+            in
+                ProofKernel.INST_TYPE lambda th thy'
+            end
+          | rp (PSubst(prfs,ctxt,prf)) thy =
+            let
+                val (thy',ths) = fold_rev (fn p => fn (thy, ths) =>
+                                           let
+                                               val (thy',th) = rp' p thy
+                                           in
+                                               (thy',th::ths)
+                                           end) prfs (thy,[])
+                val (thy'',th) = rp' prf thy'
+            in
+                ProofKernel.SUBST ths ctxt th thy''
+            end
+          | rp (PAbs(prf,v)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.ABS v th thy'
+            end
+          | rp (PDisch(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.DISCH tm th thy'
+            end
+          | rp (PMp(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                ProofKernel.MP th1 th2 thy2
+            end
+          | rp (PHyp tm) thy = ProofKernel.ASSUME tm thy
+          | rp (PDef(seg,name,rhs)) thy =
+            (case ProofKernel.get_def seg name rhs thy of
+                 (thy',SOME res) => (thy',res)
+               | (thy',NONE) => 
+                 if seg = thyname
+                 then ProofKernel.new_definition seg name rhs thy'
+                 else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
+          | rp (POracle _) thy = NY "ORACLE"
+          | rp (PSpec(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.SPEC tm th thy'
+            end
+          | rp (PInst(prf,theta)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.INST theta th thy'
+            end
+          | rp (PGen(prf,v)) thy =
+            let
+                val (thy',th) = rp' prf thy
+                val p = ProofKernel.GEN v th thy'
+            in
+                p
+            end
+          | rp (PGenAbs(prf,opt,vl)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.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
+                ProofKernel.IMP_ANTISYM th1 th2 thy2
+            end
+          | rp (PSym prf) thy =
+            let
+                val (thy1,th) = rp' prf thy
+            in
+                ProofKernel.SYM th thy1
+            end
+          | rp (PTrans(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                ProofKernel.TRANS th1 th2 thy2
+            end
+          | rp (PComb(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                ProofKernel.COMB th1 th2 thy2
+            end
+          | rp (PEqMp(prf1,prf2)) thy =
+            let
+                val (thy1,th1) = rp' prf1 thy
+                val (thy2,th2) = rp' prf2 thy1
+            in
+                ProofKernel.EQ_MP th1 th2 thy2
+            end
+          | rp (PEqImp prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.EQ_IMP_RULE th thy'
+            end
+          | rp (PExists(prf,ex,wit)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.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
+                ProofKernel.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
+                ProofKernel.CONJ th1 th2 thy2
+            end
+          | rp (PConjunct1 prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.CONJUNCT1 th thy'
+            end
+          | rp (PConjunct2 prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.CONJUNCT2 th thy'
+            end
+          | rp (PDisj1(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.DISJ1 th tm thy'
+            end
+          | rp (PDisj2(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.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
+                ProofKernel.DISJ_CASES th th1 th2 thy2
+            end
+          | rp (PNotI prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.NOT_INTRO th thy'
+            end
+          | rp (PNotE prf) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.NOT_ELIM th thy'
+            end
+          | rp (PContr(prf,tm)) thy =
+            let
+                val (thy',th) = rp' prf thy
+            in
+                ProofKernel.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 ProofKernel.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 (_, prf) = import_proof thyname' thmname thy'
+                                                    val prf = prf thy'
+                                                    val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
+                                                    val _ = writeln ("Successfully finished replaying "^thmname^" !")
+                                                in
+                                                    case content_of prf of
+                                                        PTmSpec _ => (thy',th)
+                                                      | PTyDef  _ => (thy',th)
+                                                      | PTyIntro _ => (thy',th)
+                                                      | _ => ProofKernel.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 ProofKernel.get_axiom thyname name thy of
+                            (thy',SOME res) => (thy',res)
+                          | (thy',NONE) => ProofKernel.new_axiom name c thy')
+                  | PTmSpec(seg,names,prf') =>
+                    let
+                        val (thy',th) = rp' prf' thy
+                    in
+                        ProofKernel.new_specification seg thmname names th thy'
+                    end
+                  | PTyDef(seg,name,prf') =>
+                    let
+                        val (thy',th) = rp' prf' thy
+                    in
+                        ProofKernel.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
+                        ProofKernel.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
+                    end
+                  | _ => rp pc thy
+            end
+    in
+        rp' prf thy
+    end
+
+fun setup_int_thms thyname thy =
+    let
+        val fname =
+            case ProofKernel.get_proof_dir thyname thy of
+                SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
+              | NONE => error "Cannot find proof files"
+        val is = TextIO.openIn fname
+        val (num_int_thms,facts) =
+            let
+                fun get_facts facts =
+                    case TextIO.inputLine is of
+                        NONE => (case facts of
+                                   i::facts => (the (Int.fromString i),map ProofKernel.protect_factname (rev facts))
+                                 | _ => raise ERR "replay_thm" "Bad facts.lst file")
+                      | SOME 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 prf = mk_proof PDisk
+                val _ = set_disk_info_of prf thyname thmname
+                val _ = writeln ("Replaying "^thmname^" ...")
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in
+                p
+            end
+    in
+        replay_fact (thmname,thy)
+    end
+
+fun import_thms thyname int_thms thmnames thy =
+    let
+        fun replay_fact thmname thy = 
+            let
+                val prf = mk_proof PDisk        
+                val _ = set_disk_info_of prf thyname thmname
+                val _ = writeln ("Replaying "^thmname^" ...")
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in
+                p
+            end
+        val res_thy = fold replay_fact thmnames thy
+    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 prf = mk_proof PDisk        
+                val _ = set_disk_info_of prf thyname thmname        
+                val _ = writeln ("Replaying "^thmname^" ...")
+                val p = fst (replay_proof int_thms thyname thmname prf thy)
+            in 
+                p
+            end
+    in
+        replay_fact (thmname,thy)
+    end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/shuffler.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,585 @@
+(*  Title:      HOL/Import/shuffler.ML
+    Author:     Sebastian Skalberg, TU Muenchen
+
+Package for proving two terms equal by normalizing (hence the
+"shuffler" name).  Uses the simplifier for the normalization.
+*)
+
+signature Shuffler =
+sig
+    val debug      : bool Unsynchronized.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: Proof.context -> bool -> (string * thm) list -> int -> tactic
+    val shuffle_tac: Proof.context -> thm list -> int -> tactic
+    val search_tac : Proof.context -> int -> tactic
+
+    val print_shuffles: theory -> unit
+
+    val add_shuffle_rule: thm -> theory -> theory
+    val shuffle_attr: attribute
+
+    val setup      : theory -> theory
+end
+
+structure Shuffler : Shuffler =
+struct
+
+val debug = Unsynchronized.ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+val string_of_thm = Print_Mode.setmp [] Display.string_of_thm_without_context;
+
+structure ShuffleData = Theory_Data
+(
+  type T = thm list
+  val empty = []
+  val extend = I
+  val merge = Thm.merge_thms
+)
+
+fun print_shuffles thy =
+  Pretty.writeln (Pretty.big_list "Shuffle theorems:"
+    (map (Display.pretty_thm_global thy) (ShuffleData.get thy)))
+
+val weaken =
+    let
+        val cert = cterm_of Pure.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 cPQ = cert PQ
+        val cPPQ = cert PPQ
+        val th1 = Thm.assume cPQ |> implies_intr_list [cPQ,cP]
+        val th3 = Thm.assume cP
+        val th4 = implies_elim_list (Thm.assume cPPQ) [th3,th3]
+                                    |> implies_intr_list [cPPQ,cP]
+    in
+        Thm.equal_intr th4 th1 |> Drule.export_without_context
+    end
+
+val imp_comm =
+    let
+        val cert = cterm_of Pure.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 (Thm.assume cPQR) [Thm.assume cP,Thm.assume cQ]
+                                    |> implies_intr_list [cPQR,cQ,cP]
+        val th2 = implies_elim_list (Thm.assume cQPR) [Thm.assume cQ,Thm.assume cP]
+                                    |> implies_intr_list [cQPR,cP,cQ]
+    in
+        Thm.equal_intr th1 th2 |> Drule.export_without_context
+    end
+
+val all_comm =
+    let
+        val cert = cterm_of Pure.thy
+        val xT = TFree("'a",[])
+        val yT = TFree("'b",[])
+        val x = Free("x",xT)
+        val y = Free("y",yT)
+        val P = Free("P",xT-->yT-->propT)
+        val lhs = Logic.all x (Logic.all y (P $ x $ y))
+        val rhs = Logic.all y (Logic.all x (P $ x $ y))
+        val cl = cert lhs
+        val cr = cert rhs
+        val cx = cert x
+        val cy = cert y
+        val th1 = Thm.assume cr
+                         |> forall_elim_list [cy,cx]
+                         |> forall_intr_list [cx,cy]
+                         |> Thm.implies_intr cr
+        val th2 = Thm.assume cl
+                         |> forall_elim_list [cx,cy]
+                         |> forall_intr_list [cy,cx]
+                         |> Thm.implies_intr cl
+    in
+        Thm.equal_intr th1 th2 |> Drule.export_without_context
+    end
+
+val equiv_comm =
+    let
+        val cert = cterm_of Pure.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  = Thm.assume ctu |> Thm.symmetric |> Thm.implies_intr ctu
+        val th2  = Thm.assume cut |> Thm.symmetric |> Thm.implies_intr cut
+    in
+        Thm.equal_intr th1 th2 |> Drule.export_without_context
+    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 thy xv yv t =
+    let
+        val lhs = Logic.list_all ([xv,yv],t)
+        val rhs = Logic.list_all ([yv,xv],swap_bound 0 t)
+        val rew = Logic.mk_equals (lhs,rhs)
+        val init = Thm.trivial (cterm_of thy rew)
+    in
+        all_comm RS init
+    end
+
+fun quant_rewrite thy _ (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 thy (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    = Thm.reflexive (cterm_of thy t)
+                         val newt_th = Thm.reflexive (cterm_of thy newt)
+                     in
+                         SOME (Thm.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 = Misc_Legacy.term_tvars t
+        val tfree_names = Misc_Legacy.add_term_tfree_names(t,[])
+        val (type_inst,_) =
+            fold (fn (w as (v,_), S) => fn (inst, used) =>
+                      let
+                          val v' = singleton (Name.variant_list used) v
+                      in
+                          ((w,TFree(v',S))::inst,v'::used)
+                      end)
+                  tvars ([], tfree_names)
+        val t' = subst_TVars type_inst t
+    in
+        (t', map (fn (w,TFree(v,S)) => (v,TVar(w,S))
+                  | _ => error "Internal error in Shuffler.freeze_thaw") type_inst)
+    end
+
+fun inst_tfrees thy [] thm = thm
+  | inst_tfrees thy ((name,U)::rest) thm =
+    let
+        val cU = ctyp_of thy U
+        val tfrees = Misc_Legacy.add_term_tfrees (prop_of thm,[])
+        val (rens, thm') = Thm.varifyT_global'
+    (remove (op = o apsnd fst) name tfrees) thm
+        val mid =
+            case rens of
+                [] => thm'
+              | [((_, S), idx)] => Drule.instantiate_normalize
+            ([(ctyp_of thy (TVar (idx, S)), cU)], []) thm'
+              | _ => error "Shuffler.inst_tfrees internal error"
+    in
+        inst_tfrees thy 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 thy _ origt =
+    let
+        val (typet,Tinst) = freeze_thaw_term origt
+        val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
+        val final = inst_tfrees thy 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 (cat_lines
+                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+                      Syntax.string_of_term_global thy origt,
+                      Syntax.string_of_term_global thy lhs,
+                      Syntax.string_of_term_global thy typet,
+                      Syntax.string_of_term_global thy t] @
+                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
+                else ()
+            end
+    in
+        case t of
+            Const("all",_) $ (Abs(x,xT,Const("==",_) $ P $ Q)) =>
+            (if eta_redex P andalso eta_redex Q
+              then
+                  let
+                      val cert = cterm_of thy
+                      val v = Free (singleton (Name.variant_list (Term.add_free_names t [])) "v", xT)
+                      val cv = cert v
+                      val ct = cert t
+                      val th = (Thm.assume ct)
+                                   |> Thm.forall_elim cv
+                                   |> Thm.abstract_rule x cv
+                      val ext_th = Thm.eta_conversion (cert (Abs(x,xT,P)))
+                      val th' = Thm.transitive (Thm.symmetric ext_th) th
+                      val cu = cert (prop_of th')
+                      val uth = Thm.combination (Thm.assume cu) (Thm.reflexive cv)
+                      val uth' = (Thm.beta_conversion false (cert (Abs(x,xT,Q) $ v)))
+                                     |> Thm.transitive uth
+                                     |> Thm.forall_intr cv
+                                     |> Thm.implies_intr cu
+                      val rew_th = Thm.equal_intr (th' |> Thm.implies_intr ct) uth'
+                      val res = final rew_th
+                  in
+                       SOME res
+                  end
+              else NONE)
+          | _ => NONE
+       end
+
+fun eta_expand thy _ origt =
+    let
+        val (typet,Tinst) = freeze_thaw_term origt
+        val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
+        val final = inst_tfrees thy 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 (cat_lines
+                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+                      Syntax.string_of_term_global thy origt,
+                      Syntax.string_of_term_global thy lhs,
+                      Syntax.string_of_term_global thy typet,
+                      Syntax.string_of_term_global thy t] @
+                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
+                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,_]) =>
+                      let
+                          val cert = cterm_of thy
+                          val vname = singleton (Name.variant_list (Term.add_free_names t [])) "v"
+                          val v = Free(vname,aT)
+                          val cv = cert v
+                          val ct = cert t
+                          val th1 = (Thm.combination (Thm.assume ct) (Thm.reflexive cv))
+                                        |> Thm.forall_intr cv
+                                        |> Thm.implies_intr ct
+                          val concl = cert (concl_of th1)
+                          val th2 = (Thm.assume concl)
+                                        |> Thm.forall_elim cv
+                                        |> Thm.abstract_rule vname cv
+                          val (lhs,rhs) = Logic.dest_equals (prop_of th2)
+                          val elhs = Thm.eta_conversion (cert lhs)
+                          val erhs = Thm.eta_conversion (cert rhs)
+                          val th2' = Thm.transitive
+                                         (Thm.transitive (Thm.symmetric elhs) th2)
+                                         erhs
+                          val res = Thm.equal_intr th1 (th2' |> Thm.implies_intr concl)
+                          val res' = final res
+                      in
+                          SOME res'
+                      end
+                    | _ => NONE)
+            else NONE
+          | _ => error ("Bad eta_expand argument" ^ Syntax.string_of_term_global thy t)
+    end;
+
+fun mk_tfree s = TFree("'"^s,[])
+fun mk_free s t = Free (s,t)
+val xT = mk_tfree "a"
+val yT = mk_tfree "b"
+val x = Free ("x", xT)
+val y = Free ("y", yT)
+val P  = mk_free "P" (xT-->yT-->propT)
+val Q  = mk_free "Q" (xT-->yT)
+val R  = mk_free "R" (xT-->yT)
+in
+
+fun quant_simproc thy = Simplifier.simproc_global_i
+                           thy
+                           "Ordered rewriting of nested quantifiers"
+                           [Logic.all x (Logic.all y (P $ x $ y))]
+                           quant_rewrite
+fun eta_expand_simproc thy = Simplifier.simproc_global_i
+                         thy
+                         "Smart eta-expansion by equivalences"
+                         [Logic.mk_equals(Q,R)]
+                         eta_expand
+fun eta_contract_simproc thy = Simplifier.simproc_global_i
+                         thy
+                         "Smart handling of eta-contractions"
+                         [Logic.all x (Logic.mk_equals (Q $ x, R $ x))]
+                         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 thy 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(_,xT,t),idx) =
+            let
+                val x' = "x" ^ string_of_int idx
+                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 thy t
+        val ct' = cterm_of thy t'
+        val res = Thm.transitive (Thm.reflexive ct) (Thm.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 norms = ShuffleData.get thy
+        val ss = Simplifier.global_context thy empty_ss
+          addsimps (map (Thm.transfer thy) norms)
+          addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy]
+        fun chain f th =
+            let
+                val rhs = Thm.rhs_of th
+            in
+                Thm.transitive th (f rhs)
+            end
+        val th =
+            t |> disamb_bound thy
+              |> chain (Simplifier.full_rewrite ss)
+              |> chain Thm.eta_conversion
+              |> Thm.strip_shyps
+        val _ = message ("norm_term: " ^ (string_of_thm th))
+    in
+        th
+    end
+
+
+(* Closes a theorem with respect to free and schematic variables (does
+not touch type variables, though). *)
+
+fun close_thm th =
+    let
+        val thy = Thm.theory_of_thm th
+        val c = prop_of th
+        val vars = Misc_Legacy.add_term_frees (c, Misc_Legacy.add_term_vars(c,[]))
+    in
+        Drule.forall_intr_list (map (cterm_of thy) vars) th
+    end
+
+
+(* Normalizes a theorem's conclusion using norm_term. *)
+
+fun norm_thm thy th =
+    let
+        val c = prop_of th
+    in
+        Thm.equal_elim (norm_term thy c) th
+    end
+
+(* make_equal thy t u tries to construct the theorem t == u under the
+signature thy.  If it succeeds, SOME (t == u) is returned, otherwise
+NONE is returned. *)
+
+fun make_equal thy t u =
+    let
+        val t_is_t' = norm_term thy t
+        val u_is_u' = norm_term thy u
+        val th = Thm.transitive t_is_t' (Thm.symmetric u_is_u')
+        val _ = message ("make_equal: SOME " ^ (string_of_thm th))
+    in
+        SOME th
+    end
+    handle THM _ => (message "make_equal: NONE";NONE)
+
+fun match_consts ignore t (* th *) =
+    let
+        fun add_consts (Const (c, _), cs) =
+            if member (op =) ignore c
+            then cs
+            else insert (op =) c 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) =>
+        let
+            val th_consts = add_consts(prop_of th,[])
+        in
+            eq_set (op =) (t_consts, th_consts)
+        end
+    end
+
+val collect_ignored = fold_rev (fn thm => fn cs =>
+  let
+    val (lhs, rhs) = Logic.dest_equals (prop_of thm);
+    val consts_lhs = Term.add_const_names lhs [];
+    val consts_rhs = Term.add_const_names rhs [];
+    val ignore_lhs = subtract (op =) consts_rhs consts_lhs;
+    val ignore_rhs = subtract (op =) consts_lhs consts_rhs;
+  in
+    fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
+  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 vars = Misc_Legacy.add_term_frees (t, Misc_Legacy.add_term_vars (t,[]))
+        val closed_t = fold_rev Logic.all vars t
+        val rew_th = norm_term thy closed_t
+        val rhs = Thm.rhs_of rew_th
+
+        fun process [] = NONE
+          | process ((name,th)::thms) =
+            let
+                val norm_th = Thm.varifyT_global (norm_thm thy (close_thm (Thm.transfer thy th)))
+                val triv_th = Thm.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 (Thm.bicompose false (*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 = Thm.equal_elim (Thm.symmetric rew_th) mod_th
+                    in
+                        message ("Shuffler.set_prop succeeded by " ^ name);
+                        SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)
+                    end
+                  | NONE => process thms
+            end
+            handle THM _ => process thms
+    in
+        fn thms =>
+           case process thms of
+               res as SOME (_,th) => if (prop_of th) aconv t
+                                        then res
+                                        else error "Internal error in set_prop"
+             | NONE => NONE
+    end
+
+fun find_potential thy t =
+    let
+        val shuffles = ShuffleData.get thy
+        val ignored = collect_ignored shuffles []
+        val all_thms =
+          map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (Global_Theory.facts_of thy)))
+    in
+        filter (match_consts ignored t) all_thms
+    end
+
+fun gen_shuffle_tac ctxt search thms = SUBGOAL (fn (t, i) =>
+    let
+        val thy = Proof_Context.theory_of ctxt
+        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)
+    end)
+
+fun shuffle_tac ctxt thms =
+  gen_shuffle_tac ctxt false (map (pair "") thms);
+
+fun search_tac ctxt =
+  gen_shuffle_tac ctxt true (map (pair "premise") (Assumption.all_prems_of ctxt));
+
+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
+
+val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I);
+
+val setup =
+  Method.setup @{binding shuffle_tac}
+    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (shuffle_tac ctxt ths)))
+    "solve goal by shuffling terms around" #>
+  Method.setup @{binding search_tac}
+    (Scan.succeed (SIMPLE_METHOD' o search_tac)) "search for suitable theorems" #>
+  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.setup @{binding shuffle_rule} (Scan.succeed shuffle_attr) "declare rule for shuffler";
+
+end
--- a/src/HOL/Import/HOL_Light/Compatibility.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,350 +0,0 @@
-(*  Title:      HOL/Import/HOL_Light/Compatibility.thy
-    Author:     Steven Obua and Sebastian Skalberg, TU Muenchen
-    Author:     Cezary Kaliszyk
-*)
-
-theory Compatibility
-imports Main Fact Parity "~~/src/HOL/Library/Infinite_Set"
-  HOLLightList HOLLightReal HOLLightInt "~~/src/HOL/Import/Importer"
-begin
-
-(* list *)
-lemmas [import_rew] = list_el_def member_def list_mem_def
-(* int *)
-lemmas [import_rew] = int_coprime.simps int_gcd.simps hl_mod_def hl_div_def int_mod_def eqeq_def
-(* real *)
-lemma [import_rew]:
-  "real (0::nat) = 0" "real (1::nat) = 1" "real (2::nat) = 2"
-  by simp_all
-
-lemma one:
-  "\<forall>v. v = ()"
-  by simp
-
-lemma num_Axiom:
-  "\<exists>!fn. fn 0 = e \<and> (\<forall>n. fn (Suc n) = f (fn n) n)"
-  apply (rule ex1I[of _ "nat_rec e (%n e. f e n)"])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma SUC_INJ:
-  "\<forall>m n. Suc m = Suc n \<longleftrightarrow> m = n"
-  by simp
-
-lemma PAIR:
-  "(fst x, snd x) = x"
-  by simp
-
-lemma EXISTS_UNIQUE_THM:
-  "(Ex1 P) = (Ex P & (\<forall>x y. P x & P y --> (x = y)))"
-  by auto
-
-lemma DEF__star_:
-  "op * = (SOME mult. (\<forall>n. mult 0 n = 0) \<and> (\<forall>m n. mult (Suc m) n = mult m n + n))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply auto
-  done
-
-lemma DEF__slash__backslash_:
-  "(t1 \<and> t2) = ((\<lambda>f. f t1 t2 :: bool) = (\<lambda>f. f True True))"
-  unfolding fun_eq_iff
-  by (intro iffI, simp_all) (erule allE[of _ "(%a b. a \<and> b)"], simp)
-
-lemma DEF__lessthan__equal_:
-  "op \<le> = (SOME u. (\<forall>m. u m 0 = (m = 0)) \<and> (\<forall>m n. u m (Suc n) = (m = Suc n \<or> u m n)))"
-  apply (rule some_equality[symmetric])
-  apply auto[1]
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply auto
-  done
-
-lemma DEF__lessthan_:
-  "op < = (SOME u. (\<forall>m. u m 0 = False) \<and> (\<forall>m n. u m (Suc n) = (m = n \<or> u m n)))"
-  apply (rule some_equality[symmetric])
-  apply auto[1]
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply auto
-  done
-
-lemma DEF__greaterthan__equal_:
-  "(op \<ge>) = (%u ua. ua \<le> u)"
-  by (simp)
-
-lemma DEF__greaterthan_:
-  "(op >) = (%u ua. ua < u)"
-  by (simp)
-
-lemma DEF__equal__equal__greaterthan_:
-  "(t1 \<longrightarrow> t2) = ((t1 \<and> t2) = t1)"
-  by auto
-
-lemma DEF_WF:
-  "wfP = (\<lambda>u. \<forall>P. (\<exists>x. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))"
-  unfolding fun_eq_iff
-proof (intro allI iffI impI wfI_min[to_pred], elim exE wfE_min[to_pred])
-  fix P :: "'a \<Rightarrow> bool" and xa :: "'a"
-  assume "P xa"
-  then show "xa \<in> Collect P" by simp
-next
-  fix x P xa z
-  assume "P xa" "z \<in> {a. P a}" "\<And>y. x y z \<Longrightarrow> y \<notin> {a. P a}"
-  then show "\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y)" by auto
-next
-  fix x :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and xa :: "'a" and Q
-  assume a: "xa \<in> Q"
-  assume b: "\<forall>P. Ex P \<longrightarrow> (\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y))"
-  then have "Ex (\<lambda>x. x \<in> Q) \<longrightarrow> (\<exists>xa. (\<lambda>x. x \<in> Q) xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> (\<lambda>x. x \<in> Q) y))" by auto
-  then show "\<exists>z\<in>Q. \<forall>y. x y z \<longrightarrow> y \<notin> Q" using a by auto
-qed
-
-lemma DEF_UNIV: "top = (%x. True)"
-  by (rule ext) (simp add: top1I)
-
-lemma DEF_UNIONS:
-  "Sup = (\<lambda>u. {ua. \<exists>x. (\<exists>ua. ua \<in> u \<and> x \<in> ua) \<and> ua = x})"
-  by (auto simp add: Union_eq)
-
-lemma DEF_UNION:
-  "op \<union> = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<or> x \<in> ua) \<and> ub = x})"
-  by auto
-
-lemma DEF_SUBSET: "op \<subseteq> = (\<lambda>u ua. \<forall>x. x \<in> u \<longrightarrow> x \<in> ua)"
-  by (metis set_rev_mp subsetI)
-
-lemma DEF_SND:
-  "snd = (\<lambda>p. SOME x. EX y. p = (y, x))"
-  unfolding fun_eq_iff
-  by (rule someI2) (auto intro: snd_conv[symmetric] someI2)
-
-definition [simp, import_rew]: "SETSPEC x P y \<longleftrightarrow> P & x = y"
-
-lemma DEF_PSUBSET: "op \<subset> = (\<lambda>u ua. u \<subseteq> ua & u \<noteq> ua)"
-  by (metis psubset_eq)
-
-definition [import_rew]: "Pred n = n - (Suc 0)"
-
-lemma DEF_PRE: "Pred = (SOME PRE. PRE 0 = 0 & (\<forall>n. PRE (Suc n) = n))"
-  apply (rule some_equality[symmetric])
-  apply (simp add: Pred_def)
-  apply (rule ext)
-  apply (induct_tac x)
-  apply (auto simp add: Pred_def)
-  done
-
-lemma DEF_ONE_ONE:
-  "inj = (\<lambda>u. \<forall>x1 x2. u x1 = u x2 \<longrightarrow> x1 = x2)"
-  by (simp add: inj_on_def)
-
-definition ODD'[import_rew]: "(ODD :: nat \<Rightarrow> bool) = odd"
-
-lemma DEF_ODD:
-  "odd = (SOME ODD. ODD 0 = False \<and> (\<forall>n. ODD (Suc n) = (\<not> ODD n)))"
-  apply (rule some_equality[symmetric])
-  apply simp
-  unfolding fun_eq_iff
-  apply (intro allI)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-definition [import_rew, simp]: "NUMERAL (x :: nat) = x"
-
-lemma DEF_MOD:
-  "op mod = (SOME r. \<forall>m n. if n = (0 :: nat) then m div n = 0 \<and>
-     r m n = m else m = m div n * n + r m n \<and> r m n < n)"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (case_tac "xa = 0")
-  apply auto
-  apply (drule_tac x="x" in spec)
-  apply (drule_tac x="xa" in spec)
-  apply auto
-  by (metis mod_less mod_mult_self2 nat_add_commute nat_mult_commute)
-
-definition "MEASURE = (%u x y. (u x :: nat) < u y)"
-
-lemma [import_rew]:
-  "MEASURE u = (%a b. (a, b) \<in> measure u)"
-  unfolding MEASURE_def measure_def
-  by simp
-
-definition
-  "LET f s = f s"
-
-lemma [import_rew]:
-  "LET f s = Let s f"
-  by (simp add: LET_def Let_def)
-
-lemma DEF_INTERS:
-  "Inter = (\<lambda>u. {ua. \<exists>x. (\<forall>ua. ua \<in> u \<longrightarrow> x \<in> ua) \<and> ua = x})"
-  by auto
-
-lemma DEF_INTER:
-  "op \<inter> = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<and> x \<in> ua) \<and> ub = x})"
-  by auto
-
-lemma DEF_INSERT:
-  "insert = (\<lambda>u ua. {y. y \<in> ua | y = u})"
-  by auto
-
-lemma DEF_IMAGE:
-  "op ` = (\<lambda>u ua. {ub. \<exists>y. (\<exists>x. x \<in> ua \<and> y = u x) \<and> ub = y})"
-  by (simp add: fun_eq_iff image_def Bex_def)
-
-lemma DEF_GEQ:
-  "(op =) = (op =)"
-  by simp
-
-lemma DEF_GABS:
-  "Eps = Eps"
-  by simp
-
-lemma DEF_FST:
-  "fst = (%p. SOME x. EX y. p = (x, y))"
-  unfolding fun_eq_iff
-  by (rule someI2) (auto intro: fst_conv[symmetric] someI2)
-
-lemma DEF_FINITE:
-  "finite = (\<lambda>a. \<forall>FP. (\<forall>a. a = {} \<or> (\<exists>x s. a = insert x s \<and> FP s) \<longrightarrow> FP a) \<longrightarrow> FP a)"
-  unfolding fun_eq_iff
-  apply (intro allI iffI impI)
-  apply (erule finite_induct)
-  apply auto[2]
-  apply (drule_tac x="finite" in spec)
-  by (metis finite_insert infinite_imp_nonempty infinite_super predicate1I)
-
-lemma DEF_FACT:
-  "fact = (SOME FACT. FACT 0 = 1 & (\<forall>n. FACT (Suc n) = Suc n * FACT n))"
-  apply (rule some_equality[symmetric])
-  apply (simp_all)
-  unfolding fun_eq_iff
-  apply (intro allI)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_EXP:
-  "power = (SOME EXP. (\<forall>m. EXP m 0 = 1) \<and> (\<forall>m n. EXP m (Suc n) = m * EXP m n))"
-  apply (rule some_equality[symmetric])
-  apply (simp_all)
-  unfolding fun_eq_iff
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply simp_all
-  done
-
-lemma DEF_EVEN:
-  "even = (SOME EVEN. EVEN 0 = True \<and> (\<forall>n. EVEN (Suc n) = (\<not> EVEN n)))"
-  apply (rule some_equality[symmetric])
-  apply simp
-  unfolding fun_eq_iff
-  apply (intro allI)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_EMPTY: "bot = (%x. False)"
-  by (rule ext) auto
-  
-lemma DEF_DIV:
-  "op div = (SOME q. \<exists>r. \<forall>m n. if n = (0 :: nat) then q m n = 0 \<and> r m n = m
-     else m = q m n * n + r m n \<and> r m n < n)"
-  apply (rule some_equality[symmetric])
-  apply (rule_tac x="op mod" in exI)
-  apply (auto simp add: fun_eq_iff)
-  apply (case_tac "xa = 0")
-  apply auto
-  apply (drule_tac x="x" in spec)
-  apply (drule_tac x="xa" in spec)
-  apply auto
-  by (metis div_mult_self2 gr_implies_not0 mod_div_trivial mod_less
-      nat_add_commute nat_mult_commute plus_nat.add_0)
-
-definition [import_rew]: "DISJOINT a b \<longleftrightarrow> a \<inter> b = {}"
-
-lemma DEF_DISJOINT:
-  "DISJOINT = (%u ua. u \<inter> ua = {})"
-  by (auto simp add: DISJOINT_def [abs_def])
-
-lemma DEF_DIFF:
-  "op - = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<and> x \<notin> ua) \<and> ub = x})"
-  by (metis set_diff_eq)
-
-definition [import_rew]: "DELETE s e = s - {e}"
-
-lemma DEF_DELETE:
-  "DELETE = (\<lambda>u ua. {ub. \<exists>y. (y \<in> u \<and> y \<noteq> ua) \<and> ub = y})"
-  by (auto simp add: DELETE_def [abs_def])
-
-lemma COND_DEF:
-  "(if b then t else f) = (SOME x. (b = True \<longrightarrow> x = t) \<and> (b = False \<longrightarrow> x = f))"
-  by auto
-
-definition [simp]: "NUMERAL_BIT1 n = n + (n + Suc 0)"
-
-lemma BIT1_DEF:
-  "NUMERAL_BIT1 = (%u. Suc (2 * u))"
-  by (auto)
-
-definition [simp]: "NUMERAL_BIT0 (n :: nat) = n + n"
-
-lemma BIT0_DEF:
-  "NUMERAL_BIT0 = (SOME BIT0. BIT0 0 = 0 \<and> (\<forall>n. BIT0 (Suc n) = Suc (Suc (BIT0 n))))"
-  apply (rule some_equality[symmetric])
-  apply auto
-  apply (rule ext)
-  apply (induct_tac x)
-  apply auto
-  done
-
-lemma [import_rew]:
-  "NUMERAL_BIT0 n = 2 * n"
-  "NUMERAL_BIT1 n = 2 * n + 1"
-  "2 * 0 = (0 :: nat)"
-  "2 * 1 = (2 :: nat)"
-  "0 + 1 = (1 :: nat)"
-  by simp_all
-
-lemma DEF_MINUS: "op - = (SOME sub. (\<forall>m. sub m 0 = m) & (\<forall>m n. sub m (Suc n) = sub m n - Suc 0))"
-  apply (rule some_equality[symmetric])
-  apply auto
-  apply (rule ext)+
-  apply (induct_tac xa)
-  apply auto
-  done
-
-lemma DEF_PLUS: "op + = (SOME add. (\<forall>n. add 0 n = n) & (\<forall>m n. add (Suc m) n = Suc (add m n)))"
-  apply (rule some_equality[symmetric])
-  apply auto
-  apply (rule ext)+
-  apply (induct_tac x)
-  apply auto
-  done
-
-lemmas [import_rew] = id_apply
-
-lemma DEF_CHOICE: "Eps = (%u. SOME x. u x)"
-  by simp
-
-definition dotdot :: "nat => nat => nat set"
-  where "dotdot u ua = {ub. EX x. (u <= x & x <= ua) & ub = x}"
-
-lemma [import_rew]: "dotdot a b = {a..b}"
-  unfolding fun_eq_iff atLeastAtMost_def atLeast_def atMost_def dotdot_def
-  by (simp add: Collect_conj_eq)
-
-definition [import_rew,simp]: "INFINITE S \<longleftrightarrow> \<not> finite S"
-
-lemma DEF_INFINITE: "INFINITE = (\<lambda>u. \<not>finite u)"
-  by (simp add: INFINITE_def [abs_def])
-
-end
-
--- a/src/HOL/Import/HOL_Light/Generate.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-theory Generate
-imports "Template/GenHOLLight"
-begin
-
-end
--- a/src/HOL/Import/HOL_Light/Generated/HOLLight.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,7472 +0,0 @@
-(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)
-
-theory HOLLight
-imports "../../Importer" "../Compatibility"
-begin 
-
-setup_theory "~~/src/HOL/Import/HOL_Light/Generated" hollight
-
-consts
-  "_FALSITY_" :: "bool" ("'_FALSITY'_")
-
-defs
-  "_FALSITY__def": "_FALSITY_ == False"
-
-lemma DEF__FALSITY_: "_FALSITY_ = False"
-  by (import hollight DEF__FALSITY_)
-
-lemma CONJ_ACI: "(p & q) = (q & p) &
-((p & q) & r) = (p & q & r) &
-(p & q & r) = (q & p & r) & (p & p) = p & (p & p & q) = (p & q)"
-  by (import hollight CONJ_ACI)
-
-lemma DISJ_ACI: "(p | q) = (q | p) &
-((p | q) | r) = (p | q | r) &
-(p | q | r) = (q | p | r) & (p | p) = p & (p | p | q) = (p | q)"
-  by (import hollight DISJ_ACI)
-
-lemma IMP_CONJ_ALT: "(p & q --> r) = (q --> p --> r)"
-  by (import hollight IMP_CONJ_ALT)
-
-lemma EQ_CLAUSES: "(True = t) = t & (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
-  by (import hollight EQ_CLAUSES)
-
-lemma NOT_CLAUSES_WEAK: "(~ True) = False & (~ False) = True"
-  by (import hollight NOT_CLAUSES_WEAK)
-
-lemma AND_CLAUSES: "(True & t) = t &
-(t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
-  by (import hollight AND_CLAUSES)
-
-lemma OR_CLAUSES: "(True | t) = True &
-(t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
-  by (import hollight OR_CLAUSES)
-
-lemma IMP_CLAUSES: "(True --> t) = t &
-(t --> True) = True &
-(False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
-  by (import hollight IMP_CLAUSES)
-
-lemma IMP_EQ_CLAUSE: "((x::'q_851) = x --> (p::bool)) = p"
-  by (import hollight IMP_EQ_CLAUSE)
-
-lemma TRIV_EXISTS_AND_THM: "(EX x::'A. (P::bool) & (Q::bool)) = ((EX x::'A. P) & (EX x::'A. Q))"
-  by (import hollight TRIV_EXISTS_AND_THM)
-
-lemma TRIV_AND_EXISTS_THM: "((EX x::'A. (P::bool)) & (EX x::'A. (Q::bool))) = (EX x::'A. P & Q)"
-  by (import hollight TRIV_AND_EXISTS_THM)
-
-lemma TRIV_FORALL_OR_THM: "(ALL x::'A. (P::bool) | (Q::bool)) = ((ALL x::'A. P) | (ALL x::'A. Q))"
-  by (import hollight TRIV_FORALL_OR_THM)
-
-lemma TRIV_OR_FORALL_THM: "((ALL x::'A. (P::bool)) | (ALL x::'A. (Q::bool))) = (ALL x::'A. P | Q)"
-  by (import hollight TRIV_OR_FORALL_THM)
-
-lemma TRIV_FORALL_IMP_THM: "(ALL x::'A. (P::bool) --> (Q::bool)) = ((EX x::'A. P) --> (ALL x::'A. Q))"
-  by (import hollight TRIV_FORALL_IMP_THM)
-
-lemma TRIV_EXISTS_IMP_THM: "(EX x::'A. (P::bool) --> (Q::bool)) = ((ALL x::'A. P) --> (EX x::'A. Q))"
-  by (import hollight TRIV_EXISTS_IMP_THM)
-
-lemma EXISTS_UNIQUE_ALT: "Ex1 (P::'A => bool) = (EX x::'A. ALL y::'A. P y = (x = y))"
-  by (import hollight EXISTS_UNIQUE_ALT)
-
-lemma SELECT_UNIQUE: "(!!y::'A. (P::'A => bool) y = (y = (x::'A))) ==> Eps P = x"
-  by (import hollight SELECT_UNIQUE)
-
-lemma EXCLUDED_MIDDLE: "t | ~ t"
-  by (import hollight EXCLUDED_MIDDLE)
-
-lemma COND_CLAUSES: "(if True then x::'A else (xa::'A)) = x & (if False then x else xa) = xa"
-  by (import hollight COND_CLAUSES)
-
-lemma COND_EXPAND: "(if b then t1 else t2) = ((~ b | t1) & (b | t2))"
-  by (import hollight COND_EXPAND)
-
-lemma COND_RATOR: "(if b::bool then f::'A => 'B else (g::'A => 'B)) (x::'A) =
-(if b then f x else g x)"
-  by (import hollight COND_RATOR)
-
-lemma COND_ABS: "(%x::'A. if b::bool then (f::'A => 'B) x else (g::'A => 'B) x) =
-(if b then f else g)"
-  by (import hollight COND_ABS)
-
-lemma MONO_COND: "[| (A --> B) & (C --> D); if b then A else C |] ==> if b then B else D"
-  by (import hollight MONO_COND)
-
-lemma SKOLEM_THM: "(ALL x::'A. Ex ((P::'A => 'B => bool) x)) =
-(EX x::'A => 'B. ALL xa::'A. P xa (x xa))"
-  by (import hollight SKOLEM_THM)
-
-lemma UNIQUE_SKOLEM_ALT: "(ALL x::'A. Ex1 ((P::'A => 'B => bool) x)) =
-(EX f::'A => 'B. ALL (x::'A) y::'B. P x y = (f x = y))"
-  by (import hollight UNIQUE_SKOLEM_ALT)
-
-lemma COND_EQ_CLAUSE: "(if (x::'q_2963) = x then y::'q_2956 else (z::'q_2956)) = y"
-  by (import hollight COND_EQ_CLAUSE)
-
-lemma bool_RECURSION: "EX x::bool => 'A. x False = (a::'A) & x True = (b::'A)"
-  by (import hollight bool_RECURSION)
-
-lemma o_ASSOC: "(f::'C => 'D) o ((g::'B => 'C) o (h::'A => 'B)) = f o g o h"
-  by (import hollight o_ASSOC)
-
-lemma I_O_ID: "id o (f::'A => 'B) = f & f o id = f"
-  by (import hollight I_O_ID)
-
-lemma EXISTS_ONE_REP: "EX x. x"
-  by (import hollight EXISTS_ONE_REP)
-
-lemma one_axiom: "(f::'A => unit) = (x::'A => unit)"
-  by (import hollight one_axiom)
-
-lemma one_RECURSION: "EX x::unit => 'A. x () = (e::'A)"
-  by (import hollight one_RECURSION)
-
-lemma one_Axiom: "EX! fn::unit => 'A. fn () = (e::'A)"
-  by (import hollight one_Axiom)
-
-lemma th_cond: "(b = False --> x = x0) & (b = True --> x = x1) ==> x = (b & x1 | ~ b & x0)"
-  by (import hollight th_cond)
-
-definition
-  LET_END :: "'A => 'A"  where
-  "LET_END == %t::'A. t"
-
-lemma DEF_LET_END: "LET_END = (%t::'A. t)"
-  by (import hollight DEF_LET_END)
-
-consts
-  "_SEQPATTERN" :: "('q_4007 => 'q_4004 => bool)
-=> ('q_4007 => 'q_4004 => bool) => 'q_4007 => 'q_4004 => bool" ("'_SEQPATTERN")
-
-defs
-  "_SEQPATTERN_def": "_SEQPATTERN ==
-%(r::'q_4007 => 'q_4004 => bool) (s::'q_4007 => 'q_4004 => bool) x::'q_4007.
-   if Ex (r x) then r x else s x"
-
-lemma DEF__SEQPATTERN: "_SEQPATTERN =
-(%(r::'q_4007 => 'q_4004 => bool) (s::'q_4007 => 'q_4004 => bool)
-    x::'q_4007. if Ex (r x) then r x else s x)"
-  by (import hollight DEF__SEQPATTERN)
-
-consts
-  "_UNGUARDED_PATTERN" :: "bool => bool => bool" ("'_UNGUARDED'_PATTERN")
-
-defs
-  "_UNGUARDED_PATTERN_def": "_UNGUARDED_PATTERN == op &"
-
-lemma DEF__UNGUARDED_PATTERN: "_UNGUARDED_PATTERN = op &"
-  by (import hollight DEF__UNGUARDED_PATTERN)
-
-consts
-  "_GUARDED_PATTERN" :: "bool => bool => bool => bool" ("'_GUARDED'_PATTERN")
-
-defs
-  "_GUARDED_PATTERN_def": "_GUARDED_PATTERN == %p g r. p & g & r"
-
-lemma DEF__GUARDED_PATTERN: "_GUARDED_PATTERN = (%p g r. p & g & r)"
-  by (import hollight DEF__GUARDED_PATTERN)
-
-consts
-  "_MATCH" :: "'q_4049 => ('q_4049 => 'q_4053 => bool) => 'q_4053" ("'_MATCH")
-
-defs
-  "_MATCH_def": "_MATCH ==
-%(e::'q_4049) r::'q_4049 => 'q_4053 => bool.
-   if Ex1 (r e) then Eps (r e) else SOME z::'q_4053. False"
-
-lemma DEF__MATCH: "_MATCH =
-(%(e::'q_4049) r::'q_4049 => 'q_4053 => bool.
-    if Ex1 (r e) then Eps (r e) else SOME z::'q_4053. False)"
-  by (import hollight DEF__MATCH)
-
-consts
-  "_FUNCTION" :: "('q_4071 => 'q_4075 => bool) => 'q_4071 => 'q_4075" ("'_FUNCTION")
-
-defs
-  "_FUNCTION_def": "_FUNCTION ==
-%(r::'q_4071 => 'q_4075 => bool) x::'q_4071.
-   if Ex1 (r x) then Eps (r x) else SOME z::'q_4075. False"
-
-lemma DEF__FUNCTION: "_FUNCTION =
-(%(r::'q_4071 => 'q_4075 => bool) x::'q_4071.
-    if Ex1 (r x) then Eps (r x) else SOME z::'q_4075. False)"
-  by (import hollight DEF__FUNCTION)
-
-lemma PAIR_EXISTS_THM: "EX (x::'A => 'B => bool) (a::'A) b::'B. x = Pair_Rep a b"
-  by (import hollight PAIR_EXISTS_THM)
-
-lemma pair_RECURSION: "EX x::'A * 'B => 'C.
-   ALL (a0::'A) a1::'B. x (a0, a1) = (PAIR'::'A => 'B => 'C) a0 a1"
-  by (import hollight pair_RECURSION)
-
-definition
-  UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C"  where
-  "UNCURRY == %(u::'A => 'B => 'C) ua::'A * 'B. u (fst ua) (snd ua)"
-
-lemma DEF_UNCURRY: "UNCURRY = (%(u::'A => 'B => 'C) ua::'A * 'B. u (fst ua) (snd ua))"
-  by (import hollight DEF_UNCURRY)
-
-definition
-  PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D"  where
-  "PASSOC ==
-%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C.
-   u ((fst ua, fst (snd ua)), snd (snd ua))"
-
-lemma DEF_PASSOC: "PASSOC =
-(%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C.
-    u ((fst ua, fst (snd ua)), snd (snd ua)))"
-  by (import hollight DEF_PASSOC)
-
-lemma LAMBDA_PAIR_THM: "(x::'q_4547 * 'q_4546 => 'q_4539) =
-(SOME f::'q_4547 * 'q_4546 => 'q_4539.
-    ALL (xa::'q_4547) y::'q_4546. f (xa, y) = x (xa, y))"
-  by (import hollight LAMBDA_PAIR_THM)
-
-lemma FORALL_PAIRED_THM: "All (SOME f::'q_4576 * 'q_4575 => bool.
-        ALL (x::'q_4576) y::'q_4575.
-           f (x, y) = (P::'q_4576 => 'q_4575 => bool) x y) =
-(ALL x::'q_4576. All (P x))"
-  by (import hollight FORALL_PAIRED_THM)
-
-lemma EXISTS_PAIRED_THM: "Ex (SOME f::'q_4612 * 'q_4611 => bool.
-       ALL (x::'q_4612) y::'q_4611.
-          f (x, y) = (P::'q_4612 => 'q_4611 => bool) x y) =
-(EX x::'q_4612. Ex (P x))"
-  by (import hollight EXISTS_PAIRED_THM)
-
-lemma FORALL_TRIPLED_THM: "All (SOME f::'q_4649 * 'q_4648 * 'q_4647 => bool.
-        ALL (x::'q_4649) (y::'q_4648) z::'q_4647.
-           f (x, y, z) = (P::'q_4649 => 'q_4648 => 'q_4647 => bool) x y z) =
-(ALL (x::'q_4649) y::'q_4648. All (P x y))"
-  by (import hollight FORALL_TRIPLED_THM)
-
-lemma EXISTS_TRIPLED_THM: "Ex (SOME f::'q_4695 * 'q_4694 * 'q_4693 => bool.
-       ALL (x::'q_4695) (y::'q_4694) z::'q_4693.
-          f (x, y, z) = (P::'q_4695 => 'q_4694 => 'q_4693 => bool) x y z) =
-(EX (x::'q_4695) y::'q_4694. Ex (P x y))"
-  by (import hollight EXISTS_TRIPLED_THM)
-
-lemma IND_SUC_0_EXISTS: "EX (x::ind => ind) z::ind.
-   (ALL (x1::ind) x2::ind. (x x1 = x x2) = (x1 = x2)) &
-   (ALL xa::ind. x xa ~= z)"
-  by (import hollight IND_SUC_0_EXISTS)
-
-definition
-  IND_SUC :: "ind => ind"  where
-  "IND_SUC ==
-SOME f. EX z. (ALL x1 x2. (f x1 = f x2) = (x1 = x2)) & (ALL x. f x ~= z)"
-
-lemma DEF_IND_SUC: "IND_SUC =
-(SOME f. EX z. (ALL x1 x2. (f x1 = f x2) = (x1 = x2)) & (ALL x. f x ~= z))"
-  by (import hollight DEF_IND_SUC)
-
-definition
-  IND_0 :: "ind"  where
-  "IND_0 ==
-SOME z.
-   (ALL x1 x2. (IND_SUC x1 = IND_SUC x2) = (x1 = x2)) &
-   (ALL x. IND_SUC x ~= z)"
-
-lemma DEF_IND_0: "IND_0 =
-(SOME z.
-    (ALL x1 x2. (IND_SUC x1 = IND_SUC x2) = (x1 = x2)) &
-    (ALL x. IND_SUC x ~= z))"
-  by (import hollight DEF_IND_0)
-
-definition
-  NUM_REP :: "ind => bool"  where
-  "NUM_REP ==
-%a. ALL NUM_REP'.
-       (ALL a.
-           a = IND_0 | (EX i. a = IND_SUC i & NUM_REP' i) -->
-           NUM_REP' a) -->
-       NUM_REP' a"
-
-lemma DEF_NUM_REP: "NUM_REP =
-(%a. ALL NUM_REP'.
-        (ALL a.
-            a = IND_0 | (EX i. a = IND_SUC i & NUM_REP' i) -->
-            NUM_REP' a) -->
-        NUM_REP' a)"
-  by (import hollight DEF_NUM_REP)
-
-lemma num_RECURSION_STD: "EX fn::nat => 'Z.
-   fn (0::nat) = (e::'Z) &
-   (ALL n::nat. fn (Suc n) = (f::nat => 'Z => 'Z) n (fn n))"
-  by (import hollight num_RECURSION_STD)
-
-lemma ADD_CLAUSES: "(ALL x::nat. (0::nat) + x = x) &
-(ALL x::nat. x + (0::nat) = x) &
-(ALL (x::nat) xa::nat. Suc x + xa = Suc (x + xa)) &
-(ALL (x::nat) xa::nat. x + Suc xa = Suc (x + xa))"
-  by (import hollight ADD_CLAUSES)
-
-lemma ADD_AC: "(m::nat) + (n::nat) = n + m &
-m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)"
-  by (import hollight ADD_AC)
-
-lemma EQ_ADD_LCANCEL_0: "((m::nat) + (n::nat) = m) = (n = (0::nat))"
-  by (import hollight EQ_ADD_LCANCEL_0)
-
-lemma EQ_ADD_RCANCEL_0: "((x::nat) + (xa::nat) = xa) = (x = (0::nat))"
-  by (import hollight EQ_ADD_RCANCEL_0)
-
-lemma BIT1: "2 * x + 1 = Suc (x + x)"
-  by (import hollight BIT1)
-
-lemma BIT1_THM: "2 * x + 1 = Suc (x + x)"
-  by (import hollight BIT1_THM)
-
-lemma TWO: "2 = Suc 1"
-  by (import hollight TWO)
-
-lemma MULT_CLAUSES: "(ALL x::nat. (0::nat) * x = (0::nat)) &
-(ALL x::nat. x * (0::nat) = (0::nat)) &
-(ALL x::nat. (1::nat) * x = x) &
-(ALL x::nat. x * (1::nat) = x) &
-(ALL (x::nat) xa::nat. Suc x * xa = x * xa + xa) &
-(ALL (x::nat) xa::nat. x * Suc xa = x + x * xa)"
-  by (import hollight MULT_CLAUSES)
-
-lemma MULT_AC: "(m::nat) * (n::nat) = n * m &
-m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)"
-  by (import hollight MULT_AC)
-
-lemma EXP_EQ_1: "((x::nat) ^ (n::nat) = (1::nat)) = (x = (1::nat) | n = (0::nat))"
-  by (import hollight EXP_EQ_1)
-
-lemma LT_ANTISYM: "~ ((m::nat) < (n::nat) & n < m)"
-  by (import hollight LT_ANTISYM)
-
-lemma LET_ANTISYM: "~ ((m::nat) <= (n::nat) & n < m)"
-  by (import hollight LET_ANTISYM)
-
-lemma LTE_ANTISYM: "~ ((x::nat) < (xa::nat) & xa <= x)"
-  by (import hollight LTE_ANTISYM)
-
-lemma LT_CASES: "(m::nat) < (n::nat) | n < m | m = n"
-  by (import hollight LT_CASES)
-
-lemma LTE_CASES: "(x::nat) < (xa::nat) | xa <= x"
-  by (import hollight LTE_CASES)
-
-lemma LE_1: "(ALL x::nat. x ~= (0::nat) --> (0::nat) < x) &
-(ALL x::nat. x ~= (0::nat) --> (1::nat) <= x) &
-(ALL x>0::nat. x ~= (0::nat)) &
-(ALL x>0::nat. (1::nat) <= x) &
-(ALL x>=1::nat. (0::nat) < x) & (ALL x>=1::nat. x ~= (0::nat))"
-  by (import hollight LE_1)
-
-lemma LT_EXISTS: "(m < n) = (EX d. n = m + Suc d)"
-  by (import hollight LT_EXISTS)
-
-lemma LT_ADD: "((m::nat) < m + (n::nat)) = ((0::nat) < n)"
-  by (import hollight LT_ADD)
-
-lemma LT_ADDR: "((xa::nat) < (x::nat) + xa) = ((0::nat) < x)"
-  by (import hollight LT_ADDR)
-
-lemma LT_LMULT: "(m::nat) ~= (0::nat) & (n::nat) < (p::nat) ==> m * n < m * p"
-  by (import hollight LT_LMULT)
-
-lemma LE_MULT_LCANCEL: "((m::nat) * (n::nat) <= m * (p::nat)) = (m = (0::nat) | n <= p)"
-  by (import hollight LE_MULT_LCANCEL)
-
-lemma LE_MULT_RCANCEL: "((x::nat) * (xb::nat) <= (xa::nat) * xb) = (x <= xa | xb = (0::nat))"
-  by (import hollight LE_MULT_RCANCEL)
-
-lemma LT_MULT_LCANCEL: "((m::nat) * (n::nat) < m * (p::nat)) = (m ~= (0::nat) & n < p)"
-  by (import hollight LT_MULT_LCANCEL)
-
-lemma LT_MULT_RCANCEL: "((x::nat) * (xb::nat) < (xa::nat) * xb) = (x < xa & xb ~= (0::nat))"
-  by (import hollight LT_MULT_RCANCEL)
-
-lemma LT_MULT2: "(m::nat) < (n::nat) & (p::nat) < (q::nat) ==> m * p < n * q"
-  by (import hollight LT_MULT2)
-
-lemma WLOG_LE: "(ALL (m::nat) n::nat. (P::nat => nat => bool) m n = P n m) &
-(ALL (m::nat) n::nat. m <= n --> P m n)
-==> P (m::nat) (x::nat)"
-  by (import hollight WLOG_LE)
-
-lemma WLOG_LT: "(ALL m::nat. (P::nat => nat => bool) m m) &
-(ALL (m::nat) n::nat. P m n = P n m) &
-(ALL (m::nat) n::nat. m < n --> P m n)
-==> P (m::nat) (x::nat)"
-  by (import hollight WLOG_LT)
-
-lemma num_WOP: "Ex (P::nat => bool) = (EX n::nat. P n & (ALL m<n. ~ P m))"
-  by (import hollight num_WOP)
-
-lemma num_MAX: "(Ex (P::nat => bool) & (EX M::nat. ALL x::nat. P x --> x <= M)) =
-(EX m::nat. P m & (ALL x::nat. P x --> x <= m))"
-  by (import hollight num_MAX)
-
-lemma NOT_EVEN: "odd (n::nat) = odd n"
-  by (import hollight NOT_EVEN)
-
-lemma NOT_ODD: "(~ odd (n::nat)) = even n"
-  by (import hollight NOT_ODD)
-
-lemma EVEN_OR_ODD: "even (n::nat) | odd n"
-  by (import hollight EVEN_OR_ODD)
-
-lemma EVEN_AND_ODD: "~ (even (x::nat) & odd x)"
-  by (import hollight EVEN_AND_ODD)
-
-lemma EVEN_EXP: "even ((m::nat) ^ (n::nat)) = (even m & n ~= (0::nat))"
-  by (import hollight EVEN_EXP)
-
-lemma ODD_MULT: "odd ((m::nat) * (n::nat)) = (odd m & odd n)"
-  by (import hollight ODD_MULT)
-
-lemma ODD_EXP: "odd ((m::nat) ^ (n::nat)) = (odd m | n = (0::nat))"
-  by (import hollight ODD_EXP)
-
-lemma EVEN_DOUBLE: "even ((2::nat) * (n::nat))"
-  by (import hollight EVEN_DOUBLE)
-
-lemma ODD_DOUBLE: "odd (Suc (2 * x))"
-  by (import hollight ODD_DOUBLE)
-
-lemma EVEN_EXISTS_LEMMA: "(even n --> (EX m. n = 2 * m)) & (odd n --> (EX m. n = Suc (2 * m)))"
-  by (import hollight EVEN_EXISTS_LEMMA)
-
-lemma EVEN_ODD_DECOMPOSITION: "(EX (k::nat) m::nat. odd m & (n::nat) = (2::nat) ^ k * m) = (n ~= (0::nat))"
-  by (import hollight EVEN_ODD_DECOMPOSITION)
-
-lemma SUB_0: "(0::nat) - (x::nat) = (0::nat) & x - (0::nat) = x"
-  by (import hollight SUB_0)
-
-lemma SUB_PRESUC: "Suc m - n - Suc 0 = m - n"
-  by (import hollight SUB_PRESUC)
-
-lemma ADD_SUBR: "(xa::nat) - ((x::nat) + xa) = (0::nat)"
-  by (import hollight ADD_SUBR)
-
-lemma EVEN_SUB: "even ((m::nat) - (n::nat)) = (m <= n | even m = even n)"
-  by (import hollight EVEN_SUB)
-
-lemma ODD_SUB: "odd ((x::nat) - (xa::nat)) = (xa < x & odd x ~= odd xa)"
-  by (import hollight ODD_SUB)
-
-lemma EXP_LT_0: "((0::nat) < (xa::nat) ^ (x::nat)) = (xa ~= (0::nat) | x = (0::nat))"
-  by (import hollight EXP_LT_0)
-
-lemma LT_EXP: "((x::nat) ^ (m::nat) < x ^ (n::nat)) =
-((2::nat) <= x & m < n | x = (0::nat) & m ~= (0::nat) & n = (0::nat))"
-  by (import hollight LT_EXP)
-
-lemma LE_EXP: "((x::nat) ^ (m::nat) <= x ^ (n::nat)) =
-(if x = (0::nat) then m = (0::nat) --> n = (0::nat)
- else x = (1::nat) | m <= n)"
-  by (import hollight LE_EXP)
-
-lemma EQ_EXP: "((x::nat) ^ (m::nat) = x ^ (n::nat)) =
-(if x = (0::nat) then (m = (0::nat)) = (n = (0::nat))
- else x = (1::nat) | m = n)"
-  by (import hollight EQ_EXP)
-
-lemma EXP_MONO_LE_IMP: "(x::nat) <= (xa::nat) ==> x ^ (xb::nat) <= xa ^ xb"
-  by (import hollight EXP_MONO_LE_IMP)
-
-lemma EXP_MONO_LT_IMP: "(x::nat) < (y::nat) & (n::nat) ~= (0::nat) ==> x ^ n < y ^ n"
-  by (import hollight EXP_MONO_LT_IMP)
-
-lemma EXP_MONO_LE: "((x::nat) ^ (n::nat) <= (y::nat) ^ n) = (x <= y | n = (0::nat))"
-  by (import hollight EXP_MONO_LE)
-
-lemma EXP_MONO_LT: "((x::nat) ^ (xb::nat) < (xa::nat) ^ xb) = (x < xa & xb ~= (0::nat))"
-  by (import hollight EXP_MONO_LT)
-
-lemma EXP_MONO_EQ: "((x::nat) ^ (xb::nat) = (xa::nat) ^ xb) = (x = xa | xb = (0::nat))"
-  by (import hollight EXP_MONO_EQ)
-
-lemma DIVMOD_EXIST: "(n::nat) ~= (0::nat) ==> EX (q::nat) r::nat. (m::nat) = q * n + r & r < n"
-  by (import hollight DIVMOD_EXIST)
-
-lemma DIVMOD_EXIST_0: "EX (x::nat) xa::nat.
-   if (n::nat) = (0::nat) then x = (0::nat) & xa = (m::nat)
-   else m = x * n + xa & xa < n"
-  by (import hollight DIVMOD_EXIST_0)
-
-lemma DIVISION: "(n::nat) ~= (0::nat) ==> (m::nat) = m div n * n + m mod n & m mod n < n"
-  by (import hollight DIVISION)
-
-lemma DIVMOD_UNIQ_LEMMA: "((m::nat) = (q1::nat) * (n::nat) + (r1::nat) & r1 < n) &
-m = (q2::nat) * n + (r2::nat) & r2 < n
-==> q1 = q2 & r1 = r2"
-  by (import hollight DIVMOD_UNIQ_LEMMA)
-
-lemma DIVMOD_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n
-==> m div n = q & m mod n = r"
-  by (import hollight DIVMOD_UNIQ)
-
-lemma MOD_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n ==> m mod n = r"
-  by (import hollight MOD_UNIQ)
-
-lemma DIV_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n ==> m div n = q"
-  by (import hollight DIV_UNIQ)
-
-lemma MOD_EQ: "(m::nat) = (n::nat) + (q::nat) * (p::nat) ==> m mod p = n mod p"
-  by (import hollight MOD_EQ)
-
-lemma DIV_LE: "(n::nat) ~= (0::nat) ==> (m::nat) div n <= m"
-  by (import hollight DIV_LE)
-
-lemma DIV_MUL_LE: "(n::nat) * ((m::nat) div n) <= m"
-  by (import hollight DIV_MUL_LE)
-
-lemma MOD_MOD: "(n::nat) * (p::nat) ~= (0::nat) ==> (m::nat) mod (n * p) mod n = m mod n"
-  by (import hollight MOD_MOD)
-
-lemma MOD_MOD_REFL: "(n::nat) ~= (0::nat) ==> (m::nat) mod n mod n = m mod n"
-  by (import hollight MOD_MOD_REFL)
-
-lemma DIV_MULT2: "(x::nat) * (xb::nat) ~= (0::nat) ==> x * (xa::nat) div (x * xb) = xa div xb"
-  by (import hollight DIV_MULT2)
-
-lemma MOD_MULT2: "(x::nat) * (xb::nat) ~= (0::nat)
-==> x * (xa::nat) mod (x * xb) = x * (xa mod xb)"
-  by (import hollight MOD_MULT2)
-
-lemma MOD_EXISTS: "(EX q::nat. (m::nat) = (n::nat) * q) =
-(if n = (0::nat) then m = (0::nat) else m mod n = (0::nat))"
-  by (import hollight MOD_EXISTS)
-
-lemma LE_RDIV_EQ: "(a::nat) ~= (0::nat) ==> ((n::nat) <= (b::nat) div a) = (a * n <= b)"
-  by (import hollight LE_RDIV_EQ)
-
-lemma LE_LDIV_EQ: "(a::nat) ~= (0::nat)
-==> ((b::nat) div a <= (n::nat)) = (b < a * (n + (1::nat)))"
-  by (import hollight LE_LDIV_EQ)
-
-lemma LE_LDIV: "(x::nat) ~= (0::nat) & (xa::nat) <= x * (xb::nat) ==> xa div x <= xb"
-  by (import hollight LE_LDIV)
-
-lemma DIV_MONO: "(p::nat) ~= (0::nat) & (m::nat) <= (n::nat) ==> m div p <= n div p"
-  by (import hollight DIV_MONO)
-
-lemma DIV_MONO_LT: "(p::nat) ~= (0::nat) & (m::nat) + p <= (n::nat) ==> m div p < n div p"
-  by (import hollight DIV_MONO_LT)
-
-lemma DIV_EQ_0: "(n::nat) ~= (0::nat) ==> ((m::nat) div n = (0::nat)) = (m < n)"
-  by (import hollight DIV_EQ_0)
-
-lemma MOD_EQ_0: "(n::nat) ~= (0::nat)
-==> ((m::nat) mod n = (0::nat)) = (EX q::nat. m = q * n)"
-  by (import hollight MOD_EQ_0)
-
-lemma EVEN_MOD: "even (n::nat) = (n mod (2::nat) = (0::nat))"
-  by (import hollight EVEN_MOD)
-
-lemma ODD_MOD: "odd (n::nat) = (n mod (2::nat) = (1::nat))"
-  by (import hollight ODD_MOD)
-
-lemma MOD_MULT_RMOD: "(n::nat) ~= (0::nat) ==> (m::nat) * ((p::nat) mod n) mod n = m * p mod n"
-  by (import hollight MOD_MULT_RMOD)
-
-lemma MOD_MULT_LMOD: "(xa::nat) ~= (0::nat) ==> (x::nat) mod xa * (xb::nat) mod xa = x * xb mod xa"
-  by (import hollight MOD_MULT_LMOD)
-
-lemma MOD_MULT_MOD2: "(xa::nat) ~= (0::nat)
-==> (x::nat) mod xa * ((xb::nat) mod xa) mod xa = x * xb mod xa"
-  by (import hollight MOD_MULT_MOD2)
-
-lemma MOD_EXP_MOD: "(n::nat) ~= (0::nat) ==> ((m::nat) mod n) ^ (p::nat) mod n = m ^ p mod n"
-  by (import hollight MOD_EXP_MOD)
-
-lemma MOD_ADD_MOD: "(n::nat) ~= (0::nat)
-==> ((a::nat) mod n + (b::nat) mod n) mod n = (a + b) mod n"
-  by (import hollight MOD_ADD_MOD)
-
-lemma DIV_ADD_MOD: "(n::nat) ~= (0::nat)
-==> (((a::nat) + (b::nat)) mod n = a mod n + b mod n) =
-    ((a + b) div n = a div n + b div n)"
-  by (import hollight DIV_ADD_MOD)
-
-lemma MOD_LE: "(n::nat) ~= (0::nat) ==> (m::nat) mod n <= m"
-  by (import hollight MOD_LE)
-
-lemma DIV_MONO2: "(p::nat) ~= (0::nat) & p <= (m::nat) ==> (n::nat) div m <= n div p"
-  by (import hollight DIV_MONO2)
-
-lemma DIV_LE_EXCLUSION: "(b::nat) ~= (0::nat) & b * (c::nat) < ((a::nat) + (1::nat)) * (d::nat)
-==> c div d <= a div b"
-  by (import hollight DIV_LE_EXCLUSION)
-
-lemma DIV_EQ_EXCLUSION: "(b::nat) * (c::nat) < ((a::nat) + (1::nat)) * (d::nat) &
-a * d < (c + (1::nat)) * b
-==> a div b = c div d"
-  by (import hollight DIV_EQ_EXCLUSION)
-
-lemma MULT_DIV_LE: "(p::nat) ~= (0::nat) ==> (m::nat) * ((n::nat) div p) <= m * n div p"
-  by (import hollight MULT_DIV_LE)
-
-lemma DIV_DIV: "(xa::nat) * (xb::nat) ~= (0::nat)
-==> (x::nat) div xa div xb = x div (xa * xb)"
-  by (import hollight DIV_DIV)
-
-lemma DIV_MOD: "(xa::nat) * (xb::nat) ~= (0::nat)
-==> (x::nat) div xa mod xb = x mod (xa * xb) div xa"
-  by (import hollight DIV_MOD)
-
-lemma PRE_ELIM_THM: "P (n - Suc 0) = (ALL m. n = Suc m | m = 0 & n = 0 --> P m)"
-  by (import hollight PRE_ELIM_THM)
-
-lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
-(ALL d::nat. a = b + d | a < b & d = (0::nat) --> P d)"
-  by (import hollight SUB_ELIM_THM)
-
-lemma DIVMOD_ELIM_THM: "(P::nat => nat => bool) ((m::nat) div (n::nat)) (m mod n) =
-(ALL (x::nat) xa::nat.
-    n = (0::nat) & x = (0::nat) & xa = m | m = x * n + xa & xa < n -->
-    P x xa)"
-  by (import hollight DIVMOD_ELIM_THM)
-
-definition
-  minimal :: "(nat => bool) => nat"  where
-  "minimal == %u. SOME n. u n & (ALL m<n. ~ u m)"
-
-lemma DEF_minimal: "minimal = (%u. SOME n. u n & (ALL m<n. ~ u m))"
-  by (import hollight DEF_minimal)
-
-lemma MINIMAL: "Ex P = (P (minimal P) & (ALL x<minimal P. ~ P x))"
-  by (import hollight MINIMAL)
-
-lemma TRANSITIVE_STEPWISE_LT_EQ: "(!!x y z. R x y & R y z ==> R x z)
-==> (ALL m n. m < n --> R m n) = (ALL n. R n (Suc n))"
-  by (import hollight TRANSITIVE_STEPWISE_LT_EQ)
-
-lemma TRANSITIVE_STEPWISE_LT: "[| (ALL x y z. R x y & R y z --> R x z) & (ALL n. R n (Suc n)); m < n |]
-==> R m n"
-  by (import hollight TRANSITIVE_STEPWISE_LT)
-
-lemma TRANSITIVE_STEPWISE_LE_EQ: "(ALL x. R x x) & (ALL x y z. R x y & R y z --> R x z)
-==> (ALL m n. m <= n --> R m n) = (ALL n. R n (Suc n))"
-  by (import hollight TRANSITIVE_STEPWISE_LE_EQ)
-
-lemma TRANSITIVE_STEPWISE_LE: "[| (ALL x. R x x) &
-   (ALL x y z. R x y & R y z --> R x z) & (ALL n. R n (Suc n));
-   m <= n |]
-==> R m n"
-  by (import hollight TRANSITIVE_STEPWISE_LE)
-
-lemma WF_EQ: "wfP (u_556::'A => 'A => bool) =
-(ALL P::'A => bool.
-    Ex P = (EX x::'A. P x & (ALL y::'A. u_556 y x --> ~ P y)))"
-  by (import hollight WF_EQ)
-
-lemma WF_IND: "wfP (u_556::'A => 'A => bool) =
-(ALL P::'A => bool.
-    (ALL x::'A. (ALL y::'A. u_556 y x --> P y) --> P x) --> All P)"
-  by (import hollight WF_IND)
-
-lemma WF_DCHAIN: "wfP (u_556::'A => 'A => bool) =
-(~ (EX s::nat => 'A. ALL n::nat. u_556 (s (Suc n)) (s n)))"
-  by (import hollight WF_DCHAIN)
-
-lemma WF_UREC: "[| wfP (u_556::'A => 'A => bool);
-   !!(f::'A => 'B) (g::'A => 'B) x::'A.
-      (!!z::'A. u_556 z x ==> f z = g z)
-      ==> (H::('A => 'B) => 'A => 'B) f x = H g x;
-   (ALL x::'A. (f::'A => 'B) x = H f x) &
-   (ALL x::'A. (g::'A => 'B) x = H g x) |]
-==> f = g"
-  by (import hollight WF_UREC)
-
-lemma WF_UREC_WF: "(!!(H::('A => bool) => 'A => bool) (f::'A => bool) g::'A => bool.
-    [| !!(f::'A => bool) (g::'A => bool) x::'A.
-          (!!z::'A. (u_556::'A => 'A => bool) z x ==> f z = g z)
-          ==> H f x = H g x;
-       (ALL x::'A. f x = H f x) & (ALL x::'A. g x = H g x) |]
-    ==> f = g)
-==> wfP u_556"
-  by (import hollight WF_UREC_WF)
-
-lemma WF_REC_INVARIANT: "[| wfP (u_556::'A => 'A => bool);
-   !!(f::'A => 'B) (g::'A => 'B) x::'A.
-      (!!z::'A. u_556 z x ==> f z = g z & (S::'A => 'B => bool) z (f z))
-      ==> (H::('A => 'B) => 'A => 'B) f x = H g x & S x (H f x) |]
-==> EX f::'A => 'B. ALL x::'A. f x = H f x"
-  by (import hollight WF_REC_INVARIANT)
-
-lemma WF_REC: "[| wfP (u_556::'A => 'A => bool);
-   !!(f::'A => 'B) (g::'A => 'B) x::'A.
-      (!!z::'A. u_556 z x ==> f z = g z)
-      ==> (H::('A => 'B) => 'A => 'B) f x = H g x |]
-==> EX f::'A => 'B. ALL x::'A. f x = H f x"
-  by (import hollight WF_REC)
-
-lemma WF_REC_WF: "(!!H::('A => nat) => 'A => nat.
-    (!!(f::'A => nat) (g::'A => nat) x::'A.
-        (!!z::'A. (u_556::'A => 'A => bool) z x ==> f z = g z)
-        ==> H f x = H g x)
-    ==> EX f::'A => nat. ALL x::'A. f x = H f x)
-==> wfP u_556"
-  by (import hollight WF_REC_WF)
-
-lemma WF_EREC: "[| wfP (u_556::'A => 'A => bool);
-   !!(f::'A => 'B) (g::'A => 'B) x::'A.
-      (!!z::'A. u_556 z x ==> f z = g z)
-      ==> (H::('A => 'B) => 'A => 'B) f x = H g x |]
-==> EX! f::'A => 'B. ALL x::'A. f x = H f x"
-  by (import hollight WF_EREC)
-
-lemma WF_SUBSET: "(ALL (x::'A) y::'A.
-    (u_556::'A => 'A => bool) x y --> (u_670::'A => 'A => bool) x y) &
-wfP u_670
-==> wfP u_556"
-  by (import hollight WF_SUBSET)
-
-lemma WF_MEASURE_GEN: "wfP (u_556::'B => 'B => bool)
-==> wfP (%(x::'A) x'::'A. u_556 ((m::'A => 'B) x) (m x'))"
-  by (import hollight WF_MEASURE_GEN)
-
-lemma WF_LEX_DEPENDENT: "wfP (R::'A => 'A => bool) & (ALL x::'A. wfP ((S::'A => 'B => 'B => bool) x))
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
-            ALL (r1::'A) s1::'B.
-               f (r1, s1) =
-               (SOME f::'A * 'B => bool.
-                   ALL (r2::'A) s2::'B.
-                      f (r2, s2) = (R r1 r2 | r1 = r2 & S r1 s1 s2)))"
-  by (import hollight WF_LEX_DEPENDENT)
-
-lemma WF_LEX: "wfP (x::'A => 'A => bool) & wfP (xa::'B => 'B => bool)
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
-            ALL (r1::'A) s1::'B.
-               f (r1, s1) =
-               (SOME f::'A * 'B => bool.
-                   ALL (r2::'A) s2::'B.
-                      f (r2, s2) = (x r1 r2 | r1 = r2 & xa s1 s2)))"
-  by (import hollight WF_LEX)
-
-lemma WF_POINTWISE: "wfP (u_556::'A => 'A => bool) & wfP (u_670::'B => 'B => bool)
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
-            ALL (x1::'A) y1::'B.
-               f (x1, y1) =
-               (SOME f::'A * 'B => bool.
-                   ALL (x2::'A) y2::'B.
-                      f (x2, y2) = (u_556 x1 x2 & u_670 y1 y2)))"
-  by (import hollight WF_POINTWISE)
-
-lemma WF_num: "(wfP::(nat => nat => bool) => bool) (op <::nat => nat => bool)"
-  by (import hollight WF_num)
-
-lemma WF_REC_num: "(!!(f::nat => 'A) (g::nat => 'A) x::nat.
-    (!!z::nat. z < x ==> f z = g z)
-    ==> (H::(nat => 'A) => nat => 'A) f x = H g x)
-==> EX f::nat => 'A. ALL x::nat. f x = H f x"
-  by (import hollight WF_REC_num)
-
-lemma WF_MEASURE: "wfP (%(a::'A) b::'A. (a, b) : measure (m::'A => nat))"
-  by (import hollight WF_MEASURE)
-
-lemma MEASURE_LE: "(ALL x::'q_12099.
-    (x, a::'q_12099) : measure (m::'q_12099 => nat) -->
-    (x, b::'q_12099) : measure m) =
-(m a <= m b)"
-  by (import hollight MEASURE_LE)
-
-lemma WF_REFL: "wfP (u_556::'A => 'A => bool) ==> ~ u_556 (x::'A) x"
-  by (import hollight WF_REFL)
-
-lemma WF_REC_TAIL: "EX f::'A => 'B.
-   ALL x::'A.
-      f x =
-      (if (P::'A => bool) x then f ((g::'A => 'A) x) else (h::'A => 'B) x)"
-  by (import hollight WF_REC_TAIL)
-
-lemma WF_REC_TAIL_GENERAL: "wfP (u_556::'A => 'A => bool) &
-(ALL (f::'A => 'B) (g::'A => 'B) x::'A.
-    (ALL z::'A. u_556 z x --> f z = g z) -->
-    (P::('A => 'B) => 'A => bool) f x = P g x &
-    (G::('A => 'B) => 'A => 'A) f x = G g x &
-    (H::('A => 'B) => 'A => 'B) f x = H g x) &
-(ALL (f::'A => 'B) (g::'A => 'B) x::'A.
-    (ALL z::'A. u_556 z x --> f z = g z) --> H f x = H g x) &
-(ALL (f::'A => 'B) (x::'A) y::'A. P f x & u_556 y (G f x) --> u_556 y x)
-==> EX f::'A => 'B. ALL x::'A. f x = (if P f x then f (G f x) else H f x)"
-  by (import hollight WF_REC_TAIL_GENERAL)
-
-lemma ARITH_ZERO: "(0::nat) = (0::nat) & (0::nat) = (0::nat)"
-  by (import hollight ARITH_ZERO)
-
-lemma ARITH_SUC: "(ALL x. Suc x = Suc x) &
-Suc 0 = 1 &
-(ALL x. Suc (2 * x) = 2 * x + 1) & (ALL x. Suc (2 * x + 1) = 2 * Suc x)"
-  by (import hollight ARITH_SUC)
-
-lemma ARITH_PRE: "(ALL x. x - Suc 0 = x - Suc 0) &
-0 - Suc 0 = 0 &
-(ALL x. 2 * x - Suc 0 = (if x = 0 then 0 else 2 * (x - Suc 0) + 1)) &
-(ALL x. 2 * x + 1 - Suc 0 = 2 * x)"
-  by (import hollight ARITH_PRE)
-
-lemma ARITH_ADD: "(ALL (x::nat) xa::nat. x + xa = x + xa) &
-(0::nat) + (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) + (2::nat) * x = (2::nat) * x) &
-(ALL x::nat.
-    (0::nat) + ((2::nat) * x + (1::nat)) = (2::nat) * x + (1::nat)) &
-(ALL x::nat. (2::nat) * x + (0::nat) = (2::nat) * x) &
-(ALL x::nat. (2::nat) * x + (1::nat) + (0::nat) = (2::nat) * x + (1::nat)) &
-(ALL (x::nat) xa::nat. (2::nat) * x + (2::nat) * xa = (2::nat) * (x + xa)) &
-(ALL (x::nat) xa::nat.
-    (2::nat) * x + ((2::nat) * xa + (1::nat)) =
-    (2::nat) * (x + xa) + (1::nat)) &
-(ALL (x::nat) xa::nat.
-    (2::nat) * x + (1::nat) + (2::nat) * xa =
-    (2::nat) * (x + xa) + (1::nat)) &
-(ALL (x::nat) xa::nat.
-    (2::nat) * x + (1::nat) + ((2::nat) * xa + (1::nat)) =
-    (2::nat) * Suc (x + xa))"
-  by (import hollight ARITH_ADD)
-
-lemma ARITH_MULT: "(ALL (x::nat) xa::nat. x * xa = x * xa) &
-(0::nat) * (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) * ((2::nat) * x) = (0::nat)) &
-(ALL x::nat. (0::nat) * ((2::nat) * x + (1::nat)) = (0::nat)) &
-(ALL x::nat. (2::nat) * x * (0::nat) = (0::nat)) &
-(ALL x::nat. ((2::nat) * x + (1::nat)) * (0::nat) = (0::nat)) &
-(ALL (x::nat) xa::nat.
-    (2::nat) * x * ((2::nat) * xa) = (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
-    (2::nat) * x * ((2::nat) * xa + (1::nat)) =
-    (2::nat) * x + (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat)) * ((2::nat) * xa) =
-    (2::nat) * xa + (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat)) * ((2::nat) * xa + (1::nat)) =
-    (2::nat) * x + (1::nat) +
-    ((2::nat) * xa + (2::nat) * ((2::nat) * (x * xa))))"
-  by (import hollight ARITH_MULT)
-
-lemma ARITH_EXP: "(ALL (x::nat) xa::nat. x ^ xa = x ^ xa) &
-(0::nat) ^ (0::nat) = (1::nat) &
-(ALL m::nat. ((2::nat) * m) ^ (0::nat) = (1::nat)) &
-(ALL m::nat. ((2::nat) * m + (1::nat)) ^ (0::nat) = (1::nat)) &
-(ALL n::nat. (0::nat) ^ ((2::nat) * n) = (0::nat) ^ n * (0::nat) ^ n) &
-(ALL (m::nat) n::nat.
-    ((2::nat) * m) ^ ((2::nat) * n) =
-    ((2::nat) * m) ^ n * ((2::nat) * m) ^ n) &
-(ALL (m::nat) n::nat.
-    ((2::nat) * m + (1::nat)) ^ ((2::nat) * n) =
-    ((2::nat) * m + (1::nat)) ^ n * ((2::nat) * m + (1::nat)) ^ n) &
-(ALL n::nat. (0::nat) ^ ((2::nat) * n + (1::nat)) = (0::nat)) &
-(ALL (m::nat) n::nat.
-    ((2::nat) * m) ^ ((2::nat) * n + (1::nat)) =
-    (2::nat) * m * (((2::nat) * m) ^ n * ((2::nat) * m) ^ n)) &
-(ALL (m::nat) n::nat.
-    ((2::nat) * m + (1::nat)) ^ ((2::nat) * n + (1::nat)) =
-    ((2::nat) * m + (1::nat)) *
-    (((2::nat) * m + (1::nat)) ^ n * ((2::nat) * m + (1::nat)) ^ n))"
-  by (import hollight ARITH_EXP)
-
-lemma ARITH_EVEN: "(ALL x::nat. even x = even x) &
-even (0::nat) = True &
-(ALL x::nat. even ((2::nat) * x) = True) &
-(ALL x::nat. even ((2::nat) * x + (1::nat)) = False)"
-  by (import hollight ARITH_EVEN)
-
-lemma ARITH_ODD: "(ALL x::nat. odd x = odd x) &
-odd (0::nat) = False &
-(ALL x::nat. odd ((2::nat) * x) = False) &
-(ALL x::nat. odd ((2::nat) * x + (1::nat)) = True)"
-  by (import hollight ARITH_ODD)
-
-lemma ARITH_LE: "(ALL (x::nat) xa::nat. (x <= xa) = (x <= xa)) &
-((0::nat) <= (0::nat)) = True &
-(ALL x::nat. ((2::nat) * x <= (0::nat)) = (x <= (0::nat))) &
-(ALL x::nat. ((2::nat) * x + (1::nat) <= (0::nat)) = False) &
-(ALL x::nat. ((0::nat) <= (2::nat) * x) = True) &
-(ALL x::nat. ((0::nat) <= (2::nat) * x + (1::nat)) = True) &
-(ALL (x::nat) xa::nat. ((2::nat) * x <= (2::nat) * xa) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x <= (2::nat) * xa + (1::nat)) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat) <= (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat) <= (2::nat) * xa + (1::nat)) = (x <= xa))"
-  by (import hollight ARITH_LE)
-
-lemma ARITH_LT: "(ALL (x::nat) xa::nat. (x < xa) = (x < xa)) &
-((0::nat) < (0::nat)) = False &
-(ALL x::nat. ((2::nat) * x < (0::nat)) = False) &
-(ALL x::nat. ((2::nat) * x + (1::nat) < (0::nat)) = False) &
-(ALL x::nat. ((0::nat) < (2::nat) * x) = ((0::nat) < x)) &
-(ALL x::nat. ((0::nat) < (2::nat) * x + (1::nat)) = True) &
-(ALL (x::nat) xa::nat. ((2::nat) * x < (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x < (2::nat) * xa + (1::nat)) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat) < (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat) < (2::nat) * xa + (1::nat)) = (x < xa))"
-  by (import hollight ARITH_LT)
-
-lemma ARITH_EQ: "(ALL (x::nat) xa::nat. (x = xa) = (x = xa)) &
-((0::nat) = (0::nat)) = True &
-(ALL x::nat. ((2::nat) * x = (0::nat)) = (x = (0::nat))) &
-(ALL x::nat. ((2::nat) * x + (1::nat) = (0::nat)) = False) &
-(ALL x::nat. ((0::nat) = (2::nat) * x) = ((0::nat) = x)) &
-(ALL x::nat. ((0::nat) = (2::nat) * x + (1::nat)) = False) &
-(ALL (x::nat) xa::nat. ((2::nat) * x = (2::nat) * xa) = (x = xa)) &
-(ALL (x::nat) xa::nat. ((2::nat) * x = (2::nat) * xa + (1::nat)) = False) &
-(ALL (x::nat) xa::nat. ((2::nat) * x + (1::nat) = (2::nat) * xa) = False) &
-(ALL (x::nat) xa::nat.
-    ((2::nat) * x + (1::nat) = (2::nat) * xa + (1::nat)) = (x = xa))"
-  by (import hollight ARITH_EQ)
-
-lemma ARITH_SUB: "(ALL (x::nat) xa::nat. x - xa = x - xa) &
-(0::nat) - (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) - (2::nat) * x = (0::nat)) &
-(ALL x::nat. (0::nat) - ((2::nat) * x + (1::nat)) = (0::nat)) &
-(ALL x::nat. (2::nat) * x - (0::nat) = (2::nat) * x) &
-(ALL x::nat. (2::nat) * x + (1::nat) - (0::nat) = (2::nat) * x + (1::nat)) &
-(ALL (m::nat) n::nat. (2::nat) * m - (2::nat) * n = (2::nat) * (m - n)) &
-(ALL (m::nat) n::nat.
-    (2::nat) * m - ((2::nat) * n + (1::nat)) =
-    (2::nat) * (m - n) - Suc (0::nat)) &
-(ALL (m::nat) n::nat.
-    (2::nat) * m + (1::nat) - (2::nat) * n =
-    (if n <= m then (2::nat) * (m - n) + (1::nat) else (0::nat))) &
-(ALL (m::nat) n::nat.
-    (2::nat) * m + (1::nat) - ((2::nat) * n + (1::nat)) =
-    (2::nat) * (m - n))"
-  by (import hollight ARITH_SUB)
-
-lemma right_th: "(s::nat) * ((2::nat) * (x::nat) + (1::nat)) = s + (2::nat) * (s * x)"
-  by (import hollight right_th)
-
-lemma SEMIRING_PTHS: "(ALL (x::'A) (y::'A) z::'A.
-    (add::'A => 'A => 'A) x (add y z) = add (add x y) z) &
-(ALL (x::'A) y::'A. add x y = add y x) &
-(ALL x::'A. add (r0::'A) x = x) &
-(ALL (x::'A) (y::'A) z::'A.
-    (mul::'A => 'A => 'A) x (mul y z) = mul (mul x y) z) &
-(ALL (x::'A) y::'A. mul x y = mul y x) &
-(ALL x::'A. mul (r1::'A) x = x) &
-(ALL x::'A. mul r0 x = r0) &
-(ALL (x::'A) (y::'A) z::'A. mul x (add y z) = add (mul x y) (mul x z)) &
-(ALL x::'A. (pwr::'A => nat => 'A) x (0::nat) = r1) &
-(ALL (x::'A) n::nat. pwr x (Suc n) = mul x (pwr x n))
-==> mul r1 (x::'A) = x &
-    add (mul (a::'A) (m::'A)) (mul (b::'A) m) = mul (add a b) m &
-    add (mul a m) m = mul (add a r1) m &
-    add m (mul a m) = mul (add a r1) m &
-    add m m = mul (add r1 r1) m &
-    mul r0 m = r0 &
-    add r0 a = a &
-    add a r0 = a &
-    mul a b = mul b a &
-    mul (add a b) (c::'A) = add (mul a c) (mul b c) &
-    mul r0 a = r0 &
-    mul a r0 = r0 &
-    mul r1 a = a &
-    mul a r1 = a &
-    mul (mul (lx::'A) (ly::'A)) (mul (rx::'A) (ry::'A)) =
-    mul (mul lx rx) (mul ly ry) &
-    mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry)) &
-    mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry) &
-    mul (mul lx ly) rx = mul (mul lx rx) ly &
-    mul (mul lx ly) rx = mul lx (mul ly rx) &
-    mul lx rx = mul rx lx &
-    mul lx (mul rx ry) = mul (mul lx rx) ry &
-    mul lx (mul rx ry) = mul rx (mul lx ry) &
-    add (add a b) (add c (d::'A)) = add (add a c) (add b d) &
-    add (add a b) c = add a (add b c) &
-    add a (add c d) = add c (add a d) &
-    add (add a b) c = add (add a c) b &
-    add a c = add c a &
-    add a (add c d) = add (add a c) d &
-    mul (pwr x (p::nat)) (pwr x (q::nat)) = pwr x (p + q) &
-    mul x (pwr x q) = pwr x (Suc q) &
-    mul (pwr x q) x = pwr x (Suc q) &
-    mul x x = pwr x (2::nat) &
-    pwr (mul x (y::'A)) q = mul (pwr x q) (pwr y q) &
-    pwr (pwr x p) q = pwr x (p * q) &
-    pwr x (0::nat) = r1 &
-    pwr x (1::nat) = x &
-    mul x (add y (z::'A)) = add (mul x y) (mul x z) &
-    pwr x (Suc q) = mul x (pwr x q)"
-  by (import hollight SEMIRING_PTHS)
-
-lemma NUM_INTEGRAL_LEMMA: "(w::nat) = (x::nat) + (d::nat) & (y::nat) = (z::nat) + (e::nat)
-==> (w * y + x * z = w * z + x * y) = (w = x | y = z)"
-  by (import hollight NUM_INTEGRAL_LEMMA)
-
-lemma NUM_INTEGRAL: "(ALL x::nat. (0::nat) * x = (0::nat)) &
-(ALL (x::nat) (xa::nat) xb::nat. (x + xa = x + xb) = (xa = xb)) &
-(ALL (w::nat) (x::nat) (y::nat) z::nat.
-    (w * y + x * z = w * z + x * y) = (w = x | y = z))"
-  by (import hollight NUM_INTEGRAL)
-
-lemma INJ_INVERSE2: "(!!(x1::'A) (y1::'B) (x2::'A) y2::'B.
-    ((P::'A => 'B => 'C) x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2))
-==> EX (x::'C => 'A) Y::'C => 'B.
-       ALL (xa::'A) y::'B. x (P xa y) = xa & Y (P xa y) = y"
-  by (import hollight INJ_INVERSE2)
-
-definition
-  NUMPAIR :: "nat => nat => nat"  where
-  "NUMPAIR == %u ua. 2 ^ u * (2 * ua + 1)"
-
-lemma DEF_NUMPAIR: "NUMPAIR = (%u ua. 2 ^ u * (2 * ua + 1))"
-  by (import hollight DEF_NUMPAIR)
-
-lemma NUMPAIR_INJ_LEMMA: "NUMPAIR x xa = NUMPAIR xb xc ==> x = xb"
-  by (import hollight NUMPAIR_INJ_LEMMA)
-
-lemma NUMPAIR_INJ: "(NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
-  by (import hollight NUMPAIR_INJ)
-
-definition
-  NUMFST :: "nat => nat"  where
-  "NUMFST == SOME X. EX Y. ALL x y. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
-
-lemma DEF_NUMFST: "NUMFST = (SOME X. EX Y. ALL x y. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
-  by (import hollight DEF_NUMFST)
-
-definition
-  NUMSND :: "nat => nat"  where
-  "NUMSND == SOME Y. ALL x y. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
-
-lemma DEF_NUMSND: "NUMSND = (SOME Y. ALL x y. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
-  by (import hollight DEF_NUMSND)
-
-definition
-  NUMSUM :: "bool => nat => nat"  where
-  "NUMSUM == %u ua. if u then Suc (2 * ua) else 2 * ua"
-
-lemma DEF_NUMSUM: "NUMSUM = (%u ua. if u then Suc (2 * ua) else 2 * ua)"
-  by (import hollight DEF_NUMSUM)
-
-lemma NUMSUM_INJ: "(NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
-  by (import hollight NUMSUM_INJ)
-
-definition
-  NUMLEFT :: "nat => bool"  where
-  "NUMLEFT == SOME X. EX Y. ALL x y. X (NUMSUM x y) = x & Y (NUMSUM x y) = y"
-
-lemma DEF_NUMLEFT: "NUMLEFT = (SOME X. EX Y. ALL x y. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
-  by (import hollight DEF_NUMLEFT)
-
-definition
-  NUMRIGHT :: "nat => nat"  where
-  "NUMRIGHT == SOME Y. ALL x y. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y"
-
-lemma DEF_NUMRIGHT: "NUMRIGHT = (SOME Y. ALL x y. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
-  by (import hollight DEF_NUMRIGHT)
-
-definition
-  INJN :: "nat => nat => 'A => bool"  where
-  "INJN == %(u::nat) (n::nat) a::'A. n = u"
-
-lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A. n = u)"
-  by (import hollight DEF_INJN)
-
-lemma INJN_INJ: "(op =::bool => bool => bool)
- ((op =::(nat => 'A::type => bool) => (nat => 'A::type => bool) => bool)
-   ((INJN::nat => nat => 'A::type => bool) (n1::nat))
-   ((INJN::nat => nat => 'A::type => bool) (n2::nat)))
- ((op =::nat => nat => bool) n1 n2)"
-  by (import hollight INJN_INJ)
-
-definition
-  INJA :: "'A => nat => 'A => bool"  where
-  "INJA == %(u::'A) (n::nat) b::'A. b = u"
-
-lemma DEF_INJA: "INJA = (%(u::'A) (n::nat) b::'A. b = u)"
-  by (import hollight DEF_INJA)
-
-lemma INJA_INJ: "(INJA (a1::'A) = INJA (a2::'A)) = (a1 = a2)"
-  by (import hollight INJA_INJ)
-
-definition
-  INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool"  where
-  "INJF == %(u::nat => nat => 'A => bool) n::nat. u (NUMFST n) (NUMSND n)"
-
-lemma DEF_INJF: "INJF = (%(u::nat => nat => 'A => bool) n::nat. u (NUMFST n) (NUMSND n))"
-  by (import hollight DEF_INJF)
-
-lemma INJF_INJ: "(INJF (f1::nat => nat => 'A => bool) =
- INJF (f2::nat => nat => 'A => bool)) =
-(f1 = f2)"
-  by (import hollight INJF_INJ)
-
-definition
-  INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool"  where
-  "INJP ==
-%(u::nat => 'A => bool) (ua::nat => 'A => bool) (n::nat) a::'A.
-   if NUMLEFT n then u (NUMRIGHT n) a else ua (NUMRIGHT n) a"
-
-lemma DEF_INJP: "INJP =
-(%(u::nat => 'A => bool) (ua::nat => 'A => bool) (n::nat) a::'A.
-    if NUMLEFT n then u (NUMRIGHT n) a else ua (NUMRIGHT n) a)"
-  by (import hollight DEF_INJP)
-
-lemma INJP_INJ: "(INJP (f1::nat => 'A => bool) (f2::nat => 'A => bool) =
- INJP (f1'::nat => 'A => bool) (f2'::nat => 'A => bool)) =
-(f1 = f1' & f2 = f2')"
-  by (import hollight INJP_INJ)
-
-definition
-  ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool"  where
-  "ZCONSTR ==
-%(u::nat) (ua::'A) ub::nat => nat => 'A => bool.
-   INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))"
-
-lemma DEF_ZCONSTR: "ZCONSTR =
-(%(u::nat) (ua::'A) ub::nat => nat => 'A => bool.
-    INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))"
-  by (import hollight DEF_ZCONSTR)
-
-definition
-  ZBOT :: "nat => 'A => bool"  where
-  "ZBOT == INJP (INJN (0::nat)) (SOME z::nat => 'A => bool. True)"
-
-lemma DEF_ZBOT: "ZBOT = INJP (INJN (0::nat)) (SOME z::nat => 'A => bool. True)"
-  by (import hollight DEF_ZBOT)
-
-lemma ZCONSTR_ZBOT: "ZCONSTR (x::nat) (xa::'A) (xb::nat => nat => 'A => bool) ~= ZBOT"
-  by (import hollight ZCONSTR_ZBOT)
-
-definition
-  ZRECSPACE :: "(nat => 'A => bool) => bool"  where
-  "ZRECSPACE ==
-%a::nat => 'A => bool.
-   ALL ZRECSPACE'::(nat => 'A => bool) => bool.
-      (ALL a::nat => 'A => bool.
-          a = ZBOT |
-          (EX (c::nat) (i::'A) r::nat => nat => 'A => bool.
-              a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
-          ZRECSPACE' a) -->
-      ZRECSPACE' a"
-
-lemma DEF_ZRECSPACE: "ZRECSPACE =
-(%a::nat => 'A => bool.
-    ALL ZRECSPACE'::(nat => 'A => bool) => bool.
-       (ALL a::nat => 'A => bool.
-           a = ZBOT |
-           (EX (c::nat) (i::'A) r::nat => nat => 'A => bool.
-               a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
-           ZRECSPACE' a) -->
-       ZRECSPACE' a)"
-  by (import hollight DEF_ZRECSPACE)
-
-typedef (open) 'a recspace = "Collect ZRECSPACE :: (nat \<Rightarrow> 'a \<Rightarrow> bool) set" 
-  morphisms "_dest_rec" "_mk_rec"
-  apply (rule light_ex_imp_nonempty [where t="ZBOT"])
-  by (import hollight TYDEF_recspace)
-
-syntax
-  "_dest_rec" :: _ ("'_dest'_rec")
-
-syntax
-  "_mk_rec" :: _ ("'_mk'_rec")
-
-lemmas "TYDEF_recspace_@intern" = typedef_hol2hollight 
-  [where a="a :: 'A recspace" and r=r ,
-   OF type_definition_recspace]
-
-definition
-  BOTTOM :: "'A recspace"  where
-  "BOTTOM == _mk_rec ZBOT"
-
-lemma DEF_BOTTOM: "BOTTOM = _mk_rec ZBOT"
-  by (import hollight DEF_BOTTOM)
-
-definition
-  CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace"  where
-  "CONSTR ==
-%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
-   _mk_rec (ZCONSTR u ua (%n::nat. _dest_rec (ub n)))"
-
-lemma DEF_CONSTR: "CONSTR =
-(%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
-    _mk_rec (ZCONSTR u ua (%n::nat. _dest_rec (ub n))))"
-  by (import hollight DEF_CONSTR)
-
-lemma MK_REC_INJ: "[| _mk_rec (x::nat => 'A::type => bool) =
-   _mk_rec (y::nat => 'A::type => bool);
-   ZRECSPACE x & ZRECSPACE y |]
-==> x = y"
-  by (import hollight MK_REC_INJ)
-
-lemma CONSTR_BOT: "CONSTR (c::nat) (i::'A) (r::nat => 'A recspace) ~= BOTTOM"
-  by (import hollight CONSTR_BOT)
-
-lemma CONSTR_INJ: "(CONSTR (c1::nat) (i1::'A) (r1::nat => 'A recspace) =
- CONSTR (c2::nat) (i2::'A) (r2::nat => 'A recspace)) =
-(c1 = c2 & i1 = i2 & r1 = r2)"
-  by (import hollight CONSTR_INJ)
-
-lemma CONSTR_IND: "(P::'A recspace => bool) BOTTOM &
-(ALL (c::nat) (i::'A) r::nat => 'A recspace.
-    (ALL n::nat. P (r n)) --> P (CONSTR c i r))
-==> P (x::'A recspace)"
-  by (import hollight CONSTR_IND)
-
-lemma CONSTR_REC: "EX f::'A recspace => 'B.
-   ALL (c::nat) (i::'A) r::nat => 'A recspace.
-      f (CONSTR c i r) =
-      (Fn::nat => 'A => (nat => 'A recspace) => (nat => 'B) => 'B) c i r
-       (%n::nat. f (r n))"
-  by (import hollight CONSTR_REC)
-
-definition
-  FCONS :: "'A => (nat => 'A) => nat => 'A"  where
-  "FCONS ==
-SOME FCONS::'A => (nat => 'A) => nat => 'A.
-   (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)"
-
-lemma DEF_FCONS: "FCONS =
-(SOME FCONS::'A => (nat => 'A) => nat => 'A.
-    (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 hollight DEF_FCONS)
-
-lemma FCONS_UNDO: "(f::nat => 'A) = FCONS (f (0::nat)) (f o Suc)"
-  by (import hollight FCONS_UNDO)
-
-definition
-  FNIL :: "nat => 'A"  where
-  "FNIL == %u::nat. SOME x::'A. True"
-
-lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A. True)"
-  by (import hollight DEF_FNIL)
-
-definition
-  ISO :: "('A => 'B) => ('B => 'A) => bool"  where
-  "ISO ==
-%(u::'A => 'B) ua::'B => 'A.
-   (ALL x::'B. u (ua x) = x) & (ALL y::'A. ua (u y) = y)"
-
-lemma DEF_ISO: "ISO =
-(%(u::'A => 'B) ua::'B => 'A.
-    (ALL x::'B. u (ua x) = x) & (ALL y::'A. ua (u y) = y))"
-  by (import hollight DEF_ISO)
-
-lemma ISO_REFL: "ISO (%x::'A. x) (%x::'A. x)"
-  by (import hollight ISO_REFL)
-
-lemma ISO_FUN: "ISO (f::'A => 'A') (f'::'A' => 'A) & ISO (g::'B => 'B') (g'::'B' => 'B)
-==> ISO (%(h::'A => 'B) a'::'A'. g (h (f' a')))
-     (%(h::'A' => 'B') a::'A. g' (h (f a)))"
-  by (import hollight ISO_FUN)
-
-lemma ISO_USAGE: "ISO (f::'q_17485 => 'q_17482) (g::'q_17482 => 'q_17485)
-==> (ALL P::'q_17485 => bool. All P = (ALL x::'q_17482. P (g x))) &
-    (ALL P::'q_17485 => bool. Ex P = (EX x::'q_17482. P (g x))) &
-    (ALL (a::'q_17485) b::'q_17482. (a = g b) = (f a = b))"
-  by (import hollight ISO_USAGE)
-
-typedef (open) char = "{a. ALL char'.
-       (ALL a.
-           (EX a0 a1 a2 a3 a4 a5 a6 a7.
-               a =
-               CONSTR (NUMERAL 0) (a0 :: bool, a1 :: bool, a2 :: bool, a3 :: bool, a4 :: bool, a5 :: bool, a6 :: bool, a7:: bool)
-                (%n. BOTTOM)) -->
-           char' a) -->
-       char' a}"
-  morphisms "_dest_char" "_mk_char"
-  apply (rule light_ex_imp_nonempty [where t="CONSTR (NUMERAL 0) (a0, a1, a2, a3, a4, a5, a6, a7) (%n. BOTTOM)"])
-  by (import hollight TYDEF_char)
-
-syntax
-  "_dest_char" :: _ ("'_dest'_char")
-
-syntax
-  "_mk_char" :: _ ("'_mk'_char")
-
-lemmas "TYDEF_char_@intern" = typedef_hol2hollight 
-  [where a="a :: hollight.char" and r=r ,
-   OF type_definition_char]
-
-consts
-  "_11937" :: "bool
-=> bool => bool => bool => bool => bool => bool => bool => hollight.char" ("'_11937")
-
-defs
-  "_11937_def": "_11937 ==
-%(a0::bool) (a1::bool) (a2::bool) (a3::bool) (a4::bool) (a5::bool)
-   (a6::bool) a7::bool.
-   _mk_char
-    (CONSTR (0::nat) (a0, a1, a2, a3, a4, a5, a6, a7) (%n::nat. BOTTOM))"
-
-lemma DEF__11937: "_11937 =
-(%(a0::bool) (a1::bool) (a2::bool) (a3::bool) (a4::bool) (a5::bool)
-    (a6::bool) a7::bool.
-    _mk_char
-     (CONSTR (0::nat) (a0, a1, a2, a3, a4, a5, a6, a7) (%n::nat. BOTTOM)))"
-  by (import hollight DEF__11937)
-
-definition
-  ASCII :: "bool
-=> bool => bool => bool => bool => bool => bool => bool => hollight.char"  where
-  "ASCII == _11937"
-
-lemma DEF_ASCII: "ASCII = _11937"
-  by (import hollight DEF_ASCII)
-
-consts
-  dist :: "nat * nat => nat" 
-
-defs
-  dist_def: "hollight.dist == %u. fst u - snd u + (snd u - fst u)"
-
-lemma DEF_dist: "hollight.dist = (%u. fst u - snd u + (snd u - fst u))"
-  by (import hollight DEF_dist)
-
-lemma DIST_REFL: "hollight.dist (x, x) = 0"
-  by (import hollight DIST_REFL)
-
-lemma DIST_LZERO: "hollight.dist (0, x) = x"
-  by (import hollight DIST_LZERO)
-
-lemma DIST_RZERO: "hollight.dist (x, 0) = x"
-  by (import hollight DIST_RZERO)
-
-lemma DIST_SYM: "hollight.dist (x, xa) = hollight.dist (xa, x)"
-  by (import hollight DIST_SYM)
-
-lemma DIST_LADD: "hollight.dist (x + xb, x + xa) = hollight.dist (xb, xa)"
-  by (import hollight DIST_LADD)
-
-lemma DIST_RADD: "hollight.dist (x + xa, xb + xa) = hollight.dist (x, xb)"
-  by (import hollight DIST_RADD)
-
-lemma DIST_LADD_0: "hollight.dist (x + xa, x) = xa"
-  by (import hollight DIST_LADD_0)
-
-lemma DIST_RADD_0: "hollight.dist (x, x + xa) = xa"
-  by (import hollight DIST_RADD_0)
-
-lemma DIST_LMUL: "x * hollight.dist (xa, xb) = hollight.dist (x * xa, x * xb)"
-  by (import hollight DIST_LMUL)
-
-lemma DIST_RMUL: "hollight.dist (x, xa) * xb = hollight.dist (x * xb, xa * xb)"
-  by (import hollight DIST_RMUL)
-
-lemma DIST_EQ_0: "(hollight.dist (x, xa) = 0) = (x = xa)"
-  by (import hollight DIST_EQ_0)
-
-lemma DIST_ELIM_THM: "P (hollight.dist (x, y)) =
-(ALL d. (x = y + d --> P d) & (y = x + d --> P d))"
-  by (import hollight DIST_ELIM_THM)
-
-lemma DIST_LE_CASES: "(hollight.dist (m, n) <= p) = (m <= n + p & n <= m + p)"
-  by (import hollight DIST_LE_CASES)
-
-lemma DIST_TRIANGLE_LE: "hollight.dist (m, n) + hollight.dist (n, p) <= q
-==> hollight.dist (m, p) <= q"
-  by (import hollight DIST_TRIANGLE_LE)
-
-lemma DIST_TRIANGLES_LE: "hollight.dist (m, n) <= r & hollight.dist (p, q) <= s
-==> hollight.dist (m, p) <= hollight.dist (n, q) + (r + s)"
-  by (import hollight DIST_TRIANGLES_LE)
-
-lemma BOUNDS_LINEAR: "(ALL n::nat. (A::nat) * n <= (B::nat) * n + (C::nat)) = (A <= B)"
-  by (import hollight BOUNDS_LINEAR)
-
-lemma BOUNDS_LINEAR_0: "(ALL n::nat. (A::nat) * n <= (B::nat)) = (A = (0::nat))"
-  by (import hollight BOUNDS_LINEAR_0)
-
-lemma BOUNDS_DIVIDED: "(EX B::nat. ALL n::nat. (P::nat => nat) n <= B) =
-(EX (x::nat) B::nat. ALL n::nat. n * P n <= x * n + B)"
-  by (import hollight BOUNDS_DIVIDED)
-
-lemma BOUNDS_NOTZERO: "(P::nat => nat => nat) (0::nat) (0::nat) = (0::nat) &
-(ALL (m::nat) n::nat. P m n <= (A::nat) * (m + n) + (B::nat))
-==> EX x::nat. ALL (m::nat) n::nat. P m n <= x * (m + n)"
-  by (import hollight BOUNDS_NOTZERO)
-
-lemma BOUNDS_IGNORE: "(EX B::nat. ALL i::nat. (P::nat => nat) i <= (Q::nat => nat) i + B) =
-(EX (x::nat) N::nat. ALL i>=N. P i <= Q i + x)"
-  by (import hollight BOUNDS_IGNORE)
-
-definition
-  is_nadd :: "(nat => nat) => bool"  where
-  "is_nadd ==
-%u. EX B. ALL m n. hollight.dist (m * u n, n * u m) <= B * (m + n)"
-
-lemma DEF_is_nadd: "is_nadd =
-(%u. EX B. ALL m n. hollight.dist (m * u n, n * u m) <= B * (m + n))"
-  by (import hollight DEF_is_nadd)
-
-lemma is_nadd_0: "is_nadd (%n. 0)"
-  by (import hollight is_nadd_0)
-
-typedef (open) nadd = "Collect is_nadd"  morphisms "dest_nadd" "mk_nadd"
-  apply (rule light_ex_imp_nonempty[where t="%n. NUMERAL 0"])
-  by (import hollight TYDEF_nadd)
-
-syntax
-  dest_nadd :: _ 
-
-syntax
-  mk_nadd :: _ 
-
-lemmas "TYDEF_nadd_@intern" = typedef_hol2hollight 
-  [where a="a :: nadd" and r=r ,
-   OF type_definition_nadd]
-
-lemma NADD_CAUCHY: "EX xa.
-   ALL xb xc.
-      hollight.dist (xb * dest_nadd x xc, xc * dest_nadd x xb)
-      <= xa * (xb + xc)"
-  by (import hollight NADD_CAUCHY)
-
-lemma NADD_BOUND: "EX xa B. ALL n. dest_nadd x n <= xa * n + B"
-  by (import hollight NADD_BOUND)
-
-lemma NADD_MULTIPLICATIVE: "EX xa.
-   ALL m n.
-      hollight.dist (dest_nadd x (m * n), m * dest_nadd x n) <= xa * m + xa"
-  by (import hollight NADD_MULTIPLICATIVE)
-
-lemma NADD_ADDITIVE: "EX xa.
-   ALL m n.
-      hollight.dist (dest_nadd x (m + n), dest_nadd x m + dest_nadd x n)
-      <= xa"
-  by (import hollight NADD_ADDITIVE)
-
-lemma NADD_SUC: "EX xa. ALL n. hollight.dist (dest_nadd x (Suc n), dest_nadd x n) <= xa"
-  by (import hollight NADD_SUC)
-
-lemma NADD_DIST_LEMMA: "EX xa. ALL m n. hollight.dist (dest_nadd x (m + n), dest_nadd x m) <= xa * n"
-  by (import hollight NADD_DIST_LEMMA)
-
-lemma NADD_DIST: "EX xa.
-   ALL m n.
-      hollight.dist (dest_nadd x m, dest_nadd x n)
-      <= xa * hollight.dist (m, n)"
-  by (import hollight NADD_DIST)
-
-lemma NADD_ALTMUL: "EX A B.
-   ALL n.
-      hollight.dist
-       (n * dest_nadd x (dest_nadd y n), dest_nadd x n * dest_nadd y n)
-      <= A * n + B"
-  by (import hollight NADD_ALTMUL)
-
-definition
-  nadd_eq :: "nadd => nadd => bool"  where
-  "nadd_eq ==
-%u ua. EX B. ALL n. hollight.dist (dest_nadd u n, dest_nadd ua n) <= B"
-
-lemma DEF_nadd_eq: "nadd_eq =
-(%u ua. EX B. ALL n. hollight.dist (dest_nadd u n, dest_nadd ua n) <= B)"
-  by (import hollight DEF_nadd_eq)
-
-lemma NADD_EQ_REFL: "nadd_eq x x"
-  by (import hollight NADD_EQ_REFL)
-
-lemma NADD_EQ_SYM: "nadd_eq x y = nadd_eq y x"
-  by (import hollight NADD_EQ_SYM)
-
-lemma NADD_EQ_TRANS: "nadd_eq x y & nadd_eq y z ==> nadd_eq x z"
-  by (import hollight NADD_EQ_TRANS)
-
-definition
-  nadd_of_num :: "nat => nadd"  where
-  "nadd_of_num == %u. mk_nadd (op * u)"
-
-lemma DEF_nadd_of_num: "nadd_of_num = (%u. mk_nadd (op * u))"
-  by (import hollight DEF_nadd_of_num)
-
-lemma NADD_OF_NUM: "dest_nadd (nadd_of_num x) = op * x"
-  by (import hollight NADD_OF_NUM)
-
-lemma NADD_OF_NUM_WELLDEF: "m = n ==> nadd_eq (nadd_of_num m) (nadd_of_num n)"
-  by (import hollight NADD_OF_NUM_WELLDEF)
-
-lemma NADD_OF_NUM_EQ: "nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)"
-  by (import hollight NADD_OF_NUM_EQ)
-
-definition
-  nadd_le :: "nadd => nadd => bool"  where
-  "nadd_le == %u ua. EX B. ALL n. dest_nadd u n <= dest_nadd ua n + B"
-
-lemma DEF_nadd_le: "nadd_le = (%u ua. EX B. ALL n. dest_nadd u n <= dest_nadd ua n + B)"
-  by (import hollight DEF_nadd_le)
-
-lemma NADD_LE_WELLDEF_LEMMA: "nadd_eq x x' & nadd_eq y y' & nadd_le x y ==> nadd_le x' y'"
-  by (import hollight NADD_LE_WELLDEF_LEMMA)
-
-lemma NADD_LE_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_le x y = nadd_le x' y'"
-  by (import hollight NADD_LE_WELLDEF)
-
-lemma NADD_LE_REFL: "nadd_le x x"
-  by (import hollight NADD_LE_REFL)
-
-lemma NADD_LE_TRANS: "nadd_le x y & nadd_le y z ==> nadd_le x z"
-  by (import hollight NADD_LE_TRANS)
-
-lemma NADD_LE_ANTISYM: "(nadd_le x y & nadd_le y x) = nadd_eq x y"
-  by (import hollight NADD_LE_ANTISYM)
-
-lemma NADD_LE_TOTAL_LEMMA: "~ nadd_le x y ==> EX n. n ~= 0 & dest_nadd y n + B < dest_nadd x n"
-  by (import hollight NADD_LE_TOTAL_LEMMA)
-
-lemma NADD_LE_TOTAL: "nadd_le x y | nadd_le y x"
-  by (import hollight NADD_LE_TOTAL)
-
-lemma NADD_ARCH: "EX xa. nadd_le x (nadd_of_num xa)"
-  by (import hollight NADD_ARCH)
-
-lemma NADD_OF_NUM_LE: "nadd_le (nadd_of_num m) (nadd_of_num n) = (m <= n)"
-  by (import hollight NADD_OF_NUM_LE)
-
-definition
-  nadd_add :: "nadd => nadd => nadd"  where
-  "nadd_add == %u ua. mk_nadd (%n. dest_nadd u n + dest_nadd ua n)"
-
-lemma DEF_nadd_add: "nadd_add = (%u ua. mk_nadd (%n. dest_nadd u n + dest_nadd ua n))"
-  by (import hollight DEF_nadd_add)
-
-lemma NADD_ADD: "dest_nadd (nadd_add x y) = (%n. dest_nadd x n + dest_nadd y n)"
-  by (import hollight NADD_ADD)
-
-lemma NADD_ADD_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_eq (nadd_add x y) (nadd_add x' y')"
-  by (import hollight NADD_ADD_WELLDEF)
-
-lemma NADD_ADD_SYM: "nadd_eq (nadd_add x y) (nadd_add y x)"
-  by (import hollight NADD_ADD_SYM)
-
-lemma NADD_ADD_ASSOC: "nadd_eq (nadd_add x (nadd_add y z)) (nadd_add (nadd_add x y) z)"
-  by (import hollight NADD_ADD_ASSOC)
-
-lemma NADD_ADD_LID: "nadd_eq (nadd_add (nadd_of_num 0) x) x"
-  by (import hollight NADD_ADD_LID)
-
-lemma NADD_ADD_LCANCEL: "nadd_eq (nadd_add x y) (nadd_add x z) ==> nadd_eq y z"
-  by (import hollight NADD_ADD_LCANCEL)
-
-lemma NADD_LE_ADD: "nadd_le x (nadd_add x y)"
-  by (import hollight NADD_LE_ADD)
-
-lemma NADD_LE_EXISTS: "nadd_le x y ==> EX d. nadd_eq y (nadd_add x d)"
-  by (import hollight NADD_LE_EXISTS)
-
-lemma NADD_OF_NUM_ADD: "nadd_eq (nadd_add (nadd_of_num x) (nadd_of_num xa)) (nadd_of_num (x + xa))"
-  by (import hollight NADD_OF_NUM_ADD)
-
-definition
-  nadd_mul :: "nadd => nadd => nadd"  where
-  "nadd_mul == %u ua. mk_nadd (%n. dest_nadd u (dest_nadd ua n))"
-
-lemma DEF_nadd_mul: "nadd_mul = (%u ua. mk_nadd (%n. dest_nadd u (dest_nadd ua n)))"
-  by (import hollight DEF_nadd_mul)
-
-lemma NADD_MUL: "dest_nadd (nadd_mul x y) = (%n. dest_nadd x (dest_nadd y n))"
-  by (import hollight NADD_MUL)
-
-lemma NADD_MUL_SYM: "nadd_eq (nadd_mul x y) (nadd_mul y x)"
-  by (import hollight NADD_MUL_SYM)
-
-lemma NADD_MUL_ASSOC: "nadd_eq (nadd_mul x (nadd_mul y z)) (nadd_mul (nadd_mul x y) z)"
-  by (import hollight NADD_MUL_ASSOC)
-
-lemma NADD_MUL_LID: "nadd_eq (nadd_mul (nadd_of_num 1) x) x"
-  by (import hollight NADD_MUL_LID)
-
-lemma NADD_LDISTRIB: "nadd_eq (nadd_mul x (nadd_add y z)) (nadd_add (nadd_mul x y) (nadd_mul x z))"
-  by (import hollight NADD_LDISTRIB)
-
-lemma NADD_MUL_WELLDEF_LEMMA: "nadd_eq y y' ==> nadd_eq (nadd_mul x y) (nadd_mul x y')"
-  by (import hollight NADD_MUL_WELLDEF_LEMMA)
-
-lemma NADD_MUL_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_eq (nadd_mul x y) (nadd_mul x' y')"
-  by (import hollight NADD_MUL_WELLDEF)
-
-lemma NADD_OF_NUM_MUL: "nadd_eq (nadd_mul (nadd_of_num x) (nadd_of_num xa)) (nadd_of_num (x * xa))"
-  by (import hollight NADD_OF_NUM_MUL)
-
-lemma NADD_LE_0: "nadd_le (nadd_of_num 0) x"
-  by (import hollight NADD_LE_0)
-
-lemma NADD_EQ_IMP_LE: "nadd_eq x y ==> nadd_le x y"
-  by (import hollight NADD_EQ_IMP_LE)
-
-lemma NADD_LE_LMUL: "nadd_le y z ==> nadd_le (nadd_mul x y) (nadd_mul x z)"
-  by (import hollight NADD_LE_LMUL)
-
-lemma NADD_LE_RMUL: "nadd_le x y ==> nadd_le (nadd_mul x z) (nadd_mul y z)"
-  by (import hollight NADD_LE_RMUL)
-
-lemma NADD_LE_RADD: "nadd_le (nadd_add x z) (nadd_add y z) = nadd_le x y"
-  by (import hollight NADD_LE_RADD)
-
-lemma NADD_LE_LADD: "nadd_le (nadd_add x y) (nadd_add x z) = nadd_le y z"
-  by (import hollight NADD_LE_LADD)
-
-lemma NADD_RDISTRIB: "nadd_eq (nadd_mul (nadd_add x y) z) (nadd_add (nadd_mul x z) (nadd_mul y z))"
-  by (import hollight NADD_RDISTRIB)
-
-lemma NADD_ARCH_MULT: "~ nadd_eq x (nadd_of_num 0)
-==> EX xa. nadd_le (nadd_of_num k) (nadd_mul (nadd_of_num xa) x)"
-  by (import hollight NADD_ARCH_MULT)
-
-lemma NADD_ARCH_ZERO: "(!!n. nadd_le (nadd_mul (nadd_of_num n) x) k) ==> nadd_eq x (nadd_of_num 0)"
-  by (import hollight NADD_ARCH_ZERO)
-
-lemma NADD_ARCH_LEMMA: "(!!n. nadd_le (nadd_mul (nadd_of_num n) x)
-       (nadd_add (nadd_mul (nadd_of_num n) y) z))
-==> nadd_le x y"
-  by (import hollight NADD_ARCH_LEMMA)
-
-lemma NADD_COMPLETE: "Ex P & (EX M. ALL x. P x --> nadd_le x M)
-==> EX M. (ALL x. P x --> nadd_le x M) &
-          (ALL M'. (ALL x. P x --> nadd_le x M') --> nadd_le M M')"
-  by (import hollight NADD_COMPLETE)
-
-lemma NADD_UBOUND: "EX xa N. ALL n>=N. dest_nadd x n <= xa * n"
-  by (import hollight NADD_UBOUND)
-
-lemma NADD_NONZERO: "~ nadd_eq x (nadd_of_num 0) ==> EX N. ALL n>=N. dest_nadd x n ~= 0"
-  by (import hollight NADD_NONZERO)
-
-lemma NADD_LBOUND: "~ nadd_eq x (nadd_of_num 0) ==> EX A N. ALL n>=N. n <= A * dest_nadd x n"
-  by (import hollight NADD_LBOUND)
-
-definition
-  nadd_rinv :: "nadd => nat => nat"  where
-  "nadd_rinv == %u n. n * n div dest_nadd u n"
-
-lemma DEF_nadd_rinv: "nadd_rinv = (%u n. n * n div dest_nadd u n)"
-  by (import hollight DEF_nadd_rinv)
-
-lemma NADD_MUL_LINV_LEMMA0: "~ nadd_eq x (nadd_of_num 0) ==> EX xa B. ALL i. nadd_rinv x i <= xa * i + B"
-  by (import hollight NADD_MUL_LINV_LEMMA0)
-
-lemma NADD_MUL_LINV_LEMMA1: "dest_nadd x n ~= 0
-==> hollight.dist (dest_nadd x n * nadd_rinv x n, n * n) <= dest_nadd x n"
-  by (import hollight NADD_MUL_LINV_LEMMA1)
-
-lemma NADD_MUL_LINV_LEMMA2: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL n>=N.
-             hollight.dist (dest_nadd x n * nadd_rinv x n, n * n)
-             <= dest_nadd x n"
-  by (import hollight NADD_MUL_LINV_LEMMA2)
-
-lemma NADD_MUL_LINV_LEMMA3: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL m n.
-             N <= n -->
-             hollight.dist
-              (m * (dest_nadd x m * (dest_nadd x n * nadd_rinv x n)),
-               m * (dest_nadd x m * (n * n)))
-             <= m * (dest_nadd x m * dest_nadd x n)"
-  by (import hollight NADD_MUL_LINV_LEMMA3)
-
-lemma NADD_MUL_LINV_LEMMA4: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL m n.
-             N <= m & N <= n -->
-             dest_nadd x m * dest_nadd x n *
-             hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
-             <= m * n *
-                hollight.dist (m * dest_nadd x n, n * dest_nadd x m) +
-                dest_nadd x m * dest_nadd x n * (m + n)"
-  by (import hollight NADD_MUL_LINV_LEMMA4)
-
-lemma NADD_MUL_LINV_LEMMA5: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
-       ALL m n.
-          N <= m & N <= n -->
-          dest_nadd x m * dest_nadd x n *
-          hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
-          <= B * (m * n * (m + n))"
-  by (import hollight NADD_MUL_LINV_LEMMA5)
-
-lemma NADD_MUL_LINV_LEMMA6: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
-       ALL m n.
-          N <= m & N <= n -->
-          m * n * hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
-          <= B * (m * n * (m + n))"
-  by (import hollight NADD_MUL_LINV_LEMMA6)
-
-lemma NADD_MUL_LINV_LEMMA7: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
-       ALL m n.
-          N <= m & N <= n -->
-          hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
-          <= B * (m + n)"
-  by (import hollight NADD_MUL_LINV_LEMMA7)
-
-lemma NADD_MUL_LINV_LEMMA7a: "~ nadd_eq x (nadd_of_num 0)
-==> EX A B.
-       ALL m n.
-          m <= N -->
-          hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m) <= A * n + B"
-  by (import hollight NADD_MUL_LINV_LEMMA7a)
-
-lemma NADD_MUL_LINV_LEMMA8: "~ nadd_eq x (nadd_of_num 0)
-==> EX B. ALL m n.
-             hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
-             <= B * (m + n)"
-  by (import hollight NADD_MUL_LINV_LEMMA8)
-
-definition
-  nadd_inv :: "nadd => nadd"  where
-  "nadd_inv ==
-%u. if nadd_eq u (nadd_of_num 0) then nadd_of_num 0
-    else mk_nadd (nadd_rinv u)"
-
-lemma DEF_nadd_inv: "nadd_inv =
-(%u. if nadd_eq u (nadd_of_num 0) then nadd_of_num 0
-     else mk_nadd (nadd_rinv u))"
-  by (import hollight DEF_nadd_inv)
-
-lemma NADD_INV: "dest_nadd (nadd_inv x) =
-(if nadd_eq x (nadd_of_num 0) then %n. 0 else nadd_rinv x)"
-  by (import hollight NADD_INV)
-
-lemma NADD_MUL_LINV: "~ nadd_eq x (nadd_of_num 0)
-==> nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num 1)"
-  by (import hollight NADD_MUL_LINV)
-
-lemma NADD_INV_0: "nadd_eq (nadd_inv (nadd_of_num 0)) (nadd_of_num 0)"
-  by (import hollight NADD_INV_0)
-
-lemma NADD_INV_WELLDEF: "nadd_eq x y ==> nadd_eq (nadd_inv x) (nadd_inv y)"
-  by (import hollight NADD_INV_WELLDEF)
-
-typedef (open) hreal = "{s. EX x. s = nadd_eq x}"  morphisms "dest_hreal" "mk_hreal"
-  apply (rule light_ex_imp_nonempty[where t="nadd_eq x"])
-  by (import hollight TYDEF_hreal)
-
-syntax
-  dest_hreal :: _ 
-
-syntax
-  mk_hreal :: _ 
-
-lemmas "TYDEF_hreal_@intern" = typedef_hol2hollight 
-  [where a="a :: hreal" and r=r ,
-   OF type_definition_hreal]
-
-definition
-  hreal_of_num :: "nat => hreal"  where
-  "hreal_of_num == %m. mk_hreal (nadd_eq (nadd_of_num m))"
-
-lemma DEF_hreal_of_num: "hreal_of_num = (%m. mk_hreal (nadd_eq (nadd_of_num m)))"
-  by (import hollight DEF_hreal_of_num)
-
-definition
-  hreal_add :: "hreal => hreal => hreal"  where
-  "hreal_add ==
-%x y. mk_hreal
-       (%u. EX xa ya.
-               nadd_eq (nadd_add xa ya) u &
-               dest_hreal x xa & dest_hreal y ya)"
-
-lemma DEF_hreal_add: "hreal_add =
-(%x y. mk_hreal
-        (%u. EX xa ya.
-                nadd_eq (nadd_add xa ya) u &
-                dest_hreal x xa & dest_hreal y ya))"
-  by (import hollight DEF_hreal_add)
-
-definition
-  hreal_mul :: "hreal => hreal => hreal"  where
-  "hreal_mul ==
-%x y. mk_hreal
-       (%u. EX xa ya.
-               nadd_eq (nadd_mul xa ya) u &
-               dest_hreal x xa & dest_hreal y ya)"
-
-lemma DEF_hreal_mul: "hreal_mul =
-(%x y. mk_hreal
-        (%u. EX xa ya.
-                nadd_eq (nadd_mul xa ya) u &
-                dest_hreal x xa & dest_hreal y ya))"
-  by (import hollight DEF_hreal_mul)
-
-definition
-  hreal_le :: "hreal => hreal => bool"  where
-  "hreal_le ==
-%x y. SOME u.
-         EX xa ya. nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya"
-
-lemma DEF_hreal_le: "hreal_le =
-(%x y. SOME u.
-          EX xa ya. nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)"
-  by (import hollight DEF_hreal_le)
-
-definition
-  hreal_inv :: "hreal => hreal"  where
-  "hreal_inv ==
-%x. mk_hreal (%u. EX xa. nadd_eq (nadd_inv xa) u & dest_hreal x xa)"
-
-lemma DEF_hreal_inv: "hreal_inv =
-(%x. mk_hreal (%u. EX xa. nadd_eq (nadd_inv xa) u & dest_hreal x xa))"
-  by (import hollight DEF_hreal_inv)
-
-lemma HREAL_LE_EXISTS_DEF: "hreal_le m n = (EX d. n = hreal_add m d)"
-  by (import hollight HREAL_LE_EXISTS_DEF)
-
-lemma HREAL_EQ_ADD_LCANCEL: "(hreal_add m n = hreal_add m p) = (n = p)"
-  by (import hollight HREAL_EQ_ADD_LCANCEL)
-
-lemma HREAL_EQ_ADD_RCANCEL: "(hreal_add x xb = hreal_add xa xb) = (x = xa)"
-  by (import hollight HREAL_EQ_ADD_RCANCEL)
-
-lemma HREAL_LE_ADD_LCANCEL: "hreal_le (hreal_add x xa) (hreal_add x xb) = hreal_le xa xb"
-  by (import hollight HREAL_LE_ADD_LCANCEL)
-
-lemma HREAL_LE_ADD_RCANCEL: "hreal_le (hreal_add x xb) (hreal_add xa xb) = hreal_le x xa"
-  by (import hollight HREAL_LE_ADD_RCANCEL)
-
-lemma HREAL_ADD_RID: "hreal_add x (hreal_of_num 0) = x"
-  by (import hollight HREAL_ADD_RID)
-
-lemma HREAL_ADD_RDISTRIB: "hreal_mul (hreal_add x xa) xb = hreal_add (hreal_mul x xb) (hreal_mul xa xb)"
-  by (import hollight HREAL_ADD_RDISTRIB)
-
-lemma HREAL_MUL_LZERO: "hreal_mul (hreal_of_num 0) m = hreal_of_num 0"
-  by (import hollight HREAL_MUL_LZERO)
-
-lemma HREAL_MUL_RZERO: "hreal_mul x (hreal_of_num 0) = hreal_of_num 0"
-  by (import hollight HREAL_MUL_RZERO)
-
-lemma HREAL_ADD_AC: "hreal_add m n = hreal_add n m &
-hreal_add (hreal_add m n) p = hreal_add m (hreal_add n p) &
-hreal_add m (hreal_add n p) = hreal_add n (hreal_add m p)"
-  by (import hollight HREAL_ADD_AC)
-
-lemma HREAL_LE_ADD2: "hreal_le a b & hreal_le c d ==> hreal_le (hreal_add a c) (hreal_add b d)"
-  by (import hollight HREAL_LE_ADD2)
-
-lemma HREAL_LE_MUL_RCANCEL_IMP: "hreal_le a b ==> hreal_le (hreal_mul a c) (hreal_mul b c)"
-  by (import hollight HREAL_LE_MUL_RCANCEL_IMP)
-
-definition
-  treal_of_num :: "nat => hreal * hreal"  where
-  "treal_of_num == %u. (hreal_of_num u, hreal_of_num 0)"
-
-lemma DEF_treal_of_num: "treal_of_num = (%u. (hreal_of_num u, hreal_of_num 0))"
-  by (import hollight DEF_treal_of_num)
-
-definition
-  treal_neg :: "hreal * hreal => hreal * hreal"  where
-  "treal_neg == %u. (snd u, fst u)"
-
-lemma DEF_treal_neg: "treal_neg = (%u. (snd u, fst u))"
-  by (import hollight DEF_treal_neg)
-
-definition
-  treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal"  where
-  "treal_add == %u ua. (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))"
-
-lemma DEF_treal_add: "treal_add =
-(%u ua. (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))"
-  by (import hollight DEF_treal_add)
-
-definition
-  treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal"  where
-  "treal_mul ==
-%u ua.
-   (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
-    hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua)))"
-
-lemma DEF_treal_mul: "treal_mul =
-(%u ua.
-    (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
-     hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))"
-  by (import hollight DEF_treal_mul)
-
-definition
-  treal_le :: "hreal * hreal => hreal * hreal => bool"  where
-  "treal_le ==
-%u ua. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))"
-
-lemma DEF_treal_le: "treal_le =
-(%u ua. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))"
-  by (import hollight DEF_treal_le)
-
-definition
-  treal_inv :: "hreal * hreal => hreal * hreal"  where
-  "treal_inv ==
-%u. if fst u = snd u then (hreal_of_num 0, hreal_of_num 0)
-    else if hreal_le (snd u) (fst u)
-         then (hreal_inv (SOME d. fst u = hreal_add (snd u) d),
-               hreal_of_num 0)
-         else (hreal_of_num 0,
-               hreal_inv (SOME d. snd u = hreal_add (fst u) d))"
-
-lemma DEF_treal_inv: "treal_inv =
-(%u. if fst u = snd u then (hreal_of_num 0, hreal_of_num 0)
-     else if hreal_le (snd u) (fst u)
-          then (hreal_inv (SOME d. fst u = hreal_add (snd u) d),
-                hreal_of_num 0)
-          else (hreal_of_num 0,
-                hreal_inv (SOME d. snd u = hreal_add (fst u) d)))"
-  by (import hollight DEF_treal_inv)
-
-definition
-  treal_eq :: "hreal * hreal => hreal * hreal => bool"  where
-  "treal_eq == %u ua. hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)"
-
-lemma DEF_treal_eq: "treal_eq = (%u ua. hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u))"
-  by (import hollight DEF_treal_eq)
-
-lemma TREAL_EQ_REFL: "treal_eq x x"
-  by (import hollight TREAL_EQ_REFL)
-
-lemma TREAL_EQ_SYM: "treal_eq x y = treal_eq y x"
-  by (import hollight TREAL_EQ_SYM)
-
-lemma TREAL_EQ_TRANS: "treal_eq x y & treal_eq y z ==> treal_eq x z"
-  by (import hollight TREAL_EQ_TRANS)
-
-lemma TREAL_EQ_AP: "x = xa ==> treal_eq x xa"
-  by (import hollight TREAL_EQ_AP)
-
-lemma TREAL_OF_NUM_EQ: "treal_eq (treal_of_num x) (treal_of_num xa) = (x = xa)"
-  by (import hollight TREAL_OF_NUM_EQ)
-
-lemma TREAL_OF_NUM_LE: "treal_le (treal_of_num x) (treal_of_num xa) = (x <= xa)"
-  by (import hollight TREAL_OF_NUM_LE)
-
-lemma TREAL_OF_NUM_ADD: "treal_eq (treal_add (treal_of_num x) (treal_of_num xa))
- (treal_of_num (x + xa))"
-  by (import hollight TREAL_OF_NUM_ADD)
-
-lemma TREAL_OF_NUM_MUL: "treal_eq (treal_mul (treal_of_num x) (treal_of_num xa))
- (treal_of_num (x * xa))"
-  by (import hollight TREAL_OF_NUM_MUL)
-
-lemma TREAL_ADD_SYM_EQ: "treal_add x y = treal_add y x"
-  by (import hollight TREAL_ADD_SYM_EQ)
-
-lemma TREAL_MUL_SYM_EQ: "treal_mul x y = treal_mul y x"
-  by (import hollight TREAL_MUL_SYM_EQ)
-
-lemma TREAL_ADD_SYM: "treal_eq (treal_add x y) (treal_add y x)"
-  by (import hollight TREAL_ADD_SYM)
-
-lemma TREAL_ADD_ASSOC: "treal_eq (treal_add x (treal_add y z)) (treal_add (treal_add x y) z)"
-  by (import hollight TREAL_ADD_ASSOC)
-
-lemma TREAL_ADD_LID: "treal_eq (treal_add (treal_of_num 0) x) x"
-  by (import hollight TREAL_ADD_LID)
-
-lemma TREAL_ADD_LINV: "treal_eq (treal_add (treal_neg x) x) (treal_of_num 0)"
-  by (import hollight TREAL_ADD_LINV)
-
-lemma TREAL_MUL_SYM: "treal_eq (treal_mul x xa) (treal_mul xa x)"
-  by (import hollight TREAL_MUL_SYM)
-
-lemma TREAL_MUL_ASSOC: "treal_eq (treal_mul x (treal_mul y z)) (treal_mul (treal_mul x y) z)"
-  by (import hollight TREAL_MUL_ASSOC)
-
-lemma TREAL_MUL_LID: "treal_eq (treal_mul (treal_of_num 1) x) x"
-  by (import hollight TREAL_MUL_LID)
-
-lemma TREAL_ADD_LDISTRIB: "treal_eq (treal_mul x (treal_add y z))
- (treal_add (treal_mul x y) (treal_mul x z))"
-  by (import hollight TREAL_ADD_LDISTRIB)
-
-lemma TREAL_LE_REFL: "treal_le x x"
-  by (import hollight TREAL_LE_REFL)
-
-lemma TREAL_LE_ANTISYM: "(treal_le x y & treal_le y x) = treal_eq x y"
-  by (import hollight TREAL_LE_ANTISYM)
-
-lemma TREAL_LE_TRANS: "treal_le x y & treal_le y z ==> treal_le x z"
-  by (import hollight TREAL_LE_TRANS)
-
-lemma TREAL_LE_TOTAL: "treal_le x y | treal_le y x"
-  by (import hollight TREAL_LE_TOTAL)
-
-lemma TREAL_LE_LADD_IMP: "treal_le y z ==> treal_le (treal_add x y) (treal_add x z)"
-  by (import hollight TREAL_LE_LADD_IMP)
-
-lemma TREAL_LE_MUL: "treal_le (treal_of_num 0) x & treal_le (treal_of_num 0) y
-==> treal_le (treal_of_num 0) (treal_mul x y)"
-  by (import hollight TREAL_LE_MUL)
-
-lemma TREAL_INV_0: "treal_eq (treal_inv (treal_of_num 0)) (treal_of_num 0)"
-  by (import hollight TREAL_INV_0)
-
-lemma TREAL_MUL_LINV: "~ treal_eq x (treal_of_num 0)
-==> treal_eq (treal_mul (treal_inv x) x) (treal_of_num 1)"
-  by (import hollight TREAL_MUL_LINV)
-
-lemma TREAL_OF_NUM_WELLDEF: "m = n ==> treal_eq (treal_of_num m) (treal_of_num n)"
-  by (import hollight TREAL_OF_NUM_WELLDEF)
-
-lemma TREAL_NEG_WELLDEF: "treal_eq x1 x2 ==> treal_eq (treal_neg x1) (treal_neg x2)"
-  by (import hollight TREAL_NEG_WELLDEF)
-
-lemma TREAL_ADD_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_add x1 y) (treal_add x2 y)"
-  by (import hollight TREAL_ADD_WELLDEFR)
-
-lemma TREAL_ADD_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
-==> treal_eq (treal_add x1 y1) (treal_add x2 y2)"
-  by (import hollight TREAL_ADD_WELLDEF)
-
-lemma TREAL_MUL_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_mul x1 y) (treal_mul x2 y)"
-  by (import hollight TREAL_MUL_WELLDEFR)
-
-lemma TREAL_MUL_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
-==> treal_eq (treal_mul x1 y1) (treal_mul x2 y2)"
-  by (import hollight TREAL_MUL_WELLDEF)
-
-lemma TREAL_EQ_IMP_LE: "treal_eq x y ==> treal_le x y"
-  by (import hollight TREAL_EQ_IMP_LE)
-
-lemma TREAL_LE_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2 ==> treal_le x1 y1 = treal_le x2 y2"
-  by (import hollight TREAL_LE_WELLDEF)
-
-lemma TREAL_INV_WELLDEF: "treal_eq x y ==> treal_eq (treal_inv x) (treal_inv y)"
-  by (import hollight TREAL_INV_WELLDEF)
-
-typedef (open) real = "{s. EX x. s = treal_eq x}"  morphisms "dest_real" "mk_real"
-  apply (rule light_ex_imp_nonempty[where t="treal_eq x"])
-  by (import hollight TYDEF_real)
-
-syntax
-  dest_real :: _ 
-
-syntax
-  mk_real :: _ 
-
-lemmas "TYDEF_real_@intern" = typedef_hol2hollight 
-  [where a="a :: hollight.real" and r=r ,
-   OF type_definition_real]
-
-definition
-  real_of_num :: "nat => hollight.real"  where
-  "real_of_num == %m. mk_real (treal_eq (treal_of_num m))"
-
-lemma DEF_real_of_num: "real_of_num = (%m. mk_real (treal_eq (treal_of_num m)))"
-  by (import hollight DEF_real_of_num)
-
-definition
-  real_neg :: "hollight.real => hollight.real"  where
-  "real_neg ==
-%x1. mk_real (%u. EX x1a. treal_eq (treal_neg x1a) u & dest_real x1 x1a)"
-
-lemma DEF_real_neg: "real_neg =
-(%x1. mk_real (%u. EX x1a. treal_eq (treal_neg x1a) u & dest_real x1 x1a))"
-  by (import hollight DEF_real_neg)
-
-definition
-  real_add :: "hollight.real => hollight.real => hollight.real"  where
-  "real_add ==
-%x1 y1.
-   mk_real
-    (%u. EX x1a y1a.
-            treal_eq (treal_add x1a y1a) u &
-            dest_real x1 x1a & dest_real y1 y1a)"
-
-lemma DEF_real_add: "real_add =
-(%x1 y1.
-    mk_real
-     (%u. EX x1a y1a.
-             treal_eq (treal_add x1a y1a) u &
-             dest_real x1 x1a & dest_real y1 y1a))"
-  by (import hollight DEF_real_add)
-
-definition
-  real_mul :: "hollight.real => hollight.real => hollight.real"  where
-  "real_mul ==
-%x1 y1.
-   mk_real
-    (%u. EX x1a y1a.
-            treal_eq (treal_mul x1a y1a) u &
-            dest_real x1 x1a & dest_real y1 y1a)"
-
-lemma DEF_real_mul: "real_mul =
-(%x1 y1.
-    mk_real
-     (%u. EX x1a y1a.
-             treal_eq (treal_mul x1a y1a) u &
-             dest_real x1 x1a & dest_real y1 y1a))"
-  by (import hollight DEF_real_mul)
-
-definition
-  real_le :: "hollight.real => hollight.real => bool"  where
-  "real_le ==
-%x1 y1.
-   SOME u.
-      EX x1a y1a. treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a"
-
-lemma DEF_real_le: "real_le =
-(%x1 y1.
-    SOME u.
-       EX x1a y1a.
-          treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)"
-  by (import hollight DEF_real_le)
-
-definition
-  real_inv :: "hollight.real => hollight.real"  where
-  "real_inv ==
-%x. mk_real (%u. EX xa. treal_eq (treal_inv xa) u & dest_real x xa)"
-
-lemma DEF_real_inv: "real_inv =
-(%x. mk_real (%u. EX xa. treal_eq (treal_inv xa) u & dest_real x xa))"
-  by (import hollight DEF_real_inv)
-
-definition
-  real_sub :: "hollight.real => hollight.real => hollight.real"  where
-  "real_sub == %u ua. real_add u (real_neg ua)"
-
-lemma DEF_real_sub: "real_sub = (%u ua. real_add u (real_neg ua))"
-  by (import hollight DEF_real_sub)
-
-definition
-  real_lt :: "hollight.real => hollight.real => bool"  where
-  "real_lt == %u ua. ~ real_le ua u"
-
-lemma DEF_real_lt: "real_lt = (%u ua. ~ real_le ua u)"
-  by (import hollight DEF_real_lt)
-
-definition
-  real_ge :: "hollight.real => hollight.real => bool"  where
-  "real_ge == %u ua. real_le ua u"
-
-lemma DEF_real_ge: "real_ge = (%u ua. real_le ua u)"
-  by (import hollight DEF_real_ge)
-
-definition
-  real_gt :: "hollight.real => hollight.real => bool"  where
-  "real_gt == %u ua. real_lt ua u"
-
-lemma DEF_real_gt: "real_gt = (%u ua. real_lt ua u)"
-  by (import hollight DEF_real_gt)
-
-definition
-  real_abs :: "hollight.real => hollight.real"  where
-  "real_abs == %u. if real_le (real_of_num 0) u then u else real_neg u"
-
-lemma DEF_real_abs: "real_abs = (%u. if real_le (real_of_num 0) u then u else real_neg u)"
-  by (import hollight DEF_real_abs)
-
-definition
-  real_pow :: "hollight.real => nat => hollight.real"  where
-  "real_pow ==
-SOME real_pow.
-   (ALL x. real_pow x 0 = real_of_num 1) &
-   (ALL x n. real_pow x (Suc n) = real_mul x (real_pow x n))"
-
-lemma DEF_real_pow: "real_pow =
-(SOME real_pow.
-    (ALL x. real_pow x 0 = real_of_num 1) &
-    (ALL x n. real_pow x (Suc n) = real_mul x (real_pow x n)))"
-  by (import hollight DEF_real_pow)
-
-definition
-  real_div :: "hollight.real => hollight.real => hollight.real"  where
-  "real_div == %u ua. real_mul u (real_inv ua)"
-
-lemma DEF_real_div: "real_div = (%u ua. real_mul u (real_inv ua))"
-  by (import hollight DEF_real_div)
-
-definition
-  real_max :: "hollight.real => hollight.real => hollight.real"  where
-  "real_max == %u ua. if real_le u ua then ua else u"
-
-lemma DEF_real_max: "real_max = (%u ua. if real_le u ua then ua else u)"
-  by (import hollight DEF_real_max)
-
-definition
-  real_min :: "hollight.real => hollight.real => hollight.real"  where
-  "real_min == %u ua. if real_le u ua then u else ua"
-
-lemma DEF_real_min: "real_min = (%u ua. if real_le u ua then u else ua)"
-  by (import hollight DEF_real_min)
-
-lemma REAL_HREAL_LEMMA1: "EX x. (ALL xa. real_le (real_of_num 0) xa = (EX y. xa = x y)) &
-      (ALL y z. hreal_le y z = real_le (x y) (x z))"
-  by (import hollight REAL_HREAL_LEMMA1)
-
-lemma REAL_HREAL_LEMMA2: "EX x r.
-   (ALL xa. x (r xa) = xa) &
-   (ALL xa. real_le (real_of_num 0) xa --> r (x xa) = xa) &
-   (ALL x. real_le (real_of_num 0) (r x)) &
-   (ALL x y. hreal_le x y = real_le (r x) (r y))"
-  by (import hollight REAL_HREAL_LEMMA2)
-
-lemma REAL_COMPLETE_SOMEPOS: "(EX x. P x & real_le (real_of_num 0) x) & (EX M. ALL x. P x --> real_le x M)
-==> EX M. (ALL x. P x --> real_le x M) &
-          (ALL M'. (ALL x. P x --> real_le x M') --> real_le M M')"
-  by (import hollight REAL_COMPLETE_SOMEPOS)
-
-lemma REAL_COMPLETE: "Ex P & (EX M. ALL x. P x --> real_le x M)
-==> EX M. (ALL x. P x --> real_le x M) &
-          (ALL M'. (ALL x. P x --> real_le x M') --> real_le M M')"
-  by (import hollight REAL_COMPLETE)
-
-lemma REAL_ADD_AC: "real_add m n = real_add n m &
-real_add (real_add m n) p = real_add m (real_add n p) &
-real_add m (real_add n p) = real_add n (real_add m p)"
-  by (import hollight REAL_ADD_AC)
-
-lemma REAL_ADD_RINV: "real_add x (real_neg x) = real_of_num 0"
-  by (import hollight REAL_ADD_RINV)
-
-lemma REAL_EQ_ADD_LCANCEL: "(real_add x y = real_add x z) = (y = z)"
-  by (import hollight REAL_EQ_ADD_LCANCEL)
-
-lemma REAL_EQ_ADD_RCANCEL: "(real_add x z = real_add y z) = (x = y)"
-  by (import hollight REAL_EQ_ADD_RCANCEL)
-
-lemma REAL_MUL_RZERO: "real_mul x (real_of_num 0) = real_of_num 0"
-  by (import hollight REAL_MUL_RZERO)
-
-lemma REAL_MUL_LZERO: "real_mul (real_of_num 0) x = real_of_num 0"
-  by (import hollight REAL_MUL_LZERO)
-
-lemma REAL_NEG_NEG: "real_neg (real_neg x) = x"
-  by (import hollight REAL_NEG_NEG)
-
-lemma REAL_MUL_RNEG: "real_mul x (real_neg y) = real_neg (real_mul x y)"
-  by (import hollight REAL_MUL_RNEG)
-
-lemma REAL_MUL_LNEG: "real_mul (real_neg x) y = real_neg (real_mul x y)"
-  by (import hollight REAL_MUL_LNEG)
-
-lemma REAL_NEG_ADD: "real_neg (real_add x y) = real_add (real_neg x) (real_neg y)"
-  by (import hollight REAL_NEG_ADD)
-
-lemma REAL_ADD_RID: "real_add x (real_of_num 0) = x"
-  by (import hollight REAL_ADD_RID)
-
-lemma REAL_NEG_0: "real_neg (real_of_num 0) = real_of_num 0"
-  by (import hollight REAL_NEG_0)
-
-lemma REAL_LE_LNEG: "real_le (real_neg x) y = real_le (real_of_num 0) (real_add x y)"
-  by (import hollight REAL_LE_LNEG)
-
-lemma REAL_LE_NEG2: "real_le (real_neg x) (real_neg y) = real_le y x"
-  by (import hollight REAL_LE_NEG2)
-
-lemma REAL_LE_RNEG: "real_le x (real_neg y) = real_le (real_add x y) (real_of_num 0)"
-  by (import hollight REAL_LE_RNEG)
-
-lemma REAL_OF_NUM_POW: "real_pow (real_of_num x) n = real_of_num (x ^ n)"
-  by (import hollight REAL_OF_NUM_POW)
-
-lemma REAL_POW_NEG: "real_pow (real_neg x) n =
-(if even n then real_pow x n else real_neg (real_pow x n))"
-  by (import hollight REAL_POW_NEG)
-
-lemma REAL_ABS_NUM: "real_abs (real_of_num x) = real_of_num x"
-  by (import hollight REAL_ABS_NUM)
-
-lemma REAL_ABS_NEG: "real_abs (real_neg x) = real_abs x"
-  by (import hollight REAL_ABS_NEG)
-
-lemma REAL_LTE_TOTAL: "real_lt x xa | real_le xa x"
-  by (import hollight REAL_LTE_TOTAL)
-
-lemma REAL_LET_TOTAL: "real_le x xa | real_lt xa x"
-  by (import hollight REAL_LET_TOTAL)
-
-lemma REAL_LT_IMP_LE: "real_lt x y ==> real_le x y"
-  by (import hollight REAL_LT_IMP_LE)
-
-lemma REAL_LTE_TRANS: "real_lt x y & real_le y z ==> real_lt x z"
-  by (import hollight REAL_LTE_TRANS)
-
-lemma REAL_LET_TRANS: "real_le x y & real_lt y z ==> real_lt x z"
-  by (import hollight REAL_LET_TRANS)
-
-lemma REAL_LT_TRANS: "real_lt x y & real_lt y z ==> real_lt x z"
-  by (import hollight REAL_LT_TRANS)
-
-lemma REAL_LE_ADD: "real_le (real_of_num 0) x & real_le (real_of_num 0) y
-==> real_le (real_of_num 0) (real_add x y)"
-  by (import hollight REAL_LE_ADD)
-
-lemma REAL_LTE_ANTISYM: "~ (real_lt x y & real_le y x)"
-  by (import hollight REAL_LTE_ANTISYM)
-
-lemma REAL_SUB_LE: "real_le (real_of_num 0) (real_sub x xa) = real_le xa x"
-  by (import hollight REAL_SUB_LE)
-
-lemma REAL_NEG_SUB: "real_neg (real_sub x xa) = real_sub xa x"
-  by (import hollight REAL_NEG_SUB)
-
-lemma REAL_LE_LT: "real_le x xa = (real_lt x xa | x = xa)"
-  by (import hollight REAL_LE_LT)
-
-lemma REAL_SUB_LT: "real_lt (real_of_num 0) (real_sub x xa) = real_lt xa x"
-  by (import hollight REAL_SUB_LT)
-
-lemma REAL_NOT_LT: "(~ real_lt x xa) = real_le xa x"
-  by (import hollight REAL_NOT_LT)
-
-lemma REAL_SUB_0: "(real_sub x y = real_of_num 0) = (x = y)"
-  by (import hollight REAL_SUB_0)
-
-lemma REAL_LT_LE: "real_lt x y = (real_le x y & x ~= y)"
-  by (import hollight REAL_LT_LE)
-
-lemma REAL_LT_REFL: "~ real_lt x x"
-  by (import hollight REAL_LT_REFL)
-
-lemma REAL_LTE_ADD: "real_lt (real_of_num 0) x & real_le (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
-  by (import hollight REAL_LTE_ADD)
-
-lemma REAL_LET_ADD: "real_le (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
-  by (import hollight REAL_LET_ADD)
-
-lemma REAL_LT_ADD: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
-  by (import hollight REAL_LT_ADD)
-
-lemma REAL_ENTIRE: "(real_mul x y = real_of_num 0) = (x = real_of_num 0 | y = real_of_num 0)"
-  by (import hollight REAL_ENTIRE)
-
-lemma REAL_LE_NEGTOTAL: "real_le (real_of_num 0) x | real_le (real_of_num 0) (real_neg x)"
-  by (import hollight REAL_LE_NEGTOTAL)
-
-lemma REAL_LE_SQUARE: "real_le (real_of_num 0) (real_mul x x)"
-  by (import hollight REAL_LE_SQUARE)
-
-lemma REAL_MUL_RID: "real_mul x (real_of_num 1) = x"
-  by (import hollight REAL_MUL_RID)
-
-lemma REAL_POW_2: "real_pow x 2 = real_mul x x"
-  by (import hollight REAL_POW_2)
-
-lemma REAL_POLY_CLAUSES: "(ALL x y z. real_add x (real_add y z) = real_add (real_add x y) z) &
-(ALL x y. real_add x y = real_add y x) &
-(ALL x. real_add (real_of_num 0) x = x) &
-(ALL x y z. real_mul x (real_mul y z) = real_mul (real_mul x y) z) &
-(ALL x y. real_mul x y = real_mul y x) &
-(ALL x. real_mul (real_of_num 1) x = x) &
-(ALL x. real_mul (real_of_num 0) x = real_of_num 0) &
-(ALL x xa xb.
-    real_mul x (real_add xa xb) =
-    real_add (real_mul x xa) (real_mul x xb)) &
-(ALL x. real_pow x 0 = real_of_num 1) &
-(ALL x xa. real_pow x (Suc xa) = real_mul x (real_pow x xa))"
-  by (import hollight REAL_POLY_CLAUSES)
-
-lemma REAL_POLY_NEG_CLAUSES: "(ALL x. real_neg x = real_mul (real_neg (real_of_num 1)) x) &
-(ALL x xa.
-    real_sub x xa = real_add x (real_mul (real_neg (real_of_num 1)) xa))"
-  by (import hollight REAL_POLY_NEG_CLAUSES)
-
-lemma REAL_POS: "real_le (real_of_num 0) (real_of_num x)"
-  by (import hollight REAL_POS)
-
-lemma REAL_OF_NUM_LT: "real_lt (real_of_num x) (real_of_num xa) = (x < xa)"
-  by (import hollight REAL_OF_NUM_LT)
-
-lemma REAL_OF_NUM_GE: "real_ge (real_of_num x) (real_of_num xa) = (xa <= x)"
-  by (import hollight REAL_OF_NUM_GE)
-
-lemma REAL_OF_NUM_GT: "real_gt (real_of_num x) (real_of_num xa) = (xa < x)"
-  by (import hollight REAL_OF_NUM_GT)
-
-lemma REAL_OF_NUM_MAX: "real_max (real_of_num x) (real_of_num xa) = real_of_num (max x xa)"
-  by (import hollight REAL_OF_NUM_MAX)
-
-lemma REAL_OF_NUM_MIN: "real_min (real_of_num x) (real_of_num xa) = real_of_num (min x xa)"
-  by (import hollight REAL_OF_NUM_MIN)
-
-lemma REAL_OF_NUM_SUC: "real_add (real_of_num x) (real_of_num 1) = real_of_num (Suc x)"
-  by (import hollight REAL_OF_NUM_SUC)
-
-lemma REAL_OF_NUM_SUB: "m <= n ==> real_sub (real_of_num n) (real_of_num m) = real_of_num (n - m)"
-  by (import hollight REAL_OF_NUM_SUB)
-
-lemma REAL_MUL_AC: "real_mul m n = real_mul n m &
-real_mul (real_mul m n) p = real_mul m (real_mul n p) &
-real_mul m (real_mul n p) = real_mul n (real_mul m p)"
-  by (import hollight REAL_MUL_AC)
-
-lemma REAL_ADD_RDISTRIB: "real_mul (real_add x y) z = real_add (real_mul x z) (real_mul y z)"
-  by (import hollight REAL_ADD_RDISTRIB)
-
-lemma REAL_LT_LADD_IMP: "real_lt y z ==> real_lt (real_add x y) (real_add x z)"
-  by (import hollight REAL_LT_LADD_IMP)
-
-lemma REAL_LT_MUL: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_mul x y)"
-  by (import hollight REAL_LT_MUL)
-
-lemma REAL_EQ_ADD_LCANCEL_0: "(real_add x y = x) = (y = real_of_num 0)"
-  by (import hollight REAL_EQ_ADD_LCANCEL_0)
-
-lemma REAL_EQ_ADD_RCANCEL_0: "(real_add x y = y) = (x = real_of_num 0)"
-  by (import hollight REAL_EQ_ADD_RCANCEL_0)
-
-lemma REAL_LNEG_UNIQ: "(real_add x y = real_of_num 0) = (x = real_neg y)"
-  by (import hollight REAL_LNEG_UNIQ)
-
-lemma REAL_RNEG_UNIQ: "(real_add x y = real_of_num 0) = (y = real_neg x)"
-  by (import hollight REAL_RNEG_UNIQ)
-
-lemma REAL_NEG_LMUL: "real_neg (real_mul x y) = real_mul (real_neg x) y"
-  by (import hollight REAL_NEG_LMUL)
-
-lemma REAL_NEG_RMUL: "real_neg (real_mul x y) = real_mul x (real_neg y)"
-  by (import hollight REAL_NEG_RMUL)
-
-lemma REAL_NEGNEG: "real_neg (real_neg x) = x"
-  by (import hollight REAL_NEGNEG)
-
-lemma REAL_NEG_MUL2: "real_mul (real_neg x) (real_neg y) = real_mul x y"
-  by (import hollight REAL_NEG_MUL2)
-
-lemma REAL_LT_LADD: "real_lt (real_add x y) (real_add x z) = real_lt y z"
-  by (import hollight REAL_LT_LADD)
-
-lemma REAL_LT_RADD: "real_lt (real_add x z) (real_add y z) = real_lt x y"
-  by (import hollight REAL_LT_RADD)
-
-lemma REAL_LT_ANTISYM: "~ (real_lt x y & real_lt y x)"
-  by (import hollight REAL_LT_ANTISYM)
-
-lemma REAL_LT_GT: "real_lt x y ==> ~ real_lt y x"
-  by (import hollight REAL_LT_GT)
-
-lemma REAL_NOT_EQ: "(x ~= y) = (real_lt x y | real_lt y x)"
-  by (import hollight REAL_NOT_EQ)
-
-lemma REAL_LET_ANTISYM: "~ (real_le x y & real_lt y x)"
-  by (import hollight REAL_LET_ANTISYM)
-
-lemma REAL_NEG_LT0: "real_lt (real_neg x) (real_of_num 0) = real_lt (real_of_num 0) x"
-  by (import hollight REAL_NEG_LT0)
-
-lemma REAL_NEG_GT0: "real_lt (real_of_num 0) (real_neg x) = real_lt x (real_of_num 0)"
-  by (import hollight REAL_NEG_GT0)
-
-lemma REAL_NEG_LE0: "real_le (real_neg x) (real_of_num 0) = real_le (real_of_num 0) x"
-  by (import hollight REAL_NEG_LE0)
-
-lemma REAL_NEG_GE0: "real_le (real_of_num 0) (real_neg x) = real_le x (real_of_num 0)"
-  by (import hollight REAL_NEG_GE0)
-
-lemma REAL_LT_TOTAL: "x = y | real_lt x y | real_lt y x"
-  by (import hollight REAL_LT_TOTAL)
-
-lemma REAL_LT_NEGTOTAL: "x = real_of_num 0 |
-real_lt (real_of_num 0) x | real_lt (real_of_num 0) (real_neg x)"
-  by (import hollight REAL_LT_NEGTOTAL)
-
-lemma REAL_LE_01: "real_le (real_of_num 0) (real_of_num 1)"
-  by (import hollight REAL_LE_01)
-
-lemma REAL_LT_01: "real_lt (real_of_num 0) (real_of_num 1)"
-  by (import hollight REAL_LT_01)
-
-lemma REAL_LE_LADD: "real_le (real_add x y) (real_add x z) = real_le y z"
-  by (import hollight REAL_LE_LADD)
-
-lemma REAL_LE_RADD: "real_le (real_add x z) (real_add y z) = real_le x y"
-  by (import hollight REAL_LE_RADD)
-
-lemma REAL_LT_ADD2: "real_lt w x & real_lt y z ==> real_lt (real_add w y) (real_add x z)"
-  by (import hollight REAL_LT_ADD2)
-
-lemma REAL_LE_ADD2: "real_le w x & real_le y z ==> real_le (real_add w y) (real_add x z)"
-  by (import hollight REAL_LE_ADD2)
-
-lemma REAL_LT_LNEG: "real_lt (real_neg x) xa = real_lt (real_of_num 0) (real_add x xa)"
-  by (import hollight REAL_LT_LNEG)
-
-lemma REAL_LT_RNEG: "real_lt x (real_neg xa) = real_lt (real_add x xa) (real_of_num 0)"
-  by (import hollight REAL_LT_RNEG)
-
-lemma REAL_LT_ADDNEG: "real_lt y (real_add x (real_neg z)) = real_lt (real_add y z) x"
-  by (import hollight REAL_LT_ADDNEG)
-
-lemma REAL_LT_ADDNEG2: "real_lt (real_add x (real_neg y)) z = real_lt x (real_add z y)"
-  by (import hollight REAL_LT_ADDNEG2)
-
-lemma REAL_LT_ADD1: "real_le x y ==> real_lt x (real_add y (real_of_num 1))"
-  by (import hollight REAL_LT_ADD1)
-
-lemma REAL_SUB_ADD: "real_add (real_sub x y) y = x"
-  by (import hollight REAL_SUB_ADD)
-
-lemma REAL_SUB_ADD2: "real_add y (real_sub x y) = x"
-  by (import hollight REAL_SUB_ADD2)
-
-lemma REAL_SUB_REFL: "real_sub x x = real_of_num 0"
-  by (import hollight REAL_SUB_REFL)
-
-lemma REAL_LE_DOUBLE: "real_le (real_of_num 0) (real_add x x) = real_le (real_of_num 0) x"
-  by (import hollight REAL_LE_DOUBLE)
-
-lemma REAL_LE_NEGL: "real_le (real_neg x) x = real_le (real_of_num 0) x"
-  by (import hollight REAL_LE_NEGL)
-
-lemma REAL_LE_NEGR: "real_le x (real_neg x) = real_le x (real_of_num 0)"
-  by (import hollight REAL_LE_NEGR)
-
-lemma REAL_NEG_EQ_0: "(real_neg x = real_of_num 0) = (x = real_of_num 0)"
-  by (import hollight REAL_NEG_EQ_0)
-
-lemma REAL_ADD_SUB: "real_sub (real_add x y) x = y"
-  by (import hollight REAL_ADD_SUB)
-
-lemma REAL_NEG_EQ: "(real_neg x = y) = (x = real_neg y)"
-  by (import hollight REAL_NEG_EQ)
-
-lemma REAL_NEG_MINUS1: "real_neg x = real_mul (real_neg (real_of_num 1)) x"
-  by (import hollight REAL_NEG_MINUS1)
-
-lemma REAL_LT_IMP_NE: "real_lt x y ==> x ~= y"
-  by (import hollight REAL_LT_IMP_NE)
-
-lemma REAL_LE_ADDR: "real_le x (real_add x y) = real_le (real_of_num 0) y"
-  by (import hollight REAL_LE_ADDR)
-
-lemma REAL_LE_ADDL: "real_le y (real_add x y) = real_le (real_of_num 0) x"
-  by (import hollight REAL_LE_ADDL)
-
-lemma REAL_LT_ADDR: "real_lt x (real_add x y) = real_lt (real_of_num 0) y"
-  by (import hollight REAL_LT_ADDR)
-
-lemma REAL_LT_ADDL: "real_lt y (real_add x y) = real_lt (real_of_num 0) x"
-  by (import hollight REAL_LT_ADDL)
-
-lemma REAL_SUB_SUB: "real_sub (real_sub x y) x = real_neg y"
-  by (import hollight REAL_SUB_SUB)
-
-lemma REAL_LT_ADD_SUB: "real_lt (real_add x y) z = real_lt x (real_sub z y)"
-  by (import hollight REAL_LT_ADD_SUB)
-
-lemma REAL_LT_SUB_RADD: "real_lt (real_sub x y) z = real_lt x (real_add z y)"
-  by (import hollight REAL_LT_SUB_RADD)
-
-lemma REAL_LT_SUB_LADD: "real_lt x (real_sub y z) = real_lt (real_add x z) y"
-  by (import hollight REAL_LT_SUB_LADD)
-
-lemma REAL_LE_SUB_LADD: "real_le x (real_sub y z) = real_le (real_add x z) y"
-  by (import hollight REAL_LE_SUB_LADD)
-
-lemma REAL_LE_SUB_RADD: "real_le (real_sub x y) z = real_le x (real_add z y)"
-  by (import hollight REAL_LE_SUB_RADD)
-
-lemma REAL_LT_NEG: "real_lt (real_neg x) (real_neg y) = real_lt y x"
-  by (import hollight REAL_LT_NEG)
-
-lemma REAL_LE_NEG: "real_le (real_neg x) (real_neg y) = real_le y x"
-  by (import hollight REAL_LE_NEG)
-
-lemma REAL_ADD2_SUB2: "real_sub (real_add a b) (real_add c d) =
-real_add (real_sub a c) (real_sub b d)"
-  by (import hollight REAL_ADD2_SUB2)
-
-lemma REAL_SUB_LZERO: "real_sub (real_of_num 0) x = real_neg x"
-  by (import hollight REAL_SUB_LZERO)
-
-lemma REAL_SUB_RZERO: "real_sub x (real_of_num 0) = x"
-  by (import hollight REAL_SUB_RZERO)
-
-lemma REAL_LET_ADD2: "real_le w x & real_lt y z ==> real_lt (real_add w y) (real_add x z)"
-  by (import hollight REAL_LET_ADD2)
-
-lemma REAL_LTE_ADD2: "real_lt w x & real_le y z ==> real_lt (real_add w y) (real_add x z)"
-  by (import hollight REAL_LTE_ADD2)
-
-lemma REAL_SUB_LNEG: "real_sub (real_neg x) y = real_neg (real_add x y)"
-  by (import hollight REAL_SUB_LNEG)
-
-lemma REAL_SUB_RNEG: "real_sub x (real_neg y) = real_add x y"
-  by (import hollight REAL_SUB_RNEG)
-
-lemma REAL_SUB_NEG2: "real_sub (real_neg x) (real_neg y) = real_sub y x"
-  by (import hollight REAL_SUB_NEG2)
-
-lemma REAL_SUB_TRIANGLE: "real_add (real_sub a b) (real_sub b c) = real_sub a c"
-  by (import hollight REAL_SUB_TRIANGLE)
-
-lemma REAL_EQ_SUB_LADD: "(x = real_sub y z) = (real_add x z = y)"
-  by (import hollight REAL_EQ_SUB_LADD)
-
-lemma REAL_EQ_SUB_RADD: "(real_sub x y = z) = (x = real_add z y)"
-  by (import hollight REAL_EQ_SUB_RADD)
-
-lemma REAL_SUB_SUB2: "real_sub x (real_sub x y) = y"
-  by (import hollight REAL_SUB_SUB2)
-
-lemma REAL_ADD_SUB2: "real_sub x (real_add x y) = real_neg y"
-  by (import hollight REAL_ADD_SUB2)
-
-lemma REAL_EQ_IMP_LE: "x = y ==> real_le x y"
-  by (import hollight REAL_EQ_IMP_LE)
-
-lemma REAL_POS_NZ: "real_lt (real_of_num 0) x ==> x ~= real_of_num 0"
-  by (import hollight REAL_POS_NZ)
-
-lemma REAL_DIFFSQ: "real_mul (real_add x y) (real_sub x y) =
-real_sub (real_mul x x) (real_mul y y)"
-  by (import hollight REAL_DIFFSQ)
-
-lemma REAL_EQ_NEG2: "(real_neg x = real_neg y) = (x = y)"
-  by (import hollight REAL_EQ_NEG2)
-
-lemma REAL_LT_NEG2: "real_lt (real_neg x) (real_neg y) = real_lt y x"
-  by (import hollight REAL_LT_NEG2)
-
-lemma REAL_SUB_LDISTRIB: "real_mul x (real_sub y z) = real_sub (real_mul x y) (real_mul x z)"
-  by (import hollight REAL_SUB_LDISTRIB)
-
-lemma REAL_SUB_RDISTRIB: "real_mul (real_sub x y) z = real_sub (real_mul x z) (real_mul y z)"
-  by (import hollight REAL_SUB_RDISTRIB)
-
-lemma REAL_ABS_ZERO: "(real_abs x = real_of_num 0) = (x = real_of_num 0)"
-  by (import hollight REAL_ABS_ZERO)
-
-lemma REAL_ABS_0: "real_abs (real_of_num 0) = real_of_num 0"
-  by (import hollight REAL_ABS_0)
-
-lemma REAL_ABS_1: "real_abs (real_of_num 1) = real_of_num 1"
-  by (import hollight REAL_ABS_1)
-
-lemma REAL_ABS_TRIANGLE: "real_le (real_abs (real_add x y)) (real_add (real_abs x) (real_abs y))"
-  by (import hollight REAL_ABS_TRIANGLE)
-
-lemma REAL_ABS_TRIANGLE_LE: "real_le (real_add (real_abs x) (real_abs (real_sub y x))) z
-==> real_le (real_abs y) z"
-  by (import hollight REAL_ABS_TRIANGLE_LE)
-
-lemma REAL_ABS_TRIANGLE_LT: "real_lt (real_add (real_abs x) (real_abs (real_sub y x))) z
-==> real_lt (real_abs y) z"
-  by (import hollight REAL_ABS_TRIANGLE_LT)
-
-lemma REAL_ABS_POS: "real_le (real_of_num 0) (real_abs x)"
-  by (import hollight REAL_ABS_POS)
-
-lemma REAL_ABS_SUB: "real_abs (real_sub x y) = real_abs (real_sub y x)"
-  by (import hollight REAL_ABS_SUB)
-
-lemma REAL_ABS_NZ: "(x ~= real_of_num 0) = real_lt (real_of_num 0) (real_abs x)"
-  by (import hollight REAL_ABS_NZ)
-
-lemma REAL_ABS_ABS: "real_abs (real_abs x) = real_abs x"
-  by (import hollight REAL_ABS_ABS)
-
-lemma REAL_ABS_LE: "real_le x (real_abs x)"
-  by (import hollight REAL_ABS_LE)
-
-lemma REAL_ABS_REFL: "(real_abs x = x) = real_le (real_of_num 0) x"
-  by (import hollight REAL_ABS_REFL)
-
-lemma REAL_ABS_BETWEEN: "(real_lt (real_of_num 0) d &
- real_lt (real_sub x d) y & real_lt y (real_add x d)) =
-real_lt (real_abs (real_sub y x)) d"
-  by (import hollight REAL_ABS_BETWEEN)
-
-lemma REAL_ABS_BOUND: "real_lt (real_abs (real_sub x y)) d ==> real_lt y (real_add x d)"
-  by (import hollight REAL_ABS_BOUND)
-
-lemma REAL_ABS_STILLNZ: "real_lt (real_abs (real_sub x y)) (real_abs y) ==> x ~= real_of_num 0"
-  by (import hollight REAL_ABS_STILLNZ)
-
-lemma REAL_ABS_CASES: "x = real_of_num 0 | real_lt (real_of_num 0) (real_abs x)"
-  by (import hollight REAL_ABS_CASES)
-
-lemma REAL_ABS_BETWEEN1: "real_lt x z & real_lt (real_abs (real_sub y x)) (real_sub z x)
-==> real_lt y z"
-  by (import hollight REAL_ABS_BETWEEN1)
-
-lemma REAL_ABS_SIGN: "real_lt (real_abs (real_sub x y)) y ==> real_lt (real_of_num 0) x"
-  by (import hollight REAL_ABS_SIGN)
-
-lemma REAL_ABS_SIGN2: "real_lt (real_abs (real_sub x y)) (real_neg y) ==> real_lt x (real_of_num 0)"
-  by (import hollight REAL_ABS_SIGN2)
-
-lemma REAL_ABS_CIRCLE: "real_lt (real_abs h) (real_sub (real_abs y) (real_abs x))
-==> real_lt (real_abs (real_add x h)) (real_abs y)"
-  by (import hollight REAL_ABS_CIRCLE)
-
-lemma REAL_SUB_ABS: "real_le (real_sub (real_abs x) (real_abs y)) (real_abs (real_sub x y))"
-  by (import hollight REAL_SUB_ABS)
-
-lemma REAL_ABS_SUB_ABS: "real_le (real_abs (real_sub (real_abs x) (real_abs y)))
- (real_abs (real_sub x y))"
-  by (import hollight REAL_ABS_SUB_ABS)
-
-lemma REAL_ABS_BETWEEN2: "real_lt x0 y0 &
-real_lt (real_mul (real_of_num 2) (real_abs (real_sub x x0)))
- (real_sub y0 x0) &
-real_lt (real_mul (real_of_num 2) (real_abs (real_sub y y0)))
- (real_sub y0 x0)
-==> real_lt x y"
-  by (import hollight REAL_ABS_BETWEEN2)
-
-lemma REAL_ABS_BOUNDS: "real_le (real_abs x) k = (real_le (real_neg k) x & real_le x k)"
-  by (import hollight REAL_ABS_BOUNDS)
-
-lemma REAL_BOUNDS_LE: "(real_le (real_neg k) x & real_le x k) = real_le (real_abs x) k"
-  by (import hollight REAL_BOUNDS_LE)
-
-lemma REAL_BOUNDS_LT: "(real_lt (real_neg k) x & real_lt x k) = real_lt (real_abs x) k"
-  by (import hollight REAL_BOUNDS_LT)
-
-lemma REAL_MIN_MAX: "real_min x y = real_neg (real_max (real_neg x) (real_neg y))"
-  by (import hollight REAL_MIN_MAX)
-
-lemma REAL_MAX_MIN: "real_max x y = real_neg (real_min (real_neg x) (real_neg y))"
-  by (import hollight REAL_MAX_MIN)
-
-lemma REAL_MAX_MAX: "real_le x (real_max x y) & real_le y (real_max x y)"
-  by (import hollight REAL_MAX_MAX)
-
-lemma REAL_MIN_MIN: "real_le (real_min x y) x & real_le (real_min x y) y"
-  by (import hollight REAL_MIN_MIN)
-
-lemma REAL_MAX_SYM: "real_max x y = real_max y x"
-  by (import hollight REAL_MAX_SYM)
-
-lemma REAL_MIN_SYM: "real_min x y = real_min y x"
-  by (import hollight REAL_MIN_SYM)
-
-lemma REAL_LE_MAX: "real_le z (real_max x y) = (real_le z x | real_le z y)"
-  by (import hollight REAL_LE_MAX)
-
-lemma REAL_LE_MIN: "real_le z (real_min x y) = (real_le z x & real_le z y)"
-  by (import hollight REAL_LE_MIN)
-
-lemma REAL_LT_MAX: "real_lt z (real_max x y) = (real_lt z x | real_lt z y)"
-  by (import hollight REAL_LT_MAX)
-
-lemma REAL_LT_MIN: "real_lt z (real_min x y) = (real_lt z x & real_lt z y)"
-  by (import hollight REAL_LT_MIN)
-
-lemma REAL_MAX_LE: "real_le (real_max x y) z = (real_le x z & real_le y z)"
-  by (import hollight REAL_MAX_LE)
-
-lemma REAL_MIN_LE: "real_le (real_min x y) z = (real_le x z | real_le y z)"
-  by (import hollight REAL_MIN_LE)
-
-lemma REAL_MAX_LT: "real_lt (real_max x y) z = (real_lt x z & real_lt y z)"
-  by (import hollight REAL_MAX_LT)
-
-lemma REAL_MIN_LT: "real_lt (real_min x y) z = (real_lt x z | real_lt y z)"
-  by (import hollight REAL_MIN_LT)
-
-lemma REAL_MAX_ASSOC: "real_max x (real_max y z) = real_max (real_max x y) z"
-  by (import hollight REAL_MAX_ASSOC)
-
-lemma REAL_MIN_ASSOC: "real_min x (real_min y z) = real_min (real_min x y) z"
-  by (import hollight REAL_MIN_ASSOC)
-
-lemma REAL_MAX_ACI: "real_max x y = real_max y x &
-real_max (real_max x y) z = real_max x (real_max y z) &
-real_max x (real_max y z) = real_max y (real_max x z) &
-real_max x x = x & real_max x (real_max x y) = real_max x y"
-  by (import hollight REAL_MAX_ACI)
-
-lemma REAL_MIN_ACI: "real_min x y = real_min y x &
-real_min (real_min x y) z = real_min x (real_min y z) &
-real_min x (real_min y z) = real_min y (real_min x z) &
-real_min x x = x & real_min x (real_min x y) = real_min x y"
-  by (import hollight REAL_MIN_ACI)
-
-lemma REAL_ABS_MUL: "real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)"
-  by (import hollight REAL_ABS_MUL)
-
-lemma REAL_POW_LE: "real_le (real_of_num 0) x ==> real_le (real_of_num 0) (real_pow x n)"
-  by (import hollight REAL_POW_LE)
-
-lemma REAL_POW_LT: "real_lt (real_of_num 0) x ==> real_lt (real_of_num 0) (real_pow x n)"
-  by (import hollight REAL_POW_LT)
-
-lemma REAL_ABS_POW: "real_abs (real_pow x n) = real_pow (real_abs x) n"
-  by (import hollight REAL_ABS_POW)
-
-lemma REAL_LE_LMUL: "real_le (real_of_num 0) x & real_le xa xb
-==> real_le (real_mul x xa) (real_mul x xb)"
-  by (import hollight REAL_LE_LMUL)
-
-lemma REAL_LE_RMUL: "real_le x y & real_le (real_of_num 0) z
-==> real_le (real_mul x z) (real_mul y z)"
-  by (import hollight REAL_LE_RMUL)
-
-lemma REAL_LT_LMUL: "real_lt (real_of_num 0) x & real_lt xa xb
-==> real_lt (real_mul x xa) (real_mul x xb)"
-  by (import hollight REAL_LT_LMUL)
-
-lemma REAL_LT_RMUL: "real_lt x y & real_lt (real_of_num 0) z
-==> real_lt (real_mul x z) (real_mul y z)"
-  by (import hollight REAL_LT_RMUL)
-
-lemma REAL_EQ_MUL_LCANCEL: "(real_mul x y = real_mul x z) = (x = real_of_num 0 | y = z)"
-  by (import hollight REAL_EQ_MUL_LCANCEL)
-
-lemma REAL_EQ_MUL_RCANCEL: "(real_mul x xb = real_mul xa xb) = (x = xa | xb = real_of_num 0)"
-  by (import hollight REAL_EQ_MUL_RCANCEL)
-
-lemma REAL_MUL_LINV_UNIQ: "real_mul x y = real_of_num 1 ==> real_inv y = x"
-  by (import hollight REAL_MUL_LINV_UNIQ)
-
-lemma REAL_MUL_RINV_UNIQ: "real_mul x xa = real_of_num 1 ==> real_inv x = xa"
-  by (import hollight REAL_MUL_RINV_UNIQ)
-
-lemma REAL_INV_INV: "real_inv (real_inv x) = x"
-  by (import hollight REAL_INV_INV)
-
-lemma REAL_EQ_INV2: "(real_inv x = real_inv y) = (x = y)"
-  by (import hollight REAL_EQ_INV2)
-
-lemma REAL_INV_EQ_0: "(real_inv x = real_of_num 0) = (x = real_of_num 0)"
-  by (import hollight REAL_INV_EQ_0)
-
-lemma REAL_LT_INV: "real_lt (real_of_num 0) x ==> real_lt (real_of_num 0) (real_inv x)"
-  by (import hollight REAL_LT_INV)
-
-lemma REAL_LT_INV_EQ: "real_lt (real_of_num 0) (real_inv x) = real_lt (real_of_num 0) x"
-  by (import hollight REAL_LT_INV_EQ)
-
-lemma REAL_INV_NEG: "real_inv (real_neg x) = real_neg (real_inv x)"
-  by (import hollight REAL_INV_NEG)
-
-lemma REAL_LE_INV_EQ: "real_le (real_of_num 0) (real_inv x) = real_le (real_of_num 0) x"
-  by (import hollight REAL_LE_INV_EQ)
-
-lemma REAL_LE_INV: "real_le (real_of_num 0) x ==> real_le (real_of_num 0) (real_inv x)"
-  by (import hollight REAL_LE_INV)
-
-lemma REAL_MUL_RINV: "x ~= real_of_num 0 ==> real_mul x (real_inv x) = real_of_num 1"
-  by (import hollight REAL_MUL_RINV)
-
-lemma REAL_INV_1: "real_inv (real_of_num 1) = real_of_num 1"
-  by (import hollight REAL_INV_1)
-
-lemma REAL_INV_EQ_1: "(real_inv x = real_of_num 1) = (x = real_of_num 1)"
-  by (import hollight REAL_INV_EQ_1)
-
-lemma REAL_DIV_1: "real_div x (real_of_num 1) = x"
-  by (import hollight REAL_DIV_1)
-
-lemma REAL_DIV_REFL: "x ~= real_of_num 0 ==> real_div x x = real_of_num 1"
-  by (import hollight REAL_DIV_REFL)
-
-lemma REAL_DIV_RMUL: "xa ~= real_of_num 0 ==> real_mul (real_div x xa) xa = x"
-  by (import hollight REAL_DIV_RMUL)
-
-lemma REAL_DIV_LMUL: "xa ~= real_of_num 0 ==> real_mul xa (real_div x xa) = x"
-  by (import hollight REAL_DIV_LMUL)
-
-lemma REAL_ABS_INV: "real_abs (real_inv x) = real_inv (real_abs x)"
-  by (import hollight REAL_ABS_INV)
-
-lemma REAL_ABS_DIV: "real_abs (real_div x xa) = real_div (real_abs x) (real_abs xa)"
-  by (import hollight REAL_ABS_DIV)
-
-lemma REAL_INV_MUL: "real_inv (real_mul x y) = real_mul (real_inv x) (real_inv y)"
-  by (import hollight REAL_INV_MUL)
-
-lemma REAL_INV_DIV: "real_inv (real_div x xa) = real_div xa x"
-  by (import hollight REAL_INV_DIV)
-
-lemma REAL_POW_MUL: "real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)"
-  by (import hollight REAL_POW_MUL)
-
-lemma REAL_POW_INV: "real_pow (real_inv x) n = real_inv (real_pow x n)"
-  by (import hollight REAL_POW_INV)
-
-lemma REAL_INV_POW: "real_inv (real_pow x xa) = real_pow (real_inv x) xa"
-  by (import hollight REAL_INV_POW)
-
-lemma REAL_POW_DIV: "real_pow (real_div x xa) xb = real_div (real_pow x xb) (real_pow xa xb)"
-  by (import hollight REAL_POW_DIV)
-
-lemma REAL_POW_ADD: "real_pow x (m + n) = real_mul (real_pow x m) (real_pow x n)"
-  by (import hollight REAL_POW_ADD)
-
-lemma REAL_POW_NZ: "x ~= real_of_num 0 ==> real_pow x n ~= real_of_num 0"
-  by (import hollight REAL_POW_NZ)
-
-lemma REAL_POW_SUB: "x ~= real_of_num 0 & m <= n
-==> real_pow x (n - m) = real_div (real_pow x n) (real_pow x m)"
-  by (import hollight REAL_POW_SUB)
-
-lemma REAL_LT_IMP_NZ: "real_lt (real_of_num 0) x ==> x ~= real_of_num 0"
-  by (import hollight REAL_LT_IMP_NZ)
-
-lemma REAL_LT_LCANCEL_IMP: "real_lt (real_of_num 0) x & real_lt (real_mul x y) (real_mul x z)
-==> real_lt y z"
-  by (import hollight REAL_LT_LCANCEL_IMP)
-
-lemma REAL_LT_RCANCEL_IMP: "real_lt (real_of_num 0) xb & real_lt (real_mul x xb) (real_mul xa xb)
-==> real_lt x xa"
-  by (import hollight REAL_LT_RCANCEL_IMP)
-
-lemma REAL_LE_LCANCEL_IMP: "real_lt (real_of_num 0) x & real_le (real_mul x y) (real_mul x z)
-==> real_le y z"
-  by (import hollight REAL_LE_LCANCEL_IMP)
-
-lemma REAL_LE_RCANCEL_IMP: "real_lt (real_of_num 0) xb & real_le (real_mul x xb) (real_mul xa xb)
-==> real_le x xa"
-  by (import hollight REAL_LE_RCANCEL_IMP)
-
-lemma REAL_LE_RMUL_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_mul x z) (real_mul y z) = real_le x y"
-  by (import hollight REAL_LE_RMUL_EQ)
-
-lemma REAL_LE_LMUL_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_mul z x) (real_mul z y) = real_le x y"
-  by (import hollight REAL_LE_LMUL_EQ)
-
-lemma REAL_LT_RMUL_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_mul x xb) (real_mul xa xb) = real_lt x xa"
-  by (import hollight REAL_LT_RMUL_EQ)
-
-lemma REAL_LT_LMUL_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_mul xb x) (real_mul xb xa) = real_lt x xa"
-  by (import hollight REAL_LT_LMUL_EQ)
-
-lemma REAL_LE_MUL_EQ: "(ALL x y.
-    real_lt (real_of_num 0) x -->
-    real_le (real_of_num 0) (real_mul x y) = real_le (real_of_num 0) y) &
-(ALL x y.
-    real_lt (real_of_num 0) y -->
-    real_le (real_of_num 0) (real_mul x y) = real_le (real_of_num 0) x)"
-  by (import hollight REAL_LE_MUL_EQ)
-
-lemma REAL_LT_MUL_EQ: "(ALL x y.
-    real_lt (real_of_num 0) x -->
-    real_lt (real_of_num 0) (real_mul x y) = real_lt (real_of_num 0) y) &
-(ALL x y.
-    real_lt (real_of_num 0) y -->
-    real_lt (real_of_num 0) (real_mul x y) = real_lt (real_of_num 0) x)"
-  by (import hollight REAL_LT_MUL_EQ)
-
-lemma REAL_MUL_POS_LT: "real_lt (real_of_num 0) (real_mul x y) =
-(real_lt (real_of_num 0) x & real_lt (real_of_num 0) y |
- real_lt x (real_of_num 0) & real_lt y (real_of_num 0))"
-  by (import hollight REAL_MUL_POS_LT)
-
-lemma REAL_MUL_POS_LE: "real_le (real_of_num 0) (real_mul x xa) =
-(x = real_of_num 0 |
- xa = real_of_num 0 |
- real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa |
- real_lt x (real_of_num 0) & real_lt xa (real_of_num 0))"
-  by (import hollight REAL_MUL_POS_LE)
-
-lemma REAL_LE_RDIV_EQ: "real_lt (real_of_num 0) z
-==> real_le x (real_div y z) = real_le (real_mul x z) y"
-  by (import hollight REAL_LE_RDIV_EQ)
-
-lemma REAL_LE_LDIV_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_div x z) y = real_le x (real_mul y z)"
-  by (import hollight REAL_LE_LDIV_EQ)
-
-lemma REAL_LT_RDIV_EQ: "real_lt (real_of_num 0) xb
-==> real_lt x (real_div xa xb) = real_lt (real_mul x xb) xa"
-  by (import hollight REAL_LT_RDIV_EQ)
-
-lemma REAL_LT_LDIV_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_div x xb) xa = real_lt x (real_mul xa xb)"
-  by (import hollight REAL_LT_LDIV_EQ)
-
-lemma REAL_EQ_RDIV_EQ: "real_lt (real_of_num 0) xb ==> (x = real_div xa xb) = (real_mul x xb = xa)"
-  by (import hollight REAL_EQ_RDIV_EQ)
-
-lemma REAL_EQ_LDIV_EQ: "real_lt (real_of_num 0) xb ==> (real_div x xb = xa) = (x = real_mul xa xb)"
-  by (import hollight REAL_EQ_LDIV_EQ)
-
-lemma REAL_LT_DIV2_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_div x xb) (real_div xa xb) = real_lt x xa"
-  by (import hollight REAL_LT_DIV2_EQ)
-
-lemma REAL_LE_DIV2_EQ: "real_lt (real_of_num 0) xb
-==> real_le (real_div x xb) (real_div xa xb) = real_le x xa"
-  by (import hollight REAL_LE_DIV2_EQ)
-
-lemma REAL_MUL_2: "real_mul (real_of_num 2) x = real_add x x"
-  by (import hollight REAL_MUL_2)
-
-lemma REAL_POW_EQ_0: "(real_pow x n = real_of_num 0) = (x = real_of_num 0 & n ~= 0)"
-  by (import hollight REAL_POW_EQ_0)
-
-lemma REAL_LE_MUL2: "real_le (real_of_num 0) w &
-real_le w x & real_le (real_of_num 0) y & real_le y z
-==> real_le (real_mul w y) (real_mul x z)"
-  by (import hollight REAL_LE_MUL2)
-
-lemma REAL_LT_MUL2: "real_le (real_of_num 0) w &
-real_lt w x & real_le (real_of_num 0) y & real_lt y z
-==> real_lt (real_mul w y) (real_mul x z)"
-  by (import hollight REAL_LT_MUL2)
-
-lemma REAL_LT_SQUARE: "real_lt (real_of_num 0) (real_mul x x) = (x ~= real_of_num 0)"
-  by (import hollight REAL_LT_SQUARE)
-
-lemma REAL_POW_1: "real_pow x 1 = x"
-  by (import hollight REAL_POW_1)
-
-lemma REAL_POW_ONE: "real_pow (real_of_num 1) n = real_of_num 1"
-  by (import hollight REAL_POW_ONE)
-
-lemma REAL_LT_INV2: "real_lt (real_of_num 0) x & real_lt x y
-==> real_lt (real_inv y) (real_inv x)"
-  by (import hollight REAL_LT_INV2)
-
-lemma REAL_LE_INV2: "real_lt (real_of_num 0) x & real_le x y
-==> real_le (real_inv y) (real_inv x)"
-  by (import hollight REAL_LE_INV2)
-
-lemma REAL_LT_LINV: "real_lt (real_of_num 0) y & real_lt (real_inv y) x
-==> real_lt (real_inv x) y"
-  by (import hollight REAL_LT_LINV)
-
-lemma REAL_LT_RINV: "real_lt (real_of_num 0) x & real_lt x (real_inv y)
-==> real_lt y (real_inv x)"
-  by (import hollight REAL_LT_RINV)
-
-lemma REAL_LE_LINV: "real_lt (real_of_num 0) y & real_le (real_inv y) x
-==> real_le (real_inv x) y"
-  by (import hollight REAL_LE_LINV)
-
-lemma REAL_LE_RINV: "real_lt (real_of_num 0) x & real_le x (real_inv y)
-==> real_le y (real_inv x)"
-  by (import hollight REAL_LE_RINV)
-
-lemma REAL_INV_LE_1: "real_le (real_of_num 1) x ==> real_le (real_inv x) (real_of_num 1)"
-  by (import hollight REAL_INV_LE_1)
-
-lemma REAL_INV_1_LE: "real_lt (real_of_num 0) x & real_le x (real_of_num 1)
-==> real_le (real_of_num 1) (real_inv x)"
-  by (import hollight REAL_INV_1_LE)
-
-lemma REAL_INV_LT_1: "real_lt (real_of_num 1) x ==> real_lt (real_inv x) (real_of_num 1)"
-  by (import hollight REAL_INV_LT_1)
-
-lemma REAL_INV_1_LT: "real_lt (real_of_num 0) x & real_lt x (real_of_num 1)
-==> real_lt (real_of_num 1) (real_inv x)"
-  by (import hollight REAL_INV_1_LT)
-
-lemma REAL_SUB_INV: "x ~= real_of_num 0 & xa ~= real_of_num 0
-==> real_sub (real_inv x) (real_inv xa) =
-    real_div (real_sub xa x) (real_mul x xa)"
-  by (import hollight REAL_SUB_INV)
-
-lemma REAL_DOWN: "real_lt (real_of_num 0) d ==> EX x. real_lt (real_of_num 0) x & real_lt x d"
-  by (import hollight REAL_DOWN)
-
-lemma REAL_DOWN2: "real_lt (real_of_num 0) d1 & real_lt (real_of_num 0) d2
-==> EX e. real_lt (real_of_num 0) e & real_lt e d1 & real_lt e d2"
-  by (import hollight REAL_DOWN2)
-
-lemma REAL_POW_LE2: "real_le (real_of_num 0) x & real_le x y
-==> real_le (real_pow x n) (real_pow y n)"
-  by (import hollight REAL_POW_LE2)
-
-lemma REAL_POW_LE_1: "real_le (real_of_num 1) x ==> real_le (real_of_num 1) (real_pow x n)"
-  by (import hollight REAL_POW_LE_1)
-
-lemma REAL_POW_1_LE: "real_le (real_of_num 0) x & real_le x (real_of_num 1)
-==> real_le (real_pow x n) (real_of_num 1)"
-  by (import hollight REAL_POW_1_LE)
-
-lemma REAL_POW_MONO: "real_le (real_of_num 1) x & m <= n ==> real_le (real_pow x m) (real_pow x n)"
-  by (import hollight REAL_POW_MONO)
-
-lemma REAL_POW_LT2: "n ~= 0 & real_le (real_of_num 0) x & real_lt x y
-==> real_lt (real_pow x n) (real_pow y n)"
-  by (import hollight REAL_POW_LT2)
-
-lemma REAL_POW_LT_1: "n ~= 0 & real_lt (real_of_num 1) x
-==> real_lt (real_of_num 1) (real_pow x n)"
-  by (import hollight REAL_POW_LT_1)
-
-lemma REAL_POW_1_LT: "n ~= 0 & real_le (real_of_num 0) x & real_lt x (real_of_num 1)
-==> real_lt (real_pow x n) (real_of_num 1)"
-  by (import hollight REAL_POW_1_LT)
-
-lemma REAL_POW_MONO_LT: "real_lt (real_of_num 1) x & m < n ==> real_lt (real_pow x m) (real_pow x n)"
-  by (import hollight REAL_POW_MONO_LT)
-
-lemma REAL_POW_POW: "real_pow (real_pow x m) n = real_pow x (m * n)"
-  by (import hollight REAL_POW_POW)
-
-lemma REAL_EQ_RCANCEL_IMP: "z ~= real_of_num 0 & real_mul x z = real_mul y z ==> x = y"
-  by (import hollight REAL_EQ_RCANCEL_IMP)
-
-lemma REAL_EQ_LCANCEL_IMP: "xb ~= real_of_num 0 & real_mul xb x = real_mul xb xa ==> x = xa"
-  by (import hollight REAL_EQ_LCANCEL_IMP)
-
-lemma REAL_LT_DIV: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa
-==> real_lt (real_of_num 0) (real_div x xa)"
-  by (import hollight REAL_LT_DIV)
-
-lemma REAL_LE_DIV: "real_le (real_of_num 0) x & real_le (real_of_num 0) xa
-==> real_le (real_of_num 0) (real_div x xa)"
-  by (import hollight REAL_LE_DIV)
-
-lemma REAL_DIV_POW2: "x ~= real_of_num 0
-==> real_div (real_pow x m) (real_pow x n) =
-    (if n <= m then real_pow x (m - n) else real_inv (real_pow x (n - m)))"
-  by (import hollight REAL_DIV_POW2)
-
-lemma REAL_DIV_POW2_ALT: "x ~= real_of_num 0
-==> real_div (real_pow x m) (real_pow x n) =
-    (if n < m then real_pow x (m - n) else real_inv (real_pow x (n - m)))"
-  by (import hollight REAL_DIV_POW2_ALT)
-
-lemma REAL_LT_POW2: "real_lt (real_of_num 0) (real_pow (real_of_num 2) x)"
-  by (import hollight REAL_LT_POW2)
-
-lemma REAL_LE_POW2: "real_le (real_of_num 1) (real_pow (real_of_num 2) n)"
-  by (import hollight REAL_LE_POW2)
-
-lemma REAL_POW2_ABS: "real_pow (real_abs x) 2 = real_pow x 2"
-  by (import hollight REAL_POW2_ABS)
-
-lemma REAL_LE_SQUARE_ABS: "real_le (real_abs x) (real_abs y) = real_le (real_pow x 2) (real_pow y 2)"
-  by (import hollight REAL_LE_SQUARE_ABS)
-
-lemma REAL_LT_SQUARE_ABS: "real_lt (real_abs x) (real_abs xa) = real_lt (real_pow x 2) (real_pow xa 2)"
-  by (import hollight REAL_LT_SQUARE_ABS)
-
-lemma REAL_EQ_SQUARE_ABS: "(real_abs x = real_abs xa) = (real_pow x 2 = real_pow xa 2)"
-  by (import hollight REAL_EQ_SQUARE_ABS)
-
-lemma REAL_LE_POW_2: "real_le (real_of_num 0) (real_pow x 2)"
-  by (import hollight REAL_LE_POW_2)
-
-lemma REAL_SOS_EQ_0: "(real_add (real_pow x 2) (real_pow y 2) = real_of_num 0) =
-(x = real_of_num 0 & y = real_of_num 0)"
-  by (import hollight REAL_SOS_EQ_0)
-
-lemma REAL_POW_ZERO: "real_pow (real_of_num 0) n =
-(if n = 0 then real_of_num 1 else real_of_num 0)"
-  by (import hollight REAL_POW_ZERO)
-
-lemma REAL_POW_MONO_INV: "real_le (real_of_num 0) x & real_le x (real_of_num 1) & n <= m
-==> real_le (real_pow x m) (real_pow x n)"
-  by (import hollight REAL_POW_MONO_INV)
-
-lemma REAL_POW_LE2_REV: "n ~= 0 & real_le (real_of_num 0) y & real_le (real_pow x n) (real_pow y n)
-==> real_le x y"
-  by (import hollight REAL_POW_LE2_REV)
-
-lemma REAL_POW_LT2_REV: "real_le (real_of_num 0) y & real_lt (real_pow x n) (real_pow y n)
-==> real_lt x y"
-  by (import hollight REAL_POW_LT2_REV)
-
-lemma REAL_POW_EQ: "x ~= 0 &
-real_le (real_of_num 0) xa &
-real_le (real_of_num 0) xb & real_pow xa x = real_pow xb x
-==> xa = xb"
-  by (import hollight REAL_POW_EQ)
-
-lemma REAL_POW_EQ_ABS: "n ~= 0 & real_pow x n = real_pow y n ==> real_abs x = real_abs y"
-  by (import hollight REAL_POW_EQ_ABS)
-
-lemma REAL_POW_EQ_1_IMP: "n ~= 0 & real_pow x n = real_of_num 1 ==> real_abs x = real_of_num 1"
-  by (import hollight REAL_POW_EQ_1_IMP)
-
-lemma REAL_POW_EQ_1: "(real_pow x n = real_of_num 1) =
-(real_abs x = real_of_num 1 & (real_lt x (real_of_num 0) --> even n) |
- n = 0)"
-  by (import hollight REAL_POW_EQ_1)
-
-lemma REAL_POW_LT2_ODD: "real_lt x y & odd n ==> real_lt (real_pow x n) (real_pow y n)"
-  by (import hollight REAL_POW_LT2_ODD)
-
-lemma REAL_POW_LE2_ODD: "real_le xa xb & odd x ==> real_le (real_pow xa x) (real_pow xb x)"
-  by (import hollight REAL_POW_LE2_ODD)
-
-lemma REAL_POW_LT2_ODD_EQ: "odd n ==> real_lt (real_pow x n) (real_pow y n) = real_lt x y"
-  by (import hollight REAL_POW_LT2_ODD_EQ)
-
-lemma REAL_POW_LE2_ODD_EQ: "odd n ==> real_le (real_pow x n) (real_pow y n) = real_le x y"
-  by (import hollight REAL_POW_LE2_ODD_EQ)
-
-lemma REAL_POW_EQ_ODD_EQ: "odd x ==> (real_pow xa x = real_pow xb x) = (xa = xb)"
-  by (import hollight REAL_POW_EQ_ODD_EQ)
-
-lemma REAL_POW_EQ_ODD: "odd n & real_pow x n = real_pow y n ==> x = y"
-  by (import hollight REAL_POW_EQ_ODD)
-
-lemma REAL_POW_EQ_EQ: "(real_pow x n = real_pow y n) =
-(if even n then n = 0 | real_abs x = real_abs y else x = y)"
-  by (import hollight REAL_POW_EQ_EQ)
-
-definition
-  real_sgn :: "hollight.real => hollight.real"  where
-  "real_sgn ==
-%u. if real_lt (real_of_num 0) u then real_of_num 1
-    else if real_lt u (real_of_num 0) then real_neg (real_of_num 1)
-         else real_of_num 0"
-
-lemma DEF_real_sgn: "real_sgn =
-(%u. if real_lt (real_of_num 0) u then real_of_num 1
-     else if real_lt u (real_of_num 0) then real_neg (real_of_num 1)
-          else real_of_num 0)"
-  by (import hollight DEF_real_sgn)
-
-lemma REAL_SGN_0: "real_sgn (real_of_num 0) = real_of_num 0"
-  by (import hollight REAL_SGN_0)
-
-lemma REAL_SGN_NEG: "real_sgn (real_neg x) = real_neg (real_sgn x)"
-  by (import hollight REAL_SGN_NEG)
-
-lemma REAL_SGN_ABS: "real_mul (real_sgn x) (real_abs x) = x"
-  by (import hollight REAL_SGN_ABS)
-
-lemma REAL_ABS_SGN: "real_abs (real_sgn x) = real_sgn (real_abs x)"
-  by (import hollight REAL_ABS_SGN)
-
-lemma REAL_SGN: "real_sgn x = real_div x (real_abs x)"
-  by (import hollight REAL_SGN)
-
-lemma REAL_SGN_MUL: "real_sgn (real_mul x xa) = real_mul (real_sgn x) (real_sgn xa)"
-  by (import hollight REAL_SGN_MUL)
-
-lemma REAL_SGN_INV: "real_sgn (real_inv x) = real_sgn x"
-  by (import hollight REAL_SGN_INV)
-
-lemma REAL_SGN_DIV: "real_sgn (real_div x xa) = real_div (real_sgn x) (real_sgn xa)"
-  by (import hollight REAL_SGN_DIV)
-
-lemma REAL_WLOG_LE: "(ALL x y. P x y = P y x) & (ALL x y. real_le x y --> P x y) ==> P x xa"
-  by (import hollight REAL_WLOG_LE)
-
-lemma REAL_WLOG_LT: "(ALL x. P x x) & (ALL x y. P x y = P y x) & (ALL x y. real_lt x y --> P x y)
-==> P x xa"
-  by (import hollight REAL_WLOG_LT)
-
-definition
-  DECIMAL :: "nat => nat => hollight.real"  where
-  "DECIMAL == %u ua. real_div (real_of_num u) (real_of_num ua)"
-
-lemma DEF_DECIMAL: "DECIMAL = (%u ua. real_div (real_of_num u) (real_of_num ua))"
-  by (import hollight DEF_DECIMAL)
-
-lemma RAT_LEMMA1: "y1 ~= real_of_num 0 & y2 ~= real_of_num 0
-==> real_add (real_div x1 y1) (real_div x2 y2) =
-    real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
-     (real_mul (real_inv y1) (real_inv y2))"
-  by (import hollight RAT_LEMMA1)
-
-lemma RAT_LEMMA2: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_add (real_div x1 y1) (real_div x2 y2) =
-    real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
-     (real_mul (real_inv y1) (real_inv y2))"
-  by (import hollight RAT_LEMMA2)
-
-lemma RAT_LEMMA3: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_sub (real_div x1 y1) (real_div x2 y2) =
-    real_mul (real_sub (real_mul x1 y2) (real_mul x2 y1))
-     (real_mul (real_inv y1) (real_inv y2))"
-  by (import hollight RAT_LEMMA3)
-
-lemma RAT_LEMMA4: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_le (real_div x1 y1) (real_div x2 y2) =
-    real_le (real_mul x1 y2) (real_mul x2 y1)"
-  by (import hollight RAT_LEMMA4)
-
-lemma RAT_LEMMA5: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> (real_div x1 y1 = real_div x2 y2) = (real_mul x1 y2 = real_mul x2 y1)"
-  by (import hollight RAT_LEMMA5)
-
-lemma REAL_INTEGRAL: "(ALL x. real_mul (real_of_num 0) x = real_of_num 0) &
-(ALL x y z. (real_add x y = real_add x z) = (y = z)) &
-(ALL w x y z.
-    (real_add (real_mul w y) (real_mul x z) =
-     real_add (real_mul w z) (real_mul x y)) =
-    (w = x | y = z))"
-  by (import hollight REAL_INTEGRAL)
-
-definition
-  integer :: "hollight.real => bool"  where
-  "integer == %u. EX n. real_abs u = real_of_num n"
-
-lemma DEF_integer: "integer = (%u. EX n. real_abs u = real_of_num n)"
-  by (import hollight DEF_integer)
-
-lemma is_int: "integer x = (EX n. x = real_of_num n | x = real_neg (real_of_num n))"
-  by (import hollight is_int)
-
-typedef (open) int = "Collect integer"  morphisms "real_of_int" "int_of_real"
-  apply (rule light_ex_imp_nonempty[where t="Eps integer"])
-  by (import hollight TYDEF_int)
-
-syntax
-  real_of_int :: _ 
-
-syntax
-  int_of_real :: _ 
-
-lemmas "TYDEF_int_@intern" = typedef_hol2hollight 
-  [where a="a :: hollight.int" and r=r ,
-   OF type_definition_int]
-
-lemma dest_int_rep: "EX n. hollight.real_of_int x = real_of_num n |
-      hollight.real_of_int x = real_neg (real_of_num n)"
-  by (import hollight dest_int_rep)
-
-definition
-  int_le :: "hollight.int => hollight.int => bool"  where
-  "int_le == %u ua. real_le (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_le: "int_le = (%u ua. real_le (hollight.real_of_int u) (hollight.real_of_int ua))"
-  by (import hollight DEF_int_le)
-
-definition
-  int_lt :: "hollight.int => hollight.int => bool"  where
-  "int_lt == %u ua. real_lt (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_lt: "int_lt = (%u ua. real_lt (hollight.real_of_int u) (hollight.real_of_int ua))"
-  by (import hollight DEF_int_lt)
-
-definition
-  int_ge :: "hollight.int => hollight.int => bool"  where
-  "int_ge == %u ua. real_ge (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_ge: "int_ge = (%u ua. real_ge (hollight.real_of_int u) (hollight.real_of_int ua))"
-  by (import hollight DEF_int_ge)
-
-definition
-  int_gt :: "hollight.int => hollight.int => bool"  where
-  "int_gt == %u ua. real_gt (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_gt: "int_gt = (%u ua. real_gt (hollight.real_of_int u) (hollight.real_of_int ua))"
-  by (import hollight DEF_int_gt)
-
-definition
-  int_of_num :: "nat => hollight.int"  where
-  "int_of_num == %u. int_of_real (real_of_num u)"
-
-lemma DEF_int_of_num: "int_of_num = (%u. int_of_real (real_of_num u))"
-  by (import hollight DEF_int_of_num)
-
-lemma int_of_num_th: "hollight.real_of_int (int_of_num x) = real_of_num x"
-  by (import hollight int_of_num_th)
-
-definition
-  int_neg :: "hollight.int => hollight.int"  where
-  "int_neg == %u. int_of_real (real_neg (hollight.real_of_int u))"
-
-lemma DEF_int_neg: "int_neg = (%u. int_of_real (real_neg (hollight.real_of_int u)))"
-  by (import hollight DEF_int_neg)
-
-lemma int_neg_th: "hollight.real_of_int (int_neg x) = real_neg (hollight.real_of_int x)"
-  by (import hollight int_neg_th)
-
-definition
-  int_add :: "hollight.int => hollight.int => hollight.int"  where
-  "int_add ==
-%u ua.
-   int_of_real (real_add (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_add: "int_add =
-(%u ua.
-    int_of_real
-     (real_add (hollight.real_of_int u) (hollight.real_of_int ua)))"
-  by (import hollight DEF_int_add)
-
-lemma int_add_th: "hollight.real_of_int (int_add x xa) =
-real_add (hollight.real_of_int x) (hollight.real_of_int xa)"
-  by (import hollight int_add_th)
-
-definition
-  int_sub :: "hollight.int => hollight.int => hollight.int"  where
-  "int_sub ==
-%u ua.
-   int_of_real (real_sub (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_sub: "int_sub =
-(%u ua.
-    int_of_real
-     (real_sub (hollight.real_of_int u) (hollight.real_of_int ua)))"
-  by (import hollight DEF_int_sub)
-
-lemma int_sub_th: "hollight.real_of_int (int_sub x xa) =
-real_sub (hollight.real_of_int x) (hollight.real_of_int xa)"
-  by (import hollight int_sub_th)
-
-definition
-  int_mul :: "hollight.int => hollight.int => hollight.int"  where
-  "int_mul ==
-%u ua.
-   int_of_real (real_mul (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_mul: "int_mul =
-(%u ua.
-    int_of_real
-     (real_mul (hollight.real_of_int u) (hollight.real_of_int ua)))"
-  by (import hollight DEF_int_mul)
-
-lemma int_mul_th: "hollight.real_of_int (int_mul x y) =
-real_mul (hollight.real_of_int x) (hollight.real_of_int y)"
-  by (import hollight int_mul_th)
-
-definition
-  int_abs :: "hollight.int => hollight.int"  where
-  "int_abs == %u. int_of_real (real_abs (hollight.real_of_int u))"
-
-lemma DEF_int_abs: "int_abs = (%u. int_of_real (real_abs (hollight.real_of_int u)))"
-  by (import hollight DEF_int_abs)
-
-lemma int_abs_th: "hollight.real_of_int (int_abs x) = real_abs (hollight.real_of_int x)"
-  by (import hollight int_abs_th)
-
-definition
-  int_sgn :: "hollight.int => hollight.int"  where
-  "int_sgn == %u. int_of_real (real_sgn (hollight.real_of_int u))"
-
-lemma DEF_int_sgn: "int_sgn = (%u. int_of_real (real_sgn (hollight.real_of_int u)))"
-  by (import hollight DEF_int_sgn)
-
-lemma int_sgn_th: "hollight.real_of_int (int_sgn x) = real_sgn (hollight.real_of_int x)"
-  by (import hollight int_sgn_th)
-
-definition
-  int_max :: "hollight.int => hollight.int => hollight.int"  where
-  "int_max ==
-%u ua.
-   int_of_real (real_max (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_max: "int_max =
-(%u ua.
-    int_of_real
-     (real_max (hollight.real_of_int u) (hollight.real_of_int ua)))"
-  by (import hollight DEF_int_max)
-
-lemma int_max_th: "hollight.real_of_int (int_max x y) =
-real_max (hollight.real_of_int x) (hollight.real_of_int y)"
-  by (import hollight int_max_th)
-
-definition
-  int_min :: "hollight.int => hollight.int => hollight.int"  where
-  "int_min ==
-%u ua.
-   int_of_real (real_min (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_min: "int_min =
-(%u ua.
-    int_of_real
-     (real_min (hollight.real_of_int u) (hollight.real_of_int ua)))"
-  by (import hollight DEF_int_min)
-
-lemma int_min_th: "hollight.real_of_int (int_min x y) =
-real_min (hollight.real_of_int x) (hollight.real_of_int y)"
-  by (import hollight int_min_th)
-
-definition
-  int_pow :: "hollight.int => nat => hollight.int"  where
-  "int_pow == %u ua. int_of_real (real_pow (hollight.real_of_int u) ua)"
-
-lemma DEF_int_pow: "int_pow = (%u ua. int_of_real (real_pow (hollight.real_of_int u) ua))"
-  by (import hollight DEF_int_pow)
-
-lemma int_pow_th: "hollight.real_of_int (int_pow x xa) = real_pow (hollight.real_of_int x) xa"
-  by (import hollight int_pow_th)
-
-lemma INT_IMAGE: "(EX n. x = int_of_num n) | (EX n. x = int_neg (int_of_num n))"
-  by (import hollight INT_IMAGE)
-
-lemma INT_LT_DISCRETE: "int_lt x y = int_le (int_add x (int_of_num 1)) y"
-  by (import hollight INT_LT_DISCRETE)
-
-lemma INT_GT_DISCRETE: "int_gt x xa = int_ge x (int_add xa (int_of_num 1))"
-  by (import hollight INT_GT_DISCRETE)
-
-lemma INT_FORALL_POS: "(ALL n. P (int_of_num n)) = (ALL i. int_le (int_of_num 0) i --> P i)"
-  by (import hollight INT_FORALL_POS)
-
-lemma INT_EXISTS_POS: "(EX n. P (int_of_num n)) = (EX i. int_le (int_of_num 0) i & P i)"
-  by (import hollight INT_EXISTS_POS)
-
-lemma INT_FORALL_ABS: "(ALL n. x (int_of_num n)) = (ALL xa. x (int_abs xa))"
-  by (import hollight INT_FORALL_ABS)
-
-lemma INT_EXISTS_ABS: "(EX n. P (int_of_num n)) = (EX x. P (int_abs x))"
-  by (import hollight INT_EXISTS_ABS)
-
-lemma INT_ABS_MUL_1: "(int_abs (int_mul x y) = int_of_num 1) =
-(int_abs x = int_of_num 1 & int_abs y = int_of_num 1)"
-  by (import hollight INT_ABS_MUL_1)
-
-lemma INT_WOP: "(EX x. int_le (int_of_num 0) x & P x) =
-(EX x. int_le (int_of_num 0) x &
-       P x & (ALL y. int_le (int_of_num 0) y & P y --> int_le x y))"
-  by (import hollight INT_WOP)
-
-lemma INT_POW: "int_pow x 0 = int_of_num 1 &
-(ALL xa. int_pow x (Suc xa) = int_mul x (int_pow x xa))"
-  by (import hollight INT_POW)
-
-lemma INT_ABS: "int_abs x = (if int_le (int_of_num 0) x then x else int_neg x)"
-  by (import hollight INT_ABS)
-
-lemma INT_GE: "int_ge x xa = int_le xa x"
-  by (import hollight INT_GE)
-
-lemma INT_GT: "int_gt x xa = int_lt xa x"
-  by (import hollight INT_GT)
-
-lemma INT_LT: "int_lt x xa = (~ int_le xa x)"
-  by (import hollight INT_LT)
-
-lemma INT_ARCH: "d ~= int_of_num 0 ==> EX c. int_lt x (int_mul c d)"
-  by (import hollight INT_ARCH)
-
-lemma INT_DIVMOD_EXIST_0: "EX x xa.
-   if n = int_of_num 0 then x = int_of_num 0 & xa = m
-   else int_le (int_of_num 0) xa &
-        int_lt xa (int_abs n) & m = int_add (int_mul x n) xa"
-  by (import hollight INT_DIVMOD_EXIST_0)
-
-consts
-  div :: "hollight.int => hollight.int => hollight.int" ("div")
-
-defs
-  div_def: "div ==
-SOME q.
-   EX r. ALL m n.
-            if n = int_of_num 0 then q m n = int_of_num 0 & r m n = m
-            else int_le (int_of_num 0) (r m n) &
-                 int_lt (r m n) (int_abs n) &
-                 m = int_add (int_mul (q m n) n) (r m n)"
-
-lemma DEF_div: "div =
-(SOME q.
-    EX r. ALL m n.
-             if n = int_of_num 0 then q m n = int_of_num 0 & r m n = m
-             else int_le (int_of_num 0) (r m n) &
-                  int_lt (r m n) (int_abs n) &
-                  m = int_add (int_mul (q m n) n) (r m n))"
-  by (import hollight DEF_div)
-
-definition
-  rem :: "hollight.int => hollight.int => hollight.int"  where
-  "rem ==
-SOME r.
-   ALL m n.
-      if n = int_of_num 0 then div m n = int_of_num 0 & r m n = m
-      else int_le (int_of_num 0) (r m n) &
-           int_lt (r m n) (int_abs n) &
-           m = int_add (int_mul (div m n) n) (r m n)"
-
-lemma DEF_rem: "rem =
-(SOME r.
-    ALL m n.
-       if n = int_of_num 0 then div m n = int_of_num 0 & r m n = m
-       else int_le (int_of_num 0) (r m n) &
-            int_lt (r m n) (int_abs n) &
-            m = int_add (int_mul (div m n) n) (r m n))"
-  by (import hollight DEF_rem)
-
-lemma INT_DIVISION: "n ~= int_of_num 0
-==> m = int_add (int_mul (div m n) n) (rem m n) &
-    int_le (int_of_num 0) (rem m n) & int_lt (rem m n) (int_abs n)"
-  by (import hollight INT_DIVISION)
-
-lemma sth: "(ALL x y z. int_add x (int_add y z) = int_add (int_add x y) z) &
-(ALL x y. int_add x y = int_add y x) &
-(ALL x. int_add (int_of_num 0) x = x) &
-(ALL x y z. int_mul x (int_mul y z) = int_mul (int_mul x y) z) &
-(ALL x y. int_mul x y = int_mul y x) &
-(ALL x. int_mul (int_of_num 1) x = x) &
-(ALL x. int_mul (int_of_num 0) x = int_of_num 0) &
-(ALL x y z. int_mul x (int_add y z) = int_add (int_mul x y) (int_mul x z)) &
-(ALL x. int_pow x 0 = int_of_num 1) &
-(ALL x xa. int_pow x (Suc xa) = int_mul x (int_pow x xa))"
-  by (import hollight sth)
-
-lemma INT_INTEGRAL: "(ALL x. int_mul (int_of_num 0) x = int_of_num 0) &
-(ALL x y z. (int_add x y = int_add x z) = (y = z)) &
-(ALL w x y z.
-    (int_add (int_mul w y) (int_mul x z) =
-     int_add (int_mul w z) (int_mul x y)) =
-    (w = x | y = z))"
-  by (import hollight INT_INTEGRAL)
-
-lemma INT_DIVMOD_UNIQ: "m = int_add (int_mul q n) r & int_le (int_of_num 0) r & int_lt r (int_abs n)
-==> div m n = q & rem m n = r"
-  by (import hollight INT_DIVMOD_UNIQ)
-
-consts
-  eqeq :: "'A => 'A => ('A => 'A => bool) => bool" 
-
-defs
-  eqeq_def: "hollight.eqeq == %(u::'A) (ua::'A) ub::'A => 'A => bool. ub u ua"
-
-lemma DEF__equal__equal_: "hollight.eqeq = (%(u::'A) (ua::'A) ub::'A => 'A => bool. ub u ua)"
-  by (import hollight DEF__equal__equal_)
-
-definition
-  real_mod :: "hollight.real => hollight.real => hollight.real => bool"  where
-  "real_mod == %u ua ub. EX q. integer q & real_sub ua ub = real_mul q u"
-
-lemma DEF_real_mod: "real_mod = (%u ua ub. EX q. integer q & real_sub ua ub = real_mul q u)"
-  by (import hollight DEF_real_mod)
-
-definition
-  int_divides :: "hollight.int => hollight.int => bool"  where
-  "int_divides == %u ua. EX x. ua = int_mul u x"
-
-lemma DEF_int_divides: "int_divides = (%u ua. EX x. ua = int_mul u x)"
-  by (import hollight DEF_int_divides)
-
-consts
-  int_mod :: "hollight.int => hollight.int => hollight.int => bool" 
-
-defs
-  int_mod_def: "hollight.int_mod == %u ua ub. int_divides u (int_sub ua ub)"
-
-lemma DEF_int_mod: "hollight.int_mod = (%u ua ub. int_divides u (int_sub ua ub))"
-  by (import hollight DEF_int_mod)
-
-lemma int_congruent: "hollight.eqeq x xa (hollight.int_mod xb) =
-(EX d. int_sub x xa = int_mul xb d)"
-  by (import hollight int_congruent)
-
-consts
-  int_coprime :: "hollight.int * hollight.int => bool" 
-
-defs
-  int_coprime_def: "hollight.int_coprime ==
-%u. EX x y. int_add (int_mul (fst u) x) (int_mul (snd u) y) = int_of_num 1"
-
-lemma DEF_int_coprime: "hollight.int_coprime =
-(%u. EX x y. int_add (int_mul (fst u) x) (int_mul (snd u) y) = int_of_num 1)"
-  by (import hollight DEF_int_coprime)
-
-lemma FORALL_UNCURRY: "All (P::('A => 'B => 'C) => bool) =
-(ALL f::'A * 'B => 'C. P (%(a::'A) b::'B. f (a, b)))"
-  by (import hollight FORALL_UNCURRY)
-
-lemma EXISTS_UNCURRY: "Ex (x::('A => 'B => 'C) => bool) =
-(EX f::'A * 'B => 'C. x (%(a::'A) b::'B. f (a, b)))"
-  by (import hollight EXISTS_UNCURRY)
-
-lemma WF_INT_MEASURE: "(ALL x::'A. int_le (int_of_num (0::nat)) ((m::'A => hollight.int) x)) &
-(ALL x::'A. (ALL y::'A. int_lt (m y) (m x) --> (P::'A => bool) y) --> P x)
-==> P (x::'A)"
-  by (import hollight WF_INT_MEASURE)
-
-lemma WF_INT_MEASURE_2: "(ALL (x::'A) y::'B.
-    int_le (int_of_num (0::nat)) ((m::'A => 'B => hollight.int) x y)) &
-(ALL (x::'A) y::'B.
-    (ALL (x'::'A) y'::'B.
-        int_lt (m x' y') (m x y) --> (P::'A => 'B => bool) x' y') -->
-    P x y)
-==> P (x::'A) (xa::'B)"
-  by (import hollight WF_INT_MEASURE_2)
-
-lemma INT_GCD_EXISTS: "EX d. int_divides d a &
-      int_divides d b & (EX x y. d = int_add (int_mul a x) (int_mul b y))"
-  by (import hollight INT_GCD_EXISTS)
-
-lemma INT_GCD_EXISTS_POS: "EX d. int_le (int_of_num 0) d &
-      int_divides d a &
-      int_divides d b & (EX x y. d = int_add (int_mul a x) (int_mul b y))"
-  by (import hollight INT_GCD_EXISTS_POS)
-
-consts
-  int_gcd :: "hollight.int * hollight.int => hollight.int" 
-
-defs
-  int_gcd_def: "hollight.int_gcd ==
-SOME d.
-   ALL a b.
-      int_le (int_of_num 0) (d (a, b)) &
-      int_divides (d (a, b)) a &
-      int_divides (d (a, b)) b &
-      (EX x y. d (a, b) = int_add (int_mul a x) (int_mul b y))"
-
-lemma DEF_int_gcd: "hollight.int_gcd =
-(SOME d.
-    ALL a b.
-       int_le (int_of_num 0) (d (a, b)) &
-       int_divides (d (a, b)) a &
-       int_divides (d (a, b)) b &
-       (EX x y. d (a, b) = int_add (int_mul a x) (int_mul b y)))"
-  by (import hollight DEF_int_gcd)
-
-definition
-  num_of_int :: "hollight.int => nat"  where
-  "num_of_int == %u. SOME n. int_of_num n = u"
-
-lemma DEF_num_of_int: "num_of_int = (%u. SOME n. int_of_num n = u)"
-  by (import hollight DEF_num_of_int)
-
-lemma NUM_OF_INT_OF_NUM: "num_of_int (int_of_num x) = x"
-  by (import hollight NUM_OF_INT_OF_NUM)
-
-lemma INT_OF_NUM_OF_INT: "int_le (int_of_num 0) x ==> int_of_num (num_of_int x) = x"
-  by (import hollight INT_OF_NUM_OF_INT)
-
-lemma NUM_OF_INT: "int_le (int_of_num 0) x = (int_of_num (num_of_int x) = x)"
-  by (import hollight NUM_OF_INT)
-
-definition
-  num_divides :: "nat => nat => bool"  where
-  "num_divides == %u ua. int_divides (int_of_num u) (int_of_num ua)"
-
-lemma DEF_num_divides: "num_divides = (%u ua. int_divides (int_of_num u) (int_of_num ua))"
-  by (import hollight DEF_num_divides)
-
-definition
-  num_mod :: "nat => nat => nat => bool"  where
-  "num_mod ==
-%u ua ub. hollight.int_mod (int_of_num u) (int_of_num ua) (int_of_num ub)"
-
-lemma DEF_num_mod: "num_mod =
-(%u ua ub. hollight.int_mod (int_of_num u) (int_of_num ua) (int_of_num ub))"
-  by (import hollight DEF_num_mod)
-
-lemma num_congruent: "hollight.eqeq x xa (num_mod xb) =
-hollight.eqeq (int_of_num x) (int_of_num xa)
- (hollight.int_mod (int_of_num xb))"
-  by (import hollight num_congruent)
-
-definition
-  num_coprime :: "nat * nat => bool"  where
-  "num_coprime ==
-%u. hollight.int_coprime (int_of_num (fst u), int_of_num (snd u))"
-
-lemma DEF_num_coprime: "num_coprime =
-(%u. hollight.int_coprime (int_of_num (fst u), int_of_num (snd u)))"
-  by (import hollight DEF_num_coprime)
-
-definition
-  num_gcd :: "nat * nat => nat"  where
-  "num_gcd ==
-%u. num_of_int (hollight.int_gcd (int_of_num (fst u), int_of_num (snd u)))"
-
-lemma DEF_num_gcd: "num_gcd =
-(%u. num_of_int (hollight.int_gcd (int_of_num (fst u), int_of_num (snd u))))"
-  by (import hollight DEF_num_gcd)
-
-lemma NUM_GCD: "int_of_num (num_gcd (x, xa)) =
-hollight.int_gcd (int_of_num x, int_of_num xa)"
-  by (import hollight NUM_GCD)
-
-lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_43295 => bool) => bool) x::'q_43295.
-    (x : {v::'q_43295. P (SETSPEC v)}) =
-    P (%(p::bool) t::'q_43295. p & x = t)) &
-(ALL (p::'q_43326 => bool) x::'q_43326.
-    (x : {v::'q_43326. EX y::'q_43326. p y & v = y}) = p x) &
-(ALL (P::(bool => 'q_43354 => bool) => bool) x::'q_43354.
-    x \<in> {v::'q_43354. P (SETSPEC v)} =
-    P (%(p::bool) t::'q_43354. p & x = t)) &
-(ALL (p::'q_43383 => bool) x::'q_43383.
-    x \<in> {v::'q_43383. EX y::'q_43383. p y & v = y} = p x) &
-(ALL (p::'q_43400 => bool) x::'q_43400. p x \<longleftrightarrow> p x)"
-  by (import hollight IN_ELIM_THM)
-
-lemma INSERT: "insert (x::'A) (s::'A set) = {u::'A. EX y::'A. (y : s | y = x) & u = y}"
-  by (import hollight INSERT)
-
-definition
-  SING :: "('A set) => bool"  where
-  "SING == %u::'A set. EX x::'A. u = {x}"
-
-lemma DEF_SING: "SING = (%u::'A set. EX x::'A. u = {x})"
-  by (import hollight DEF_SING)
-
-definition
-  INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool"  where
-  "INJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
-   (ALL x::'A. x : ua --> u x : ub) &
-   (ALL (x::'A) y::'A. x : ua & y : ua & u x = u y --> x = y)"
-
-lemma DEF_INJ: "INJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
-    (ALL x::'A. x : ua --> u x : ub) &
-    (ALL (x::'A) y::'A. x : ua & y : ua & u x = u y --> x = y))"
-  by (import hollight DEF_INJ)
-
-definition
-  SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool"  where
-  "SURJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
-   (ALL x::'A. x : ua --> u x : ub) &
-   (ALL x::'B. x : ub --> (EX y::'A. y : ua & u y = x))"
-
-lemma DEF_SURJ: "SURJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
-    (ALL x::'A. x : ua --> u x : ub) &
-    (ALL x::'B. x : ub --> (EX y::'A. y : ua & u y = x)))"
-  by (import hollight DEF_SURJ)
-
-definition
-  BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool"  where
-  "BIJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub"
-
-lemma DEF_BIJ: "BIJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub)"
-  by (import hollight DEF_BIJ)
-
-definition
-  REST :: "('A => bool) => 'A => bool"  where
-  "REST == %u::'A => bool. u - {Eps u}"
-
-lemma DEF_REST: "REST = (%u::'A => bool. u - {Eps u})"
-  by (import hollight DEF_REST)
-
-lemma NOT_IN_EMPTY: "(x::'A) ~: {}"
-  by (import hollight NOT_IN_EMPTY)
-
-lemma IN_UNIONS: "((xa::'A) : Union (x::('A => bool) => bool)) =
-(EX t::'A => bool. t : x & xa : t)"
-  by (import hollight IN_UNIONS)
-
-lemma IN_INTERS: "((xa::'A) : Inter (x::('A => bool) => bool)) =
-(ALL t::'A => bool. t : x --> xa : t)"
-  by (import hollight IN_INTERS)
-
-lemma IN_DELETE: "((xa::'A) : (x::'A => bool) - {xb::'A}) = (xa : x & xa ~= xb)"
-  by (import hollight IN_DELETE)
-
-lemma IN_IMAGE: "((x::'B) : (xb::'A => 'B) ` (xa::'A => bool)) =
-(EX xc::'A. x = xb xc & xc : xa)"
-  by (import hollight IN_IMAGE)
-
-lemma IN_REST: "((x::'A) : REST (xa::'A => bool)) = (x : xa & x ~= Eps xa)"
-  by (import hollight IN_REST)
-
-lemma FORALL_IN_INSERT: "(ALL xc::'q_44214.
-    xc : insert (xa::'q_44214) (xb::'q_44214 => bool) -->
-    (x::'q_44214 => bool) xc) =
-(x xa & (ALL xa::'q_44214. xa : xb --> x xa))"
-  by (import hollight FORALL_IN_INSERT)
-
-lemma EXISTS_IN_INSERT: "(EX xc::'q_44255.
-    xc : insert (xa::'q_44255) (xb::'q_44255 => bool) &
-    (x::'q_44255 => bool) xc) =
-(x xa | (EX xa::'q_44255. xa : xb & x xa))"
-  by (import hollight EXISTS_IN_INSERT)
-
-lemma CHOICE_DEF: "(x::'A => bool) ~= {} ==> Eps x : x"
-  by (import hollight CHOICE_DEF)
-
-lemma NOT_EQUAL_SETS: "((x::'A => bool) ~= (xa::'A => bool)) = (EX xb::'A. (xb : xa) = (xb ~: x))"
-  by (import hollight NOT_EQUAL_SETS)
-
-lemma EMPTY_NOT_UNIV: "(op ~=::('A::type => bool) => ('A::type => bool) => bool)
- ({}::'A::type => bool) (UNIV::'A::type => bool)"
-  by (import hollight EMPTY_NOT_UNIV)
-
-lemma EQ_UNIV: "(ALL x::'A. x : (s::'A => bool)) = (s = UNIV)"
-  by (import hollight EQ_UNIV)
-
-lemma SING_SUBSET: "({xa::'q_44493} <= (x::'q_44493 => bool)) = (xa : x)"
-  by (import hollight SING_SUBSET)
-
-lemma PSUBSET_UNIV: "((x::'A => bool) < UNIV) = (EX xa::'A. xa ~: x)"
-  by (import hollight PSUBSET_UNIV)
-
-lemma PSUBSET_ALT: "((x::'A => bool) < (xa::'A => bool)) =
-(x <= xa & (EX a::'A. a : xa & a ~: x))"
-  by (import hollight PSUBSET_ALT)
-
-lemma SUBSET_UNION: "(ALL (x::'A => bool) xa::'A => bool. x <= x Un xa) &
-(ALL (x::'A => bool) xa::'A => bool. x <= xa Un x)"
-  by (import hollight SUBSET_UNION)
-
-lemma UNION_EMPTY: "(ALL x::'A => bool. {} Un x = x) & (ALL x::'A => bool. x Un {} = x)"
-  by (import hollight UNION_EMPTY)
-
-lemma UNION_UNIV: "(ALL x::'A => bool. UNIV Un x = UNIV) &
-(ALL x::'A => bool. x Un UNIV = UNIV)"
-  by (import hollight UNION_UNIV)
-
-lemma INTER_SUBSET: "(ALL (x::'A => bool) xa::'A => bool. x Int xa <= x) &
-(ALL (x::'A => bool) xa::'A => bool. xa Int x <= x)"
-  by (import hollight INTER_SUBSET)
-
-lemma INTER_EMPTY: "(ALL x::'A => bool. {} Int x = {}) & (ALL x::'A => bool. x Int {} = {})"
-  by (import hollight INTER_EMPTY)
-
-lemma INTER_UNIV: "(ALL x::'A => bool. UNIV Int x = x) & (ALL x::'A => bool. x Int UNIV = x)"
-  by (import hollight INTER_UNIV)
-
-lemma IN_DISJOINT: "((x::'A => bool) Int (xa::'A => bool) = {}) =
-(~ (EX xb::'A. xb : x & xb : xa))"
-  by (import hollight IN_DISJOINT)
-
-lemma DISJOINT_SYM: "((x::'A => bool) Int (xa::'A => bool) = {}) = (xa Int x = {})"
-  by (import hollight DISJOINT_SYM)
-
-lemma DISJOINT_EMPTY: "{} Int (x::'A => bool) = {} & x Int {} = {}"
-  by (import hollight DISJOINT_EMPTY)
-
-lemma DISJOINT_EMPTY_REFL: "((x::'A => bool) = {}) = (x Int x = {})"
-  by (import hollight DISJOINT_EMPTY_REFL)
-
-lemma DISJOINT_UNION: "(((x::'A => bool) Un (xa::'A => bool)) Int (xb::'A => bool) = {}) =
-(x Int xb = {} & xa Int xb = {})"
-  by (import hollight DISJOINT_UNION)
-
-lemma DECOMPOSITION: "((x::'A) : (s::'A => bool)) = (EX t::'A => bool. s = insert x t & x ~: t)"
-  by (import hollight DECOMPOSITION)
-
-lemma SET_CASES: "(s::'A => bool) = {} | (EX (x::'A) t::'A => bool. s = insert x t & x ~: t)"
-  by (import hollight SET_CASES)
-
-lemma ABSORPTION: "((x::'A) : (xa::'A => bool)) = (insert x xa = xa)"
-  by (import hollight ABSORPTION)
-
-lemma INSERT_UNIV: "insert (x::'A) UNIV = UNIV"
-  by (import hollight INSERT_UNIV)
-
-lemma INSERT_UNION: "insert (x::'A) (s::'A => bool) Un (t::'A => bool) =
-(if x : t then s Un t else insert x (s Un t))"
-  by (import hollight INSERT_UNION)
-
-lemma DISJOINT_INSERT: "(insert (x::'A) (xa::'A => bool) Int (xb::'A => bool) = {}) =
-(xa Int xb = {} & x ~: xb)"
-  by (import hollight DISJOINT_INSERT)
-
-lemma INSERT_AC: "insert (x::'q_45764) (insert (y::'q_45764) (s::'q_45764 => bool)) =
-insert y (insert x s) &
-insert x (insert x s) = insert x s"
-  by (import hollight INSERT_AC)
-
-lemma INTER_ACI: "(p::'q_45831 => bool) Int (q::'q_45831 => bool) = q Int p &
-p Int q Int (r::'q_45831 => bool) = p Int (q Int r) &
-p Int (q Int r) = q Int (p Int r) & p Int p = p & p Int (p Int q) = p Int q"
-  by (import hollight INTER_ACI)
-
-lemma UNION_ACI: "(p::'q_45897 => bool) Un (q::'q_45897 => bool) = q Un p &
-p Un q Un (r::'q_45897 => bool) = p Un (q Un r) &
-p Un (q Un r) = q Un (p Un r) & p Un p = p & p Un (p Un q) = p Un q"
-  by (import hollight UNION_ACI)
-
-lemma DELETE_NON_ELEMENT: "((x::'A) ~: (xa::'A => bool)) = (xa - {x} = xa)"
-  by (import hollight DELETE_NON_ELEMENT)
-
-lemma IN_DELETE_EQ: "(((x::'A) : (s::'A => bool)) = ((x'::'A) : s)) =
-((x : s - {x'}) = (x' : s - {x}))"
-  by (import hollight IN_DELETE_EQ)
-
-lemma EMPTY_DELETE: "{} - {x::'A} = {}"
-  by (import hollight EMPTY_DELETE)
-
-lemma DELETE_DELETE: "(xa::'A => bool) - {x::'A} - {x} = xa - {x}"
-  by (import hollight DELETE_DELETE)
-
-lemma DELETE_COMM: "(xb::'A => bool) - {x::'A} - {xa::'A} = xb - {xa} - {x}"
-  by (import hollight DELETE_COMM)
-
-lemma DELETE_SUBSET: "(xa::'A => bool) - {x::'A} <= xa"
-  by (import hollight DELETE_SUBSET)
-
-lemma SUBSET_DELETE: "((xa::'A => bool) <= (xb::'A => bool) - {x::'A}) = (x ~: xa & xa <= xb)"
-  by (import hollight SUBSET_DELETE)
-
-lemma SUBSET_INSERT_DELETE: "((xa::'A => bool) <= insert (x::'A) (xb::'A => bool)) = (xa - {x} <= xb)"
-  by (import hollight SUBSET_INSERT_DELETE)
-
-lemma PSUBSET_INSERT_SUBSET: "((x::'A => bool) < (xa::'A => bool)) =
-(EX xb::'A. xb ~: x & insert xb x <= xa)"
-  by (import hollight PSUBSET_INSERT_SUBSET)
-
-lemma PSUBSET_MEMBER: "((x::'A => bool) < (xa::'A => bool)) =
-(x <= xa & (EX y::'A. y : xa & y ~: x))"
-  by (import hollight PSUBSET_MEMBER)
-
-lemma DELETE_INSERT: "insert (x::'A) (s::'A => bool) - {y::'A} =
-(if x = y then s - {y} else insert x (s - {y}))"
-  by (import hollight DELETE_INSERT)
-
-lemma DELETE_INTER: "((x::'A => bool) - {xb::'A}) Int (xa::'A => bool) = x Int xa - {xb}"
-  by (import hollight DELETE_INTER)
-
-lemma DISJOINT_DELETE_SYM: "(((x::'A => bool) - {xb::'A}) Int (xa::'A => bool) = {}) =
-((xa - {xb}) Int x = {})"
-  by (import hollight DISJOINT_DELETE_SYM)
-
-lemma FORALL_IN_UNIONS: "(ALL x::'q_46386.
-    x : Union (s::('q_46386 => bool) => bool) --> (P::'q_46386 => bool) x) =
-(ALL (t::'q_46386 => bool) x::'q_46386. t : s & x : t --> P x)"
-  by (import hollight FORALL_IN_UNIONS)
-
-lemma EXISTS_IN_UNIONS: "(EX x::'q_46428.
-    x : Union (s::('q_46428 => bool) => bool) & (P::'q_46428 => bool) x) =
-(EX (t::'q_46428 => bool) x::'q_46428. t : s & x : t & P x)"
-  by (import hollight EXISTS_IN_UNIONS)
-
-lemma EMPTY_UNIONS: "(Union (x::('q_46454 => bool) => bool) = {}) =
-(ALL xa::'q_46454 => bool. xa : x --> xa = {})"
-  by (import hollight EMPTY_UNIONS)
-
-lemma INTER_UNIONS: "(ALL (x::('q_46493 => bool) => bool) xa::'q_46493 => bool.
-    Union x Int xa =
-    Union
-     {u::'q_46493 => bool.
-      EX xb::'q_46493 => bool. xb : x & u = xb Int xa}) &
-(ALL (x::('q_46529 => bool) => bool) xa::'q_46529 => bool.
-    xa Int Union x =
-    Union
-     {u::'q_46529 => bool. EX xb::'q_46529 => bool. xb : x & u = xa Int xb})"
-  by (import hollight INTER_UNIONS)
-
-lemma UNIONS_SUBSET: "(Union (x::('q_46545 => bool) => bool) <= (xa::'q_46545 => bool)) =
-(ALL xb::'q_46545 => bool. xb : x --> xb <= xa)"
-  by (import hollight UNIONS_SUBSET)
-
-lemma IMAGE_CLAUSES: "(f::'q_46676 => 'q_46680) ` {} = {} &
-f ` insert (x::'q_46676) (s::'q_46676 => bool) = insert (f x) (f ` s)"
-  by (import hollight IMAGE_CLAUSES)
-
-lemma IMAGE_INTER_INJ: "(!!(xa::'q_46846) y::'q_46846.
-    (x::'q_46846 => 'q_46857) xa = x y ==> xa = y)
-==> x ` ((xa::'q_46846 => bool) Int (xb::'q_46846 => bool)) =
-    x ` xa Int x ` xb"
-  by (import hollight IMAGE_INTER_INJ)
-
-lemma IMAGE_DIFF_INJ: "(!!(xa::'q_46900) y::'q_46900.
-    (x::'q_46900 => 'q_46911) xa = x y ==> xa = y)
-==> x ` ((xa::'q_46900 => bool) - (xb::'q_46900 => bool)) = x ` xa - x ` xb"
-  by (import hollight IMAGE_DIFF_INJ)
-
-lemma IMAGE_DELETE_INJ: "(!!xa::'q_46958.
-    (x::'q_46958 => 'q_46957) xa = x (xb::'q_46958) ==> xa = xb)
-==> x ` ((xa::'q_46958 => bool) - {xb}) = x ` xa - {x xb}"
-  by (import hollight IMAGE_DELETE_INJ)
-
-lemma FORALL_IN_IMAGE: "(ALL xb::'q_47016.
-    xb : (x::'q_47017 => 'q_47016) ` (xa::'q_47017 => bool) -->
-    (P::'q_47016 => bool) xb) =
-(ALL xb::'q_47017. xb : xa --> P (x xb))"
-  by (import hollight FORALL_IN_IMAGE)
-
-lemma EXISTS_IN_IMAGE: "(EX xb::'q_47052.
-    xb : (x::'q_47053 => 'q_47052) ` (xa::'q_47053 => bool) &
-    (P::'q_47052 => bool) xb) =
-(EX xb::'q_47053. xb : xa & P (x xb))"
-  by (import hollight EXISTS_IN_IMAGE)
-
-lemma FORALL_SUBSET_IMAGE: "(ALL xc<=(xa::'q_47140 => 'q_47156) ` (xb::'q_47140 => bool).
-    (x::('q_47156 => bool) => bool) xc) =
-(ALL t<=xb. x (xa ` t))"
-  by (import hollight FORALL_SUBSET_IMAGE)
-
-lemma EXISTS_SUBSET_IMAGE: "(EX xc<=(xa::'q_47183 => 'q_47199) ` (xb::'q_47183 => bool).
-    (x::('q_47199 => bool) => bool) xc) =
-(EX t<=xb. x (xa ` t))"
-  by (import hollight EXISTS_SUBSET_IMAGE)
-
-lemma SIMPLE_IMAGE: "{u::'q_47262.
- EX xb::'q_47258.
-    xb : (xa::'q_47258 => bool) & u = (x::'q_47258 => 'q_47262) xb} =
-x ` xa"
-  by (import hollight SIMPLE_IMAGE)
-
-lemma SIMPLE_IMAGE_GEN: "{u::'q_47292.
- EX xa::'q_47305.
-    (P::'q_47305 => bool) xa & u = (x::'q_47305 => 'q_47292) xa} =
-x ` {u::'q_47305. EX x::'q_47305. P x & u = x}"
-  by (import hollight SIMPLE_IMAGE_GEN)
-
-lemma IMAGE_UNIONS: "(x::'q_47323 => 'q_47332) ` Union (xa::('q_47323 => bool) => bool) =
-Union (op ` x ` xa)"
-  by (import hollight IMAGE_UNIONS)
-
-lemma SURJECTIVE_IMAGE_EQ: "(ALL y::'q_47396.
-    y : (xa::'q_47396 => bool) -->
-    (EX x::'q_47400. (f::'q_47400 => 'q_47396) x = y)) &
-(ALL xb::'q_47400. (f xb : xa) = (xb : (x::'q_47400 => bool)))
-==> f ` x = xa"
-  by (import hollight SURJECTIVE_IMAGE_EQ)
-
-lemma EMPTY_GSPEC: "{u::'q_47425. Ex (SETSPEC u False)} = {}"
-  by (import hollight EMPTY_GSPEC)
-
-lemma SING_GSPEC: "(ALL x::'q_47454. {u::'q_47454. EX xa::'q_47454. xa = x & u = xa} = {x}) &
-(ALL x::'q_47480. {u::'q_47480. EX xa::'q_47480. x = xa & u = xa} = {x})"
-  by (import hollight SING_GSPEC)
-
-lemma IN_ELIM_PAIR_THM: "((xa::'q_47526, xb::'q_47525)
- : {xa::'q_47526 * 'q_47525.
-    EX (xb::'q_47526) y::'q_47525.
-       (x::'q_47526 => 'q_47525 => bool) xb y & xa = (xb, y)}) =
-x xa xb"
-  by (import hollight IN_ELIM_PAIR_THM)
-
-lemma SET_PAIR_THM: "{u::'q_47570 * 'q_47569.
- EX p::'q_47570 * 'q_47569. (x::'q_47570 * 'q_47569 => bool) p & u = p} =
-{u::'q_47570 * 'q_47569.
- EX (a::'q_47570) b::'q_47569. x (a, b) & u = (a, b)}"
-  by (import hollight SET_PAIR_THM)
-
-lemma FORALL_IN_GSPEC: "(ALL (P::'q_47618 => bool) f::'q_47618 => 'q_47739.
-    (ALL z::'q_47739.
-        z : {u::'q_47739. EX x::'q_47618. P x & u = f x} -->
-        (Q::'q_47739 => bool) z) =
-    (ALL x::'q_47618. P x --> Q (f x))) &
-(ALL (P::'q_47675 => 'q_47674 => bool) f::'q_47675 => 'q_47674 => 'q_47739.
-    (ALL z::'q_47739.
-        z : {u::'q_47739.
-             EX (x::'q_47675) y::'q_47674. P x y & u = f x y} -->
-        Q z) =
-    (ALL (x::'q_47675) y::'q_47674. P x y --> Q (f x y))) &
-(ALL (P::'q_47742 => 'q_47741 => 'q_47740 => bool)
-    f::'q_47742 => 'q_47741 => 'q_47740 => 'q_47739.
-    (ALL z::'q_47739.
-        z : {u::'q_47739.
-             EX (w::'q_47742) (x::'q_47741) y::'q_47740.
-                P w x y & u = f w x y} -->
-        Q z) =
-    (ALL (w::'q_47742) (x::'q_47741) y::'q_47740. P w x y --> Q (f w x y)))"
-  by (import hollight FORALL_IN_GSPEC)
-
-lemma EXISTS_IN_GSPEC: "(ALL (P::'q_47788 => bool) f::'q_47788 => 'q_47909.
-    (EX z::'q_47909.
-        z : {u::'q_47909. EX x::'q_47788. P x & u = f x} &
-        (Q::'q_47909 => bool) z) =
-    (EX x::'q_47788. P x & Q (f x))) &
-(ALL (P::'q_47845 => 'q_47844 => bool) f::'q_47845 => 'q_47844 => 'q_47909.
-    (EX z::'q_47909.
-        z : {u::'q_47909. EX (x::'q_47845) y::'q_47844. P x y & u = f x y} &
-        Q z) =
-    (EX (x::'q_47845) y::'q_47844. P x y & Q (f x y))) &
-(ALL (P::'q_47912 => 'q_47911 => 'q_47910 => bool)
-    f::'q_47912 => 'q_47911 => 'q_47910 => 'q_47909.
-    (EX z::'q_47909.
-        z : {u::'q_47909.
-             EX (w::'q_47912) (x::'q_47911) y::'q_47910.
-                P w x y & u = f w x y} &
-        Q z) =
-    (EX (w::'q_47912) (x::'q_47911) y::'q_47910. P w x y & Q (f w x y)))"
-  by (import hollight EXISTS_IN_GSPEC)
-
-lemma SET_PROVE_CASES: "(P::('A => bool) => bool) {} &
-(ALL (a::'A) s::'A => bool. a ~: s --> P (insert a s))
-==> P (x::'A => bool)"
-  by (import hollight SET_PROVE_CASES)
-
-lemma UNIONS_IMAGE: "Union ((f::'q_47989 => 'q_47973 => bool) ` (s::'q_47989 => bool)) =
-{u::'q_47973. EX y::'q_47973. (EX x::'q_47989. x : s & y : f x) & u = y}"
-  by (import hollight UNIONS_IMAGE)
-
-lemma INTERS_IMAGE: "Inter ((f::'q_48032 => 'q_48016 => bool) ` (s::'q_48032 => bool)) =
-{u::'q_48016. EX y::'q_48016. (ALL x::'q_48032. x : s --> y : f x) & u = y}"
-  by (import hollight INTERS_IMAGE)
-
-lemma UNIONS_GSPEC: "(ALL (P::'q_48085 => bool) f::'q_48085 => 'q_48071 => bool.
-    Union {u::'q_48071 => bool. EX x::'q_48085. P x & u = f x} =
-    {u::'q_48071.
-     EX a::'q_48071. (EX x::'q_48085. P x & a : f x) & u = a}) &
-(ALL (P::'q_48149 => 'q_48148 => bool)
-    f::'q_48149 => 'q_48148 => 'q_48129 => bool.
-    Union
-     {u::'q_48129 => bool.
-      EX (x::'q_48149) y::'q_48148. P x y & u = f x y} =
-    {u::'q_48129.
-     EX a::'q_48129.
-        (EX (x::'q_48149) y::'q_48148. P x y & a : f x y) & u = a}) &
-(ALL (P::'q_48223 => 'q_48222 => 'q_48221 => bool)
-    f::'q_48223 => 'q_48222 => 'q_48221 => 'q_48197 => bool.
-    Union
-     {u::'q_48197 => bool.
-      EX (x::'q_48223) (y::'q_48222) z::'q_48221. P x y z & u = f x y z} =
-    {u::'q_48197.
-     EX a::'q_48197.
-        (EX (x::'q_48223) (y::'q_48222) z::'q_48221.
-            P x y z & a : f x y z) &
-        u = a})"
-  by (import hollight UNIONS_GSPEC)
-
-lemma INTERS_GSPEC: "(ALL (P::'q_48276 => bool) f::'q_48276 => 'q_48262 => bool.
-    Inter {u::'q_48262 => bool. EX x::'q_48276. P x & u = f x} =
-    {u::'q_48262.
-     EX a::'q_48262. (ALL x::'q_48276. P x --> a : f x) & u = a}) &
-(ALL (P::'q_48340 => 'q_48339 => bool)
-    f::'q_48340 => 'q_48339 => 'q_48320 => bool.
-    Inter
-     {u::'q_48320 => bool.
-      EX (x::'q_48340) y::'q_48339. P x y & u = f x y} =
-    {u::'q_48320.
-     EX a::'q_48320.
-        (ALL (x::'q_48340) y::'q_48339. P x y --> a : f x y) & u = a}) &
-(ALL (P::'q_48414 => 'q_48413 => 'q_48412 => bool)
-    f::'q_48414 => 'q_48413 => 'q_48412 => 'q_48388 => bool.
-    Inter
-     {u::'q_48388 => bool.
-      EX (x::'q_48414) (y::'q_48413) z::'q_48412. P x y z & u = f x y z} =
-    {u::'q_48388.
-     EX a::'q_48388.
-        (ALL (x::'q_48414) (y::'q_48413) z::'q_48412.
-            P x y z --> a : f x y z) &
-        u = a})"
-  by (import hollight INTERS_GSPEC)
-
-lemma DIFF_INTERS: "(x::'q_48451 => bool) - Inter (xa::('q_48451 => bool) => bool) =
-Union {u::'q_48451 => bool. EX xb::'q_48451 => bool. xb : xa & u = x - xb}"
-  by (import hollight DIFF_INTERS)
-
-lemma INTERS_UNIONS: "Inter (x::('q_48486 => bool) => bool) =
-UNIV -
-Union {u::'q_48486 => bool. EX t::'q_48486 => bool. t : x & u = UNIV - t}"
-  by (import hollight INTERS_UNIONS)
-
-lemma UNIONS_INTERS: "Union (s::('q_48521 => bool) => bool) =
-UNIV -
-Inter {u::'q_48521 => bool. EX t::'q_48521 => bool. t : s & u = UNIV - t}"
-  by (import hollight UNIONS_INTERS)
-
-lemma FINITE_SING: "finite {x::'q_48799}"
-  by (import hollight FINITE_SING)
-
-lemma FINITE_DELETE_IMP: "finite (s::'A => bool) ==> finite (s - {x::'A})"
-  by (import hollight FINITE_DELETE_IMP)
-
-lemma FINITE_DELETE: "finite ((s::'A => bool) - {x::'A}) = finite s"
-  by (import hollight FINITE_DELETE)
-
-lemma FINITE_FINITE_UNIONS: "finite (s::('q_48871 => bool) => bool)
-==> finite (Union s) = (ALL t::'q_48871 => bool. t : s --> finite t)"
-  by (import hollight FINITE_FINITE_UNIONS)
-
-lemma FINITE_IMAGE_EXPAND: "finite (s::'A => bool)
-==> finite
-     {u::'B. EX y::'B. (EX x::'A. x : s & y = (f::'A => 'B) x) & u = y}"
-  by (import hollight FINITE_IMAGE_EXPAND)
-
-lemma FINITE_IMAGE_INJ_GENERAL: "(ALL (x::'A) y::'A.
-    x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y --> x = y) &
-finite (x::'B => bool)
-==> finite {u::'A. EX xa::'A. (xa : s & f xa : x) & u = xa}"
-  by (import hollight FINITE_IMAGE_INJ_GENERAL)
-
-lemma FINITE_FINITE_PREIMAGE_GENERAL: "finite (t::'B => bool) &
-(ALL y::'B.
-    y : t -->
-    finite
-     {u::'A. EX x::'A. (x : (s::'A => bool) & (f::'A => 'B) x = y) & u = x})
-==> finite {u::'A. EX x::'A. (x : s & f x : t) & u = x}"
-  by (import hollight FINITE_FINITE_PREIMAGE_GENERAL)
-
-lemma FINITE_FINITE_PREIMAGE: "finite (t::'B => bool) &
-(ALL y::'B. y : t --> finite {u::'A. EX x::'A. (f::'A => 'B) x = y & u = x})
-==> finite {u::'A. EX x::'A. f x : t & u = x}"
-  by (import hollight FINITE_FINITE_PREIMAGE)
-
-lemma FINITE_IMAGE_INJ_EQ: "(!!(x::'A) y::'A.
-    x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y ==> x = y)
-==> finite (f ` s) = finite s"
-  by (import hollight FINITE_IMAGE_INJ_EQ)
-
-lemma FINITE_IMAGE_INJ: "(ALL (x::'A) y::'A. (f::'A => 'B) x = f y --> x = y) &
-finite (A::'B => bool)
-==> finite {u::'A. EX x::'A. f x : A & u = x}"
-  by (import hollight FINITE_IMAGE_INJ)
-
-lemma INFINITE_IMAGE_INJ: "[| !!(x::'A) y::'A. (f::'A => 'B) x = f y ==> x = y;
-   infinite (s::'A => bool) |]
-==> infinite (f ` s)"
-  by (import hollight INFINITE_IMAGE_INJ)
-
-lemma FINITE_SUBSET_IMAGE: "(finite (t::'B => bool) & t <= (f::'A => 'B) ` (s::'A => bool)) =
-(EX x::'A => bool. finite x & x <= s & t = f ` x)"
-  by (import hollight FINITE_SUBSET_IMAGE)
-
-lemma EXISTS_FINITE_SUBSET_IMAGE: "(EX xc::'q_49755 => bool.
-    finite xc &
-    xc <= (xa::'q_49735 => 'q_49755) ` (xb::'q_49735 => bool) &
-    (x::('q_49755 => bool) => bool) xc) =
-(EX xc::'q_49735 => bool. finite xc & xc <= xb & x (xa ` xc))"
-  by (import hollight EXISTS_FINITE_SUBSET_IMAGE)
-
-lemma FINITE_SUBSET_IMAGE_IMP: "finite (t::'B => bool) & t <= (f::'A => 'B) ` (s::'A => bool)
-==> EX s'::'A => bool. finite s' & s' <= s & t <= f ` s'"
-  by (import hollight FINITE_SUBSET_IMAGE_IMP)
-
-definition
-  FINREC :: "('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool"  where
-  "FINREC ==
-SOME FINREC::('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool.
-   (ALL (f::'A => 'B => 'B) (s::'A => bool) (a::'B) b::'B.
-       FINREC f b s a (0::nat) = (s = {} & a = b)) &
-   (ALL (b::'B) (s::'A => bool) (n::nat) (a::'B) f::'A => 'B => 'B.
-       FINREC f b s a (Suc n) =
-       (EX (x::'A) c::'B. x : s & FINREC f b (s - {x}) c n & a = f x c))"
-
-lemma DEF_FINREC: "FINREC =
-(SOME FINREC::('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool.
-    (ALL (f::'A => 'B => 'B) (s::'A => bool) (a::'B) b::'B.
-        FINREC f b s a (0::nat) = (s = {} & a = b)) &
-    (ALL (b::'B) (s::'A => bool) (n::nat) (a::'B) f::'A => 'B => 'B.
-        FINREC f b s a (Suc n) =
-        (EX (x::'A) c::'B. x : s & FINREC f b (s - {x}) c n & a = f x c)))"
-  by (import hollight DEF_FINREC)
-
-lemma FINREC_1_LEMMA: "FINREC (x::'q_49919 => 'q_49918 => 'q_49918) (xa::'q_49918)
- (xb::'q_49919 => bool) (xc::'q_49918) (Suc (0::nat)) =
-(EX xd::'q_49919. xb = {xd} & xc = x xd xa)"
-  by (import hollight FINREC_1_LEMMA)
-
-lemma FINREC_SUC_LEMMA: "[| !!(x::'A) (y::'A) s::'B.
-      x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s);
-   FINREC f (b::'B) (s::'A => bool) (z::'B) (Suc (n::nat)); (x::'A) : s |]
-==> EX w::'B. FINREC f b (s - {x}) w n & z = f x w"
-  by (import hollight FINREC_SUC_LEMMA)
-
-lemma FINREC_UNIQUE_LEMMA: "[| !!(x::'A) (y::'A) s::'B.
-      x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s);
-   FINREC f (b::'B) (s::'A => bool) (a1::'B) (n1::nat) &
-   FINREC f b s (a2::'B) (n2::nat) |]
-==> a1 = a2 & n1 = n2"
-  by (import hollight FINREC_UNIQUE_LEMMA)
-
-lemma FINREC_EXISTS_LEMMA: "finite (s::'A => bool)
-==> EX a::'B. Ex (FINREC (f::'A => 'B => 'B) (b::'B) s a)"
-  by (import hollight FINREC_EXISTS_LEMMA)
-
-lemma FINREC_FUN_LEMMA: "(ALL s::'A.
-    (P::'A => bool) s -->
-    (EX a::'B. Ex ((R::'A => 'B => 'C => bool) s a))) &
-(ALL (n1::'C) (n2::'C) (s::'A) (a1::'B) a2::'B.
-    R s a1 n1 & R s a2 n2 --> a1 = a2 & n1 = n2)
-==> EX x::'A => 'B. ALL (s::'A) a::'B. P s --> Ex (R s a) = (x s = a)"
-  by (import hollight FINREC_FUN_LEMMA)
-
-lemma FINREC_FUN: "(!!(x::'A) (y::'A) s::'B.
-    x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> EX g::('A => bool) => 'B.
-       g {} = (b::'B) &
-       (ALL (s::'A => bool) x::'A.
-           finite s & x : s --> g s = f x (g (s - {x})))"
-  by (import hollight FINREC_FUN)
-
-lemma SET_RECURSION_LEMMA: "(!!(x::'A) (y::'A) s::'B.
-    x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> EX g::('A => bool) => 'B.
-       g {} = (b::'B) &
-       (ALL (x::'A) s::'A => bool.
-           finite s --> g (insert x s) = (if x : s then g s else f x (g s)))"
-  by (import hollight SET_RECURSION_LEMMA)
-
-definition
-  ITSET :: "('q_50575 => 'q_50574 => 'q_50574)
-=> ('q_50575 => bool) => 'q_50574 => 'q_50574"  where
-  "ITSET ==
-%(u::'q_50575 => 'q_50574 => 'q_50574) (ua::'q_50575 => bool) ub::'q_50574.
-   (SOME g::('q_50575 => bool) => 'q_50574.
-       g {} = ub &
-       (ALL (x::'q_50575) s::'q_50575 => bool.
-           finite s -->
-           g (insert x s) = (if x : s then g s else u x (g s))))
-    ua"
-
-lemma DEF_ITSET: "ITSET =
-(%(u::'q_50575 => 'q_50574 => 'q_50574) (ua::'q_50575 => bool) ub::'q_50574.
-    (SOME g::('q_50575 => bool) => 'q_50574.
-        g {} = ub &
-        (ALL (x::'q_50575) s::'q_50575 => bool.
-            finite s -->
-            g (insert x s) = (if x : s then g s else u x (g s))))
-     ua)"
-  by (import hollight DEF_ITSET)
-
-lemma FINITE_RECURSION: "(!!(x::'A) (y::'A) s::'B.
-    x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> ITSET f {} (b::'B) = b &
-    (ALL (x::'A) xa::'A => bool.
-        finite xa -->
-        ITSET f (insert x xa) b =
-        (if x : xa then ITSET f xa b else f x (ITSET f xa b)))"
-  by (import hollight FINITE_RECURSION)
-
-lemma FINITE_RECURSION_DELETE: "(!!(x::'A) (y::'A) s::'B.
-    x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> ITSET f {} (b::'B) = b &
-    (ALL (x::'A) s::'A => bool.
-        finite s -->
-        ITSET f s b =
-        (if x : s then f x (ITSET f (s - {x}) b) else ITSET f (s - {x}) b))"
-  by (import hollight FINITE_RECURSION_DELETE)
-
-lemma ITSET_EQ: "finite (x::'q_50880 => bool) &
-(ALL xc::'q_50880.
-    xc : x -->
-    (xa::'q_50880 => 'q_50881 => 'q_50881) xc =
-    (xb::'q_50880 => 'q_50881 => 'q_50881) xc) &
-(ALL (x::'q_50880) (y::'q_50880) s::'q_50881.
-    x ~= y --> xa x (xa y s) = xa y (xa x s)) &
-(ALL (x::'q_50880) (y::'q_50880) s::'q_50881.
-    x ~= y --> xb x (xb y s) = xb y (xb x s))
-==> ITSET xa x (xc::'q_50881) = ITSET xb x xc"
-  by (import hollight ITSET_EQ)
-
-lemma SUBSET_RESTRICT: "{u::'q_50914.
- EX xb::'q_50914.
-    (xb : (x::'q_50914 => bool) & (xa::'q_50914 => bool) xb) & u = xb}
-<= x"
-  by (import hollight SUBSET_RESTRICT)
-
-lemma FINITE_RESTRICT: "finite (s::'A => bool)
-==> finite {u::'A. EX x::'A. (x : s & (P::'A => bool) x) & u = x}"
-  by (import hollight FINITE_RESTRICT)
-
-definition
-  CARD :: "('q_50968 => bool) => nat"  where
-  "CARD == %u::'q_50968 => bool. ITSET (%x::'q_50968. Suc) u (0::nat)"
-
-lemma DEF_CARD: "CARD = (%u::'q_50968 => bool. ITSET (%x::'q_50968. Suc) u (0::nat))"
-  by (import hollight DEF_CARD)
-
-lemma CARD_CLAUSES: "CARD {} = (0::nat) &
-(ALL (x::'A::type) s::'A::type => bool.
-    finite s -->
-    CARD (insert x s) = (if x : s then CARD s else Suc (CARD s)))"
-  by (import hollight CARD_CLAUSES)
-
-lemma CARD_UNION: "finite (x::'A => bool) & finite (xa::'A => bool) & x Int xa = {}
-==> CARD (x Un xa) = CARD x + CARD xa"
-  by (import hollight CARD_UNION)
-
-lemma CARD_DELETE: "finite (s::'A => bool)
-==> CARD (s - {x::'A}) = (if x : s then CARD s - (1::nat) else CARD s)"
-  by (import hollight CARD_DELETE)
-
-lemma CARD_UNION_EQ: "finite (u::'q_51213 => bool) &
-(s::'q_51213 => bool) Int (t::'q_51213 => bool) = {} & s Un t = u
-==> CARD s + CARD t = CARD u"
-  by (import hollight CARD_UNION_EQ)
-
-lemma CARD_DIFF: "finite (s::'q_51270 => bool) & (t::'q_51270 => bool) <= s
-==> CARD (s - t) = CARD s - CARD t"
-  by (import hollight CARD_DIFF)
-
-lemma CARD_EQ_0: "finite (s::'q_51308 => bool) ==> (CARD s = (0::nat)) = (s = {})"
-  by (import hollight CARD_EQ_0)
-
-lemma FINITE_INDUCT_DELETE: "[| (P::('A => bool) => bool) {} &
-   (ALL s::'A => bool.
-       finite s & s ~= {} --> (EX x::'A. x : s & (P (s - {x}) --> P s)));
-   finite (s::'A => bool) |]
-==> P s"
-  by (import hollight FINITE_INDUCT_DELETE)
-
-definition
-  HAS_SIZE :: "('q_51427 => bool) => nat => bool"  where
-  "HAS_SIZE == %(u::'q_51427 => bool) ua::nat. finite u & CARD u = ua"
-
-lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_51427 => bool) ua::nat. finite u & CARD u = ua)"
-  by (import hollight DEF_HAS_SIZE)
-
-lemma HAS_SIZE_CARD: "HAS_SIZE (x::'q_51446 => bool) (xa::nat) ==> CARD x = xa"
-  by (import hollight HAS_SIZE_CARD)
-
-lemma HAS_SIZE_0: "HAS_SIZE (s::'A => bool) (0::nat) = (s = {})"
-  by (import hollight HAS_SIZE_0)
-
-lemma HAS_SIZE_SUC: "HAS_SIZE (s::'A => bool) (Suc (n::nat)) =
-(s ~= {} & (ALL x::'A. x : s --> HAS_SIZE (s - {x}) n))"
-  by (import hollight HAS_SIZE_SUC)
-
-lemma HAS_SIZE_UNION: "HAS_SIZE (x::'q_51584 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_51584 => bool) (xc::nat) & x Int xa = {}
-==> HAS_SIZE (x Un xa) (xb + xc)"
-  by (import hollight HAS_SIZE_UNION)
-
-lemma HAS_SIZE_DIFF: "HAS_SIZE (x::'q_51620 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_51620 => bool) (xc::nat) & xa <= x
-==> HAS_SIZE (x - xa) (xb - xc)"
-  by (import hollight HAS_SIZE_DIFF)
-
-lemma HAS_SIZE_UNIONS: "HAS_SIZE (x::'A => bool) (xb::nat) &
-(ALL xb::'A. xb : x --> HAS_SIZE ((xa::'A => 'B => bool) xb) (xc::nat)) &
-(ALL (xb::'A) y::'A. xb : x & y : x & xb ~= y --> xa xb Int xa y = {})
-==> HAS_SIZE (Union {u::'B => bool. EX xb::'A. xb : x & u = xa xb})
-     (xb * xc)"
-  by (import hollight HAS_SIZE_UNIONS)
-
-lemma FINITE_HAS_SIZE: "finite (x::'q_51824 => bool) = HAS_SIZE x (CARD x)"
-  by (import hollight FINITE_HAS_SIZE)
-
-lemma HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_51886 => bool) (0::nat) = (s = {}) &
-HAS_SIZE s (Suc (n::nat)) =
-(EX (a::'q_51886) t::'q_51886 => bool.
-    HAS_SIZE t n & a ~: t & s = insert a t)"
-  by (import hollight HAS_SIZE_CLAUSES)
-
-lemma CARD_SUBSET_EQ: "finite (b::'A => bool) & (a::'A => bool) <= b & CARD a = CARD b ==> a = b"
-  by (import hollight CARD_SUBSET_EQ)
-
-lemma CARD_SUBSET: "(a::'A => bool) <= (b::'A => bool) & finite b ==> CARD a <= CARD b"
-  by (import hollight CARD_SUBSET)
-
-lemma CARD_SUBSET_LE: "finite (b::'A => bool) & (a::'A => bool) <= b & CARD b <= CARD a ==> a = b"
-  by (import hollight CARD_SUBSET_LE)
-
-lemma SUBSET_CARD_EQ: "finite (t::'q_52197 => bool) & (s::'q_52197 => bool) <= t
-==> (CARD s = CARD t) = (s = t)"
-  by (import hollight SUBSET_CARD_EQ)
-
-lemma CARD_PSUBSET: "(a::'A => bool) < (b::'A => bool) & finite b ==> CARD a < CARD b"
-  by (import hollight CARD_PSUBSET)
-
-lemma CARD_UNION_LE: "finite (s::'A => bool) & finite (t::'A => bool)
-==> CARD (s Un t) <= CARD s + CARD t"
-  by (import hollight CARD_UNION_LE)
-
-lemma CARD_UNIONS_LE: "HAS_SIZE (x::'A => bool) (xb::nat) &
-(ALL xb::'A.
-    xb : x -->
-    finite ((xa::'A => 'B => bool) xb) & CARD (xa xb) <= (xc::nat))
-==> CARD (Union {u::'B => bool. EX xb::'A. xb : x & u = xa xb}) <= xb * xc"
-  by (import hollight CARD_UNIONS_LE)
-
-lemma CARD_UNION_GEN: "finite (s::'q_52620 => bool) & finite (t::'q_52620 => bool)
-==> CARD (s Un t) = CARD s + CARD t - CARD (s Int t)"
-  by (import hollight CARD_UNION_GEN)
-
-lemma CARD_UNION_OVERLAP_EQ: "finite (s::'q_52701 => bool) & finite (t::'q_52701 => bool)
-==> (CARD (s Un t) = CARD s + CARD t) = (s Int t = {})"
-  by (import hollight CARD_UNION_OVERLAP_EQ)
-
-lemma CARD_UNION_OVERLAP: "finite (x::'q_52743 => bool) &
-finite (xa::'q_52743 => bool) & CARD (x Un xa) < CARD x + CARD xa
-==> x Int xa ~= {}"
-  by (import hollight CARD_UNION_OVERLAP)
-
-lemma CARD_IMAGE_INJ: "(ALL (xa::'A) y::'A.
-    xa : (x::'A => bool) & y : x & (f::'A => 'B) xa = f y --> xa = y) &
-finite x
-==> CARD (f ` x) = CARD x"
-  by (import hollight CARD_IMAGE_INJ)
-
-lemma HAS_SIZE_IMAGE_INJ: "(ALL (xb::'A) y::'A.
-    xb : (xa::'A => bool) & y : xa & (x::'A => 'B) xb = x y --> xb = y) &
-HAS_SIZE xa (xb::nat)
-==> HAS_SIZE (x ` xa) xb"
-  by (import hollight HAS_SIZE_IMAGE_INJ)
-
-lemma CARD_IMAGE_LE: "finite (s::'A => bool) ==> CARD ((f::'A => 'B) ` s) <= CARD s"
-  by (import hollight CARD_IMAGE_LE)
-
-lemma CARD_IMAGE_INJ_EQ: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> (f::'A => 'B) x : (t::'B => bool)) &
-(ALL y::'B. y : t --> (EX! x::'A. x : s & f x = y))
-==> CARD t = CARD s"
-  by (import hollight CARD_IMAGE_INJ_EQ)
-
-lemma CARD_SUBSET_IMAGE: "finite (t::'q_52977 => bool) &
-(s::'q_52984 => bool) <= (f::'q_52977 => 'q_52984) ` t
-==> CARD s <= CARD t"
-  by (import hollight CARD_SUBSET_IMAGE)
-
-lemma HAS_SIZE_IMAGE_INJ_EQ: "(!!(x::'q_53049) y::'q_53049.
-    x : (s::'q_53049 => bool) & y : s & (f::'q_53049 => 'q_53044) x = f y
-    ==> x = y)
-==> HAS_SIZE (f ` s) (n::nat) = HAS_SIZE s n"
-  by (import hollight HAS_SIZE_IMAGE_INJ_EQ)
-
-lemma CHOOSE_SUBSET_STRONG: "(finite (s::'A => bool) ==> (n::nat) <= CARD s) ==> EX t<=s. HAS_SIZE t n"
-  by (import hollight CHOOSE_SUBSET_STRONG)
-
-lemma CHOOSE_SUBSET: "[| finite (s::'A => bool); (n::nat) <= CARD s |] ==> EX t<=s. HAS_SIZE t n"
-  by (import hollight CHOOSE_SUBSET)
-
-lemma HAS_SIZE_PRODUCT_DEPENDENT: "HAS_SIZE (x::'A => bool) (xa::nat) &
-(ALL xa::'A. xa : x --> HAS_SIZE ((xb::'A => 'B => bool) xa) (xc::nat))
-==> HAS_SIZE
-     {u::'A * 'B. EX (xa::'A) y::'B. (xa : x & y : xb xa) & u = (xa, y)}
-     (xa * xc)"
-  by (import hollight HAS_SIZE_PRODUCT_DEPENDENT)
-
-lemma FINITE_PRODUCT_DEPENDENT: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> finite ((t::'A => 'B => bool) x))
-==> finite
-     {u::'C.
-      EX (x::'A) y::'B. (x : s & y : t x) & u = (f::'A => 'B => 'C) x y}"
-  by (import hollight FINITE_PRODUCT_DEPENDENT)
-
-lemma FINITE_PRODUCT: "finite (x::'A => bool) & finite (xa::'B => bool)
-==> finite {u::'A * 'B. EX (xb::'A) y::'B. (xb : x & y : xa) & u = (xb, y)}"
-  by (import hollight FINITE_PRODUCT)
-
-lemma CARD_PRODUCT: "finite (s::'A => bool) & finite (t::'B => bool)
-==> CARD {u::'A * 'B. EX (x::'A) y::'B. (x : s & y : t) & u = (x, y)} =
-    CARD s * CARD t"
-  by (import hollight CARD_PRODUCT)
-
-lemma HAS_SIZE_PRODUCT: "HAS_SIZE (x::'A => bool) (xa::nat) & HAS_SIZE (xb::'B => bool) (xc::nat)
-==> HAS_SIZE
-     {u::'A * 'B. EX (xa::'A) y::'B. (xa : x & y : xb) & u = (xa, y)}
-     (xa * xc)"
-  by (import hollight HAS_SIZE_PRODUCT)
-
-definition
-  CROSS :: "('q_53759 => bool) => ('q_53758 => bool) => 'q_53759 * 'q_53758 => bool"  where
-  "CROSS ==
-%(u::'q_53759 => bool) ua::'q_53758 => bool.
-   {ub::'q_53759 * 'q_53758.
-    EX (x::'q_53759) y::'q_53758. (x : u & y : ua) & ub = (x, y)}"
-
-lemma DEF_CROSS: "CROSS =
-(%(u::'q_53759 => bool) ua::'q_53758 => bool.
-    {ub::'q_53759 * 'q_53758.
-     EX (x::'q_53759) y::'q_53758. (x : u & y : ua) & ub = (x, y)})"
-  by (import hollight DEF_CROSS)
-
-lemma IN_CROSS: "((x::'q_53795, xa::'q_53798)
- : CROSS (xb::'q_53795 => bool) (xc::'q_53798 => bool)) =
-(x : xb & xa : xc)"
-  by (import hollight IN_CROSS)
-
-lemma HAS_SIZE_CROSS: "HAS_SIZE (x::'q_53823 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_53826 => bool) (xc::nat)
-==> HAS_SIZE (CROSS x xa) (xb * xc)"
-  by (import hollight HAS_SIZE_CROSS)
-
-lemma FINITE_CROSS: "finite (x::'q_53851 => bool) & finite (xa::'q_53853 => bool)
-==> finite (CROSS x xa)"
-  by (import hollight FINITE_CROSS)
-
-lemma CARD_CROSS: "finite (x::'q_53874 => bool) & finite (xa::'q_53876 => bool)
-==> CARD (CROSS x xa) = CARD x * CARD xa"
-  by (import hollight CARD_CROSS)
-
-lemma CROSS_EQ_EMPTY: "(CROSS (x::'q_53917 => bool) (xa::'q_53921 => bool) = {}) =
-(x = {} | xa = {})"
-  by (import hollight CROSS_EQ_EMPTY)
-
-lemma HAS_SIZE_FUNSPACE: "HAS_SIZE (s::'A => bool) (m::nat) & HAS_SIZE (t::'B => bool) (n::nat)
-==> HAS_SIZE
-     {u::'A => 'B.
-      EX f::'A => 'B.
-         ((ALL x::'A. x : s --> f x : t) &
-          (ALL x::'A. x ~: s --> f x = (d::'B))) &
-         u = f}
-     (n ^ m)"
-  by (import hollight HAS_SIZE_FUNSPACE)
-
-lemma CARD_FUNSPACE: "finite (s::'q_54227 => bool) & finite (t::'q_54224 => bool)
-==> CARD
-     {u::'q_54227 => 'q_54224.
-      EX f::'q_54227 => 'q_54224.
-         ((ALL x::'q_54227. x : s --> f x : t) &
-          (ALL x::'q_54227. x ~: s --> f x = (d::'q_54224))) &
-         u = f} =
-    CARD t ^ CARD s"
-  by (import hollight CARD_FUNSPACE)
-
-lemma FINITE_FUNSPACE: "finite (s::'q_54293 => bool) & finite (t::'q_54290 => bool)
-==> finite
-     {u::'q_54293 => 'q_54290.
-      EX f::'q_54293 => 'q_54290.
-         ((ALL x::'q_54293. x : s --> f x : t) &
-          (ALL x::'q_54293. x ~: s --> f x = (d::'q_54290))) &
-         u = f}"
-  by (import hollight FINITE_FUNSPACE)
-
-lemma HAS_SIZE_POWERSET: "HAS_SIZE (s::'A => bool) (n::nat)
-==> HAS_SIZE {u::'A => bool. EX t<=s. u = t} ((2::nat) ^ n)"
-  by (import hollight HAS_SIZE_POWERSET)
-
-lemma CARD_POWERSET: "finite (s::'A => bool)
-==> CARD {u::'A => bool. EX t<=s. u = t} = (2::nat) ^ CARD s"
-  by (import hollight CARD_POWERSET)
-
-lemma FINITE_POWERSET: "finite (s::'A => bool) ==> finite {u::'A => bool. EX t<=s. u = t}"
-  by (import hollight FINITE_POWERSET)
-
-lemma FINITE_UNIONS: "finite (Union (s::('A => bool) => bool)) =
-(finite s & (ALL t::'A => bool. t : s --> finite t))"
-  by (import hollight FINITE_UNIONS)
-
-lemma POWERSET_CLAUSES: "{x::'q_54515 => bool. EX xa<={}. x = xa} = {{}} &
-(ALL (x::'A) xa::'A => bool.
-    {xb::'A => bool. EX xc<=insert x xa. xb = xc} =
-    {u::'A => bool. EX s<=xa. u = s} Un
-    insert x ` {u::'A => bool. EX s<=xa. u = s})"
-  by (import hollight POWERSET_CLAUSES)
-
-lemma HAS_SIZE_NUMSEG_LT: "HAS_SIZE {u. EX m<n. u = m} n"
-  by (import hollight HAS_SIZE_NUMSEG_LT)
-
-lemma CARD_NUMSEG_LT: "CARD {u. EX m<x. u = m} = x"
-  by (import hollight CARD_NUMSEG_LT)
-
-lemma FINITE_NUMSEG_LT: "finite {u::nat. EX m<x::nat. u = m}"
-  by (import hollight FINITE_NUMSEG_LT)
-
-lemma HAS_SIZE_NUMSEG_LE: "HAS_SIZE {xa. EX xb<=x. xa = xb} (x + 1)"
-  by (import hollight HAS_SIZE_NUMSEG_LE)
-
-lemma FINITE_NUMSEG_LE: "finite {u::nat. EX m<=x::nat. u = m}"
-  by (import hollight FINITE_NUMSEG_LE)
-
-lemma CARD_NUMSEG_LE: "CARD {u. EX m<=x. u = m} = x + 1"
-  by (import hollight CARD_NUMSEG_LE)
-
-lemma num_FINITE: "finite (s::nat => bool) = (EX a::nat. ALL x::nat. x : s --> x <= a)"
-  by (import hollight num_FINITE)
-
-lemma num_FINITE_AVOID: "finite (s::nat => bool) ==> EX a::nat. a ~: s"
-  by (import hollight num_FINITE_AVOID)
-
-lemma FINITE_REAL_INTERVAL: "(ALL a. infinite {u. EX x. real_lt a x & u = x}) &
-(ALL a. infinite {u. EX x. real_le a x & u = x}) &
-(ALL b. infinite {u. EX x. real_lt x b & u = x}) &
-(ALL b. infinite {u. EX x. real_le x b & u = x}) &
-(ALL x xa.
-    finite {u. EX xb. (real_lt x xb & real_lt xb xa) & u = xb} =
-    real_le xa x) &
-(ALL a b.
-    finite {u. EX x. (real_le a x & real_lt x b) & u = x} = real_le b a) &
-(ALL a b.
-    finite {u. EX x. (real_lt a x & real_le x b) & u = x} = real_le b a) &
-(ALL a b.
-    finite {u. EX x. (real_le a x & real_le x b) & u = x} = real_le b a)"
-  by (import hollight FINITE_REAL_INTERVAL)
-
-lemma real_INFINITE: "(infinite::(hollight.real => bool) => bool) (UNIV::hollight.real => bool)"
-  by (import hollight real_INFINITE)
-
-lemma HAS_SIZE_INDEX: "HAS_SIZE (x::'A => bool) (n::nat)
-==> EX f::nat => 'A.
-       (ALL m<n. f m : x) &
-       (ALL xa::'A. xa : x --> (EX! m::nat. m < n & f m = xa))"
-  by (import hollight HAS_SIZE_INDEX)
-
-definition
-  pairwise :: "('q_55938 => 'q_55938 => bool) => ('q_55938 => bool) => bool"  where
-  "pairwise ==
-%(u::'q_55938 => 'q_55938 => bool) ua::'q_55938 => bool.
-   ALL (x::'q_55938) y::'q_55938. x : ua & y : ua & x ~= y --> u x y"
-
-lemma DEF_pairwise: "pairwise =
-(%(u::'q_55938 => 'q_55938 => bool) ua::'q_55938 => bool.
-    ALL (x::'q_55938) y::'q_55938. x : ua & y : ua & x ~= y --> u x y)"
-  by (import hollight DEF_pairwise)
-
-definition
-  PAIRWISE :: "('A => 'A => bool) => 'A list => bool"  where
-  "PAIRWISE ==
-SOME PAIRWISE::('A => 'A => bool) => 'A list => bool.
-   (ALL r::'A => 'A => bool. PAIRWISE r [] = True) &
-   (ALL (h::'A) (r::'A => 'A => bool) t::'A list.
-       PAIRWISE r (h # t) = (list_all (r h) t & PAIRWISE r t))"
-
-lemma DEF_PAIRWISE: "PAIRWISE =
-(SOME PAIRWISE::('A => 'A => bool) => 'A list => bool.
-    (ALL r::'A => 'A => bool. PAIRWISE r [] = True) &
-    (ALL (h::'A) (r::'A => 'A => bool) t::'A list.
-        PAIRWISE r (h # t) = (list_all (r h) t & PAIRWISE r t)))"
-  by (import hollight DEF_PAIRWISE)
-
-lemma PAIRWISE_EMPTY: "pairwise (x::'q_55973 => 'q_55973 => bool) {} = True"
-  by (import hollight PAIRWISE_EMPTY)
-
-lemma PAIRWISE_SING: "pairwise (x::'q_55991 => 'q_55991 => bool) {xa::'q_55991} = True"
-  by (import hollight PAIRWISE_SING)
-
-lemma PAIRWISE_MONO: "pairwise (x::'q_56011 => 'q_56011 => bool) (xa::'q_56011 => bool) &
-(xb::'q_56011 => bool) <= xa
-==> pairwise x xb"
-  by (import hollight PAIRWISE_MONO)
-
-lemma SURJECTIVE_IFF_INJECTIVE_GEN: "finite (s::'A => bool) &
-finite (t::'B => bool) & CARD s = CARD t & (f::'A => 'B) ` s <= t
-==> (ALL y::'B. y : t --> (EX x::'A. x : s & f x = y)) =
-    (ALL (x::'A) y::'A. x : s & y : s & f x = f y --> x = y)"
-  by (import hollight SURJECTIVE_IFF_INJECTIVE_GEN)
-
-lemma SURJECTIVE_IFF_INJECTIVE: "finite (x::'A => bool) & (xa::'A => 'A) ` x <= x
-==> (ALL y::'A. y : x --> (EX xb::'A. xb : x & xa xb = y)) =
-    (ALL (xb::'A) y::'A. xb : x & y : x & xa xb = xa y --> xb = y)"
-  by (import hollight SURJECTIVE_IFF_INJECTIVE)
-
-lemma IMAGE_IMP_INJECTIVE_GEN: "[| finite (s::'A => bool) &
-   CARD s = CARD (t::'B => bool) & (f::'A => 'B) ` s = t;
-   (x::'A) : s & (y::'A) : s & f x = f y |]
-==> x = y"
-  by (import hollight IMAGE_IMP_INJECTIVE_GEN)
-
-lemma IMAGE_IMP_INJECTIVE: "[| finite (s::'q_56387 => bool) & (f::'q_56387 => 'q_56387) ` s = s;
-   (x::'q_56387) : s & (y::'q_56387) : s & f x = f y |]
-==> x = y"
-  by (import hollight IMAGE_IMP_INJECTIVE)
-
-lemma CARD_LE_INJ: "finite (x::'A => bool) & finite (xa::'B => bool) & CARD x <= CARD xa
-==> EX f::'A => 'B.
-       f ` x <= xa &
-       (ALL (xa::'A) y::'A. xa : x & y : x & f xa = f y --> xa = y)"
-  by (import hollight CARD_LE_INJ)
-
-lemma FORALL_IN_CLAUSES: "(ALL x::'q_56493 => bool. (ALL xa::'q_56493. xa : {} --> x xa) = True) &
-(ALL (x::'q_56533 => bool) (xa::'q_56533) xb::'q_56533 => bool.
-    (ALL xc::'q_56533. xc : insert xa xb --> x xc) =
-    (x xa & (ALL xa::'q_56533. xa : xb --> x xa)))"
-  by (import hollight FORALL_IN_CLAUSES)
-
-lemma EXISTS_IN_CLAUSES: "(ALL x::'q_56553 => bool. (EX xa::'q_56553. xa : {} & x xa) = False) &
-(ALL (x::'q_56593 => bool) (xa::'q_56593) xb::'q_56593 => bool.
-    (EX xc::'q_56593. xc : insert xa xb & x xc) =
-    (x xa | (EX xa::'q_56593. xa : xb & x xa)))"
-  by (import hollight EXISTS_IN_CLAUSES)
-
-lemma SURJECTIVE_ON_RIGHT_INVERSE: "(ALL xb::'q_56650.
-    xb : (xa::'q_56650 => bool) -->
-    (EX xa::'q_56649.
-        xa : (s::'q_56649 => bool) & (x::'q_56649 => 'q_56650) xa = xb)) =
-(EX g::'q_56650 => 'q_56649.
-    ALL y::'q_56650. y : xa --> g y : s & x (g y) = y)"
-  by (import hollight SURJECTIVE_ON_RIGHT_INVERSE)
-
-lemma INJECTIVE_ON_LEFT_INVERSE: "(ALL (xb::'q_56743) y::'q_56743.
-    xb : (xa::'q_56743 => bool) &
-    y : xa & (x::'q_56743 => 'q_56746) xb = x y -->
-    xb = y) =
-(EX xb::'q_56746 => 'q_56743. ALL xc::'q_56743. xc : xa --> xb (x xc) = xc)"
-  by (import hollight INJECTIVE_ON_LEFT_INVERSE)
-
-lemma BIJECTIVE_ON_LEFT_RIGHT_INVERSE: "(!!x::'q_56878.
-    x : (s::'q_56878 => bool)
-    ==> (f::'q_56878 => 'q_56877) x : (t::'q_56877 => bool))
-==> ((ALL (x::'q_56878) y::'q_56878. x : s & y : s & f x = f y --> x = y) &
-     (ALL x::'q_56877. x : t --> (EX xa::'q_56878. xa : s & f xa = x))) =
-    (EX g::'q_56877 => 'q_56878.
-        (ALL y::'q_56877. y : t --> g y : s) &
-        (ALL y::'q_56877. y : t --> f (g y) = y) &
-        (ALL x::'q_56878. x : s --> g (f x) = x))"
-  by (import hollight BIJECTIVE_ON_LEFT_RIGHT_INVERSE)
-
-lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_56902. EX x::'q_56905. (f::'q_56905 => 'q_56902) x = y) =
-(EX g::'q_56902 => 'q_56905. ALL y::'q_56902. f (g y) = y)"
-  by (import hollight SURJECTIVE_RIGHT_INVERSE)
-
-lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_56939) xa::'q_56939.
-    (f::'q_56939 => 'q_56942) x = f xa --> x = xa) =
-(EX g::'q_56942 => 'q_56939. ALL x::'q_56939. g (f x) = x)"
-  by (import hollight INJECTIVE_LEFT_INVERSE)
-
-lemma BIJECTIVE_LEFT_RIGHT_INVERSE: "((ALL (x::'A) y::'A. (f::'A => 'B) x = f y --> x = y) &
- (ALL y::'B. EX x::'A. f x = y)) =
-(EX g::'B => 'A. (ALL y::'B. f (g y) = y) & (ALL x::'A. g (f x) = x))"
-  by (import hollight BIJECTIVE_LEFT_RIGHT_INVERSE)
-
-lemma FUNCTION_FACTORS_RIGHT: "(ALL xb::'q_57046.
-    EX y::'q_57034.
-       (xa::'q_57034 => 'q_57047) y = (x::'q_57046 => 'q_57047) xb) =
-(EX xb::'q_57046 => 'q_57034. x = xa o xb)"
-  by (import hollight FUNCTION_FACTORS_RIGHT)
-
-lemma FUNCTION_FACTORS_LEFT: "(ALL (xb::'q_57119) y::'q_57119.
-    (xa::'q_57119 => 'q_57099) xb = xa y -->
-    (x::'q_57119 => 'q_57120) xb = x y) =
-(EX xb::'q_57099 => 'q_57120. x = xb o xa)"
-  by (import hollight FUNCTION_FACTORS_LEFT)
-
-lemma SURJECTIVE_FORALL_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) =
-(ALL P::'B => bool. (ALL x::'A. P (f x)) = All P)"
-  by (import hollight SURJECTIVE_FORALL_THM)
-
-lemma SURJECTIVE_EXISTS_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) =
-(ALL P::'B => bool. (EX x::'A. P (f x)) = Ex P)"
-  by (import hollight SURJECTIVE_EXISTS_THM)
-
-lemma SURJECTIVE_IMAGE_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) =
-(ALL x::'B => bool.
-    f ` {u::'A. EX xa::'A. x (f xa) & u = xa} =
-    {u::'B. EX xa::'B. x xa & u = xa})"
-  by (import hollight SURJECTIVE_IMAGE_THM)
-
-lemma IMAGE_INJECTIVE_IMAGE_OF_SUBSET: "EX x<=s::'A => bool.
-   (f::'A => 'B) ` s = f ` x &
-   (ALL (xa::'A) y::'A. xa : x & y : x & f xa = f y --> xa = y)"
-  by (import hollight IMAGE_INJECTIVE_IMAGE_OF_SUBSET)
-
-lemma INJECTIVE_ON_IMAGE: "(ALL (s::'A => bool) t::'A => bool.
-    s <= (u::'A => bool) & t <= u & (f::'A => 'B) ` s = f ` t --> s = t) =
-(ALL (x::'A) y::'A. x : u & y : u & f x = f y --> x = y)"
-  by (import hollight INJECTIVE_ON_IMAGE)
-
-lemma INJECTIVE_IMAGE: "(ALL (s::'A => bool) t::'A => bool. (f::'A => 'B) ` s = f ` t --> s = t) =
-(ALL (x::'A) y::'A. f x = f y --> x = y)"
-  by (import hollight INJECTIVE_IMAGE)
-
-lemma SURJECTIVE_ON_IMAGE: "(ALL t<=v::'B => bool. EX s<=u::'A => bool. (f::'A => 'B) ` s = t) =
-(ALL y::'B. y : v --> (EX x::'A. x : u & f x = y))"
-  by (import hollight SURJECTIVE_ON_IMAGE)
-
-lemma SURJECTIVE_IMAGE: "(ALL t::'B => bool. EX s::'A => bool. (f::'A => 'B) ` s = t) =
-(ALL y::'B. EX x::'A. f x = y)"
-  by (import hollight SURJECTIVE_IMAGE)
-
-lemma CARD_EQ_BIJECTION: "finite (s::'A => bool) & finite (t::'B => bool) & CARD s = CARD t
-==> EX f::'A => 'B.
-       (ALL x::'A. x : s --> f x : t) &
-       (ALL y::'B. y : t --> (EX x::'A. x : s & f x = y)) &
-       (ALL (x::'A) y::'A. x : s & y : s & f x = f y --> x = y)"
-  by (import hollight CARD_EQ_BIJECTION)
-
-lemma CARD_EQ_BIJECTIONS: "finite (s::'A => bool) & finite (t::'B => bool) & CARD s = CARD t
-==> EX (f::'A => 'B) g::'B => 'A.
-       (ALL x::'A. x : s --> f x : t & g (f x) = x) &
-       (ALL y::'B. y : t --> g y : s & f (g y) = y)"
-  by (import hollight CARD_EQ_BIJECTIONS)
-
-lemma BIJECTIONS_HAS_SIZE: "(ALL x::'A.
-    x : (s::'A => bool) -->
-    (f::'A => 'B) x : (t::'B => bool) & (g::'B => 'A) (f x) = x) &
-(ALL y::'B. y : t --> g y : s & f (g y) = y) & HAS_SIZE s (n::nat)
-==> HAS_SIZE t n"
-  by (import hollight BIJECTIONS_HAS_SIZE)
-
-lemma BIJECTIONS_HAS_SIZE_EQ: "(ALL x::'A.
-    x : (s::'A => bool) -->
-    (f::'A => 'B) x : (t::'B => bool) & (g::'B => 'A) (f x) = x) &
-(ALL y::'B. y : t --> g y : s & f (g y) = y)
-==> HAS_SIZE s (n::nat) = HAS_SIZE t n"
-  by (import hollight BIJECTIONS_HAS_SIZE_EQ)
-
-lemma BIJECTIONS_CARD_EQ: "(finite (s::'A => bool) | finite (t::'B => bool)) &
-(ALL x::'A. x : s --> (f::'A => 'B) x : t & (g::'B => 'A) (f x) = x) &
-(ALL y::'B. y : t --> g y : s & f (g y) = y)
-==> CARD s = CARD t"
-  by (import hollight BIJECTIONS_CARD_EQ)
-
-lemma WF_FINITE: "(ALL x::'A. ~ (u_556::'A => 'A => bool) x x) &
-(ALL (x::'A) (y::'A) z::'A. u_556 x y & u_556 y z --> u_556 x z) &
-(ALL x::'A. finite {u::'A. EX y::'A. u_556 y x & u = y})
-==> wfP u_556"
-  by (import hollight WF_FINITE)
-
-consts
-  "<=_c" :: "('q_58200 => bool) => ('q_58195 => bool) => bool" ("<='_c")
-
-defs
-  "<=_c_def": "<=_c ==
-%(u::'q_58200 => bool) ua::'q_58195 => bool.
-   EX f::'q_58200 => 'q_58195.
-      (ALL x::'q_58200. x : u --> f x : ua) &
-      (ALL (x::'q_58200) y::'q_58200. x : u & y : u & f x = f y --> x = y)"
-
-lemma DEF__lessthan__equal__c: "<=_c =
-(%(u::'q_58200 => bool) ua::'q_58195 => bool.
-    EX f::'q_58200 => 'q_58195.
-       (ALL x::'q_58200. x : u --> f x : ua) &
-       (ALL (x::'q_58200) y::'q_58200. x : u & y : u & f x = f y --> x = y))"
-  by (import hollight DEF__lessthan__equal__c)
-
-consts
-  "<_c" :: "('q_58212 => bool) => ('q_58213 => bool) => bool" ("<'_c")
-
-defs
-  "<_c_def": "<_c == %(u::'q_58212 => bool) ua::'q_58213 => bool. <=_c u ua & ~ <=_c ua u"
-
-lemma DEF__lessthan__c: "<_c = (%(u::'q_58212 => bool) ua::'q_58213 => bool. <=_c u ua & ~ <=_c ua u)"
-  by (import hollight DEF__lessthan__c)
-
-consts
-  "=_c" :: "('q_58264 => bool) => ('q_58261 => bool) => bool" ("='_c")
-
-defs
-  "=_c_def": "=_c ==
-%(u::'q_58264 => bool) ua::'q_58261 => bool.
-   EX f::'q_58264 => 'q_58261.
-      (ALL x::'q_58264. x : u --> f x : ua) &
-      (ALL y::'q_58261. y : ua --> (EX! x::'q_58264. x : u & f x = y))"
-
-lemma DEF__equal__c: "=_c =
-(%(u::'q_58264 => bool) ua::'q_58261 => bool.
-    EX f::'q_58264 => 'q_58261.
-       (ALL x::'q_58264. x : u --> f x : ua) &
-       (ALL y::'q_58261. y : ua --> (EX! x::'q_58264. x : u & f x = y)))"
-  by (import hollight DEF__equal__c)
-
-consts
-  ">=_c" :: "('q_58273 => bool) => ('q_58272 => bool) => bool" (">='_c")
-
-defs
-  ">=_c_def": ">=_c == %(u::'q_58273 => bool) ua::'q_58272 => bool. <=_c ua u"
-
-lemma DEF__greaterthan__equal__c: ">=_c = (%(u::'q_58273 => bool) ua::'q_58272 => bool. <=_c ua u)"
-  by (import hollight DEF__greaterthan__equal__c)
-
-consts
-  ">_c" :: "('q_58282 => bool) => ('q_58281 => bool) => bool" (">'_c")
-
-defs
-  ">_c_def": ">_c == %(u::'q_58282 => bool) ua::'q_58281 => bool. <_c ua u"
-
-lemma DEF__greaterthan__c: ">_c = (%(u::'q_58282 => bool) ua::'q_58281 => bool. <_c ua u)"
-  by (import hollight DEF__greaterthan__c)
-
-lemma LE_C: "<=_c (x::'q_58320 => bool) (xa::'q_58323 => bool) =
-(EX xb::'q_58323 => 'q_58320.
-    ALL xc::'q_58320. xc : x --> (EX x::'q_58323. x : xa & xb x = xc))"
-  by (import hollight LE_C)
-
-lemma GE_C: ">=_c (x::'q_58364 => bool) (xa::'q_58361 => bool) =
-(EX f::'q_58364 => 'q_58361.
-    ALL y::'q_58361. y : xa --> (EX xa::'q_58364. xa : x & y = f xa))"
-  by (import hollight GE_C)
-
-definition
-  COUNTABLE :: "('q_58372 => bool) => bool"  where
-  "(op ==::(('q_58372::type => bool) => bool)
-        => (('q_58372::type => bool) => bool) => prop)
- (COUNTABLE::('q_58372::type => bool) => bool)
- ((>=_c::(nat => bool) => ('q_58372::type => bool) => bool)
-   (UNIV::nat => bool))"
-
-lemma DEF_COUNTABLE: "(op =::(('q_58372::type => bool) => bool)
-       => (('q_58372::type => bool) => bool) => bool)
- (COUNTABLE::('q_58372::type => bool) => bool)
- ((>=_c::(nat => bool) => ('q_58372::type => bool) => bool)
-   (UNIV::nat => bool))"
-  by (import hollight DEF_COUNTABLE)
-
-lemma NUMSEG_COMBINE_R: "(x::nat) <= (xa::nat) + (1::nat) & xa <= (xb::nat)
-==> {x..xa} Un {xa + (1::nat)..xb} = {x..xb}"
-  by (import hollight NUMSEG_COMBINE_R)
-
-lemma NUMSEG_COMBINE_L: "(x::nat) <= (xa::nat) & xa <= (xb::nat) + (1::nat)
-==> {x..xa - (1::nat)} Un {xa..xb} = {x..xb}"
-  by (import hollight NUMSEG_COMBINE_L)
-
-lemma NUMSEG_LREC: "(x::nat) <= (xa::nat) ==> insert x {x + (1::nat)..xa} = {x..xa}"
-  by (import hollight NUMSEG_LREC)
-
-lemma NUMSEG_RREC: "(x::nat) <= (xa::nat) ==> insert xa {x..xa - (1::nat)} = {x..xa}"
-  by (import hollight NUMSEG_RREC)
-
-lemma IN_NUMSEG_0: "((x::nat) : {0::nat..xa::nat}) = (x <= xa)"
-  by (import hollight IN_NUMSEG_0)
-
-lemma NUMSEG_EMPTY: "({x::nat..xa::nat} = {}) = (xa < x)"
-  by (import hollight NUMSEG_EMPTY)
-
-lemma CARD_NUMSEG_LEMMA: "CARD {m..m + d} = d + 1"
-  by (import hollight CARD_NUMSEG_LEMMA)
-
-lemma CARD_NUMSEG: "CARD {m..n} = n + 1 - m"
-  by (import hollight CARD_NUMSEG)
-
-lemma HAS_SIZE_NUMSEG: "HAS_SIZE {x..xa} (xa + 1 - x)"
-  by (import hollight HAS_SIZE_NUMSEG)
-
-lemma CARD_NUMSEG_1: "CARD {1..x} = x"
-  by (import hollight CARD_NUMSEG_1)
-
-lemma HAS_SIZE_NUMSEG_1: "HAS_SIZE {1..x} x"
-  by (import hollight HAS_SIZE_NUMSEG_1)
-
-lemma NUMSEG_CLAUSES: "(ALL m::nat. {m..0::nat} = (if m = (0::nat) then {0::nat} else {})) &
-(ALL (m::nat) n::nat.
-    {m..Suc n} = (if m <= Suc n then insert (Suc n) {m..n} else {m..n}))"
-  by (import hollight NUMSEG_CLAUSES)
-
-lemma FINITE_INDEX_NUMSEG: "finite (s::'A => bool) =
-(EX f::nat => 'A.
-    (ALL (i::nat) j::nat.
-        i : {1::nat..CARD s} & j : {1::nat..CARD s} & f i = f j --> i = j) &
-    s = f ` {1::nat..CARD s})"
-  by (import hollight FINITE_INDEX_NUMSEG)
-
-lemma FINITE_INDEX_NUMBERS: "finite (s::'A => bool) =
-(EX (k::nat => bool) f::nat => 'A.
-    (ALL (i::nat) j::nat. i : k & j : k & f i = f j --> i = j) &
-    finite k & s = f ` k)"
-  by (import hollight FINITE_INDEX_NUMBERS)
-
-lemma DISJOINT_NUMSEG: "({x::nat..xa::nat} Int {xb::nat..xc::nat} = {}) =
-(xa < xb | xc < x | xa < x | xc < xb)"
-  by (import hollight DISJOINT_NUMSEG)
-
-lemma NUMSEG_ADD_SPLIT: "(x::nat) <= (xa::nat) + (1::nat)
-==> {x..xa + (xb::nat)} = {x..xa} Un {xa + (1::nat)..xa + xb}"
-  by (import hollight NUMSEG_ADD_SPLIT)
-
-lemma SUBSET_NUMSEG: "({m::nat..n::nat} <= {p::nat..q::nat}) = (n < m | p <= m & n <= q)"
-  by (import hollight SUBSET_NUMSEG)
-
-lemma NUMSEG_LE: "{u::nat. EX xa<=x::nat. u = xa} = {0::nat..x}"
-  by (import hollight NUMSEG_LE)
-
-lemma NUMSEG_LT: "{u::nat. EX x<n::nat. u = x} =
-(if n = (0::nat) then {} else {0::nat..n - (1::nat)})"
-  by (import hollight NUMSEG_LT)
-
-lemma TOPOLOGICAL_SORT: "[| (ALL (x::'A) y::'A.
-       (u_556::'A => 'A => bool) x y & u_556 y x --> x = y) &
-   (ALL (x::'A) (y::'A) z::'A. u_556 x y & u_556 y z --> u_556 x z);
-   HAS_SIZE (s::'A => bool) (n::nat) |]
-==> EX f::nat => 'A.
-       s = f ` {1::nat..n} &
-       (ALL (j::nat) k::nat.
-           j : {1::nat..n} & k : {1::nat..n} & j < k -->
-           ~ u_556 (f k) (f j))"
-  by (import hollight TOPOLOGICAL_SORT)
-
-lemma FINITE_INTSEG: "(ALL l r. finite {u. EX x. (int_le l x & int_le x r) & u = x}) &
-(ALL l r. finite {u. EX x. (int_le l x & int_lt x r) & u = x}) &
-(ALL l r. finite {u. EX x. (int_lt l x & int_le x r) & u = x}) &
-(ALL l r. finite {u. EX x. (int_lt l x & int_lt x r) & u = x})"
-  by (import hollight FINITE_INTSEG)
-
-definition
-  neutral :: "('q_59899 => 'q_59899 => 'q_59899) => 'q_59899"  where
-  "neutral ==
-%u::'q_59899 => 'q_59899 => 'q_59899.
-   SOME x::'q_59899. ALL y::'q_59899. u x y = y & u y x = y"
-
-lemma DEF_neutral: "neutral =
-(%u::'q_59899 => 'q_59899 => 'q_59899.
-    SOME x::'q_59899. ALL y::'q_59899. u x y = y & u y x = y)"
-  by (import hollight DEF_neutral)
-
-definition
-  monoidal :: "('A => 'A => 'A) => bool"  where
-  "monoidal ==
-%u::'A => 'A => 'A.
-   (ALL (x::'A) y::'A. u x y = u y x) &
-   (ALL (x::'A) (y::'A) z::'A. u x (u y z) = u (u x y) z) &
-   (ALL x::'A. u (neutral u) x = x)"
-
-lemma DEF_monoidal: "monoidal =
-(%u::'A => 'A => 'A.
-    (ALL (x::'A) y::'A. u x y = u y x) &
-    (ALL (x::'A) (y::'A) z::'A. u x (u y z) = u (u x y) z) &
-    (ALL x::'A. u (neutral u) x = x))"
-  by (import hollight DEF_monoidal)
-
-lemma MONOIDAL_AC: "monoidal (x::'q_60055 => 'q_60055 => 'q_60055)
-==> (ALL a::'q_60055. x (neutral x) a = a) &
-    (ALL a::'q_60055. x a (neutral x) = a) &
-    (ALL (a::'q_60055) b::'q_60055. x a b = x b a) &
-    (ALL (a::'q_60055) (b::'q_60055) c::'q_60055.
-        x (x a b) c = x a (x b c)) &
-    (ALL (a::'q_60055) (b::'q_60055) c::'q_60055. x a (x b c) = x b (x a c))"
-  by (import hollight MONOIDAL_AC)
-
-definition
-  support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool"  where
-  "support ==
-%(u::'B => 'B => 'B) (ua::'A => 'B) ub::'A => bool.
-   {uc::'A. EX x::'A. (x : ub & ua x ~= neutral u) & uc = x}"
-
-lemma DEF_support: "support =
-(%(u::'B => 'B => 'B) (ua::'A => 'B) ub::'A => bool.
-    {uc::'A. EX x::'A. (x : ub & ua x ~= neutral u) & uc = x})"
-  by (import hollight DEF_support)
-
-definition
-  iterate :: "('q_60113 => 'q_60113 => 'q_60113)
-=> ('A => bool) => ('A => 'q_60113) => 'q_60113"  where
-  "iterate ==
-%(u::'q_60113 => 'q_60113 => 'q_60113) (ua::'A => bool) ub::'A => 'q_60113.
-   if finite (support u ub ua)
-   then ITSET (%x::'A. u (ub x)) (support u ub ua) (neutral u)
-   else neutral u"
-
-lemma DEF_iterate: "iterate =
-(%(u::'q_60113 => 'q_60113 => 'q_60113) (ua::'A => bool) ub::'A => 'q_60113.
-    if finite (support u ub ua)
-    then ITSET (%x::'A. u (ub x)) (support u ub ua) (neutral u)
-    else neutral u)"
-  by (import hollight DEF_iterate)
-
-lemma IN_SUPPORT: "((xb::'q_60163)
- : support (x::'q_60160 => 'q_60160 => 'q_60160) (xa::'q_60163 => 'q_60160)
-    (xc::'q_60163 => bool)) =
-(xb : xc & xa xb ~= neutral x)"
-  by (import hollight IN_SUPPORT)
-
-lemma SUPPORT_SUPPORT: "support (x::'q_60185 => 'q_60185 => 'q_60185) (xa::'q_60196 => 'q_60185)
- (support x xa (xb::'q_60196 => bool)) =
-support x xa xb"
-  by (import hollight SUPPORT_SUPPORT)
-
-lemma SUPPORT_EMPTY: "(ALL xc::'q_60235.
-    xc : (xb::'q_60235 => bool) -->
-    (xa::'q_60235 => 'q_60221) xc =
-    neutral (x::'q_60221 => 'q_60221 => 'q_60221)) =
-(support x xa xb = {})"
-  by (import hollight SUPPORT_EMPTY)
-
-lemma SUPPORT_SUBSET: "support (x::'q_60255 => 'q_60255 => 'q_60255) (xa::'q_60256 => 'q_60255)
- (xb::'q_60256 => bool)
-<= xb"
-  by (import hollight SUPPORT_SUBSET)
-
-lemma FINITE_SUPPORT: "finite (s::'q_60273 => bool)
-==> finite
-     (support (u::'q_60279 => 'q_60279 => 'q_60279)
-       (f::'q_60273 => 'q_60279) s)"
-  by (import hollight FINITE_SUPPORT)
-
-lemma SUPPORT_CLAUSES: "(ALL x::'q_60297 => 'q_60530.
-    support (u_4371::'q_60530 => 'q_60530 => 'q_60530) x {} = {}) &
-(ALL (x::'q_60345 => 'q_60530) (xa::'q_60345) xb::'q_60345 => bool.
-    support u_4371 x (insert xa xb) =
-    (if x xa = neutral u_4371 then support u_4371 x xb
-     else insert xa (support u_4371 x xb))) &
-(ALL (x::'q_60378 => 'q_60530) (xa::'q_60378) xb::'q_60378 => bool.
-    support u_4371 x (xb - {xa}) = support u_4371 x xb - {xa}) &
-(ALL (x::'q_60416 => 'q_60530) (xa::'q_60416 => bool) xb::'q_60416 => bool.
-    support u_4371 x (xa Un xb) =
-    support u_4371 x xa Un support u_4371 x xb) &
-(ALL (x::'q_60454 => 'q_60530) (xa::'q_60454 => bool) xb::'q_60454 => bool.
-    support u_4371 x (xa Int xb) =
-    support u_4371 x xa Int support u_4371 x xb) &
-(ALL (x::'q_60492 => 'q_60530) (xa::'q_60492 => bool) xb::'q_60492 => bool.
-    support u_4371 x (xa - xb) =
-    support u_4371 x xa - support u_4371 x xb) &
-(ALL (x::'q_60529 => 'q_60520) (xa::'q_60520 => 'q_60530)
-    xb::'q_60529 => bool.
-    support u_4371 xa (x ` xb) = x ` support u_4371 (xa o x) xb)"
-  by (import hollight SUPPORT_CLAUSES)
-
-lemma SUPPORT_DELTA: "support (x::'q_60556 => 'q_60556 => 'q_60556)
- (%xa::'q_60584.
-     if xa = (xc::'q_60584) then (xb::'q_60584 => 'q_60556) xa
-     else neutral x)
- (xa::'q_60584 => bool) =
-(if xc : xa then support x xb {xc} else {})"
-  by (import hollight SUPPORT_DELTA)
-
-lemma FINITE_SUPPORT_DELTA: "finite
- (support (x::'q_60605 => 'q_60605 => 'q_60605)
-   (%xc::'q_60614.
-       if xc = (xb::'q_60614) then (xa::'q_60614 => 'q_60605) xc
-       else neutral x)
-   (s::'q_60614 => bool))"
-  by (import hollight FINITE_SUPPORT_DELTA)
-
-lemma ITERATE_SUPPORT: "iterate (x::'q_60630 => 'q_60630 => 'q_60630)
- (support x (xa::'q_60642 => 'q_60630) (xb::'q_60642 => bool)) xa =
-iterate x xb xa"
-  by (import hollight ITERATE_SUPPORT)
-
-lemma ITERATE_EXPAND_CASES: "iterate (x::'q_60661 => 'q_60661 => 'q_60661) (xb::'q_60667 => bool)
- (xa::'q_60667 => 'q_60661) =
-(if finite (support x xa xb) then iterate x (support x xa xb) xa
- else neutral x)"
-  by (import hollight ITERATE_EXPAND_CASES)
-
-lemma ITERATE_CLAUSES_GEN: "monoidal (u_4371::'B => 'B => 'B)
-==> (ALL f::'A => 'B. iterate u_4371 {} f = neutral u_4371) &
-    (ALL (f::'A => 'B) (x::'A) s::'A => bool.
-        monoidal u_4371 & finite (support u_4371 f s) -->
-        iterate u_4371 (insert x s) f =
-        (if x : s then iterate u_4371 s f
-         else u_4371 (f x) (iterate u_4371 s f)))"
-  by (import hollight ITERATE_CLAUSES_GEN)
-
-lemma ITERATE_CLAUSES: "monoidal (x::'q_60857 => 'q_60857 => 'q_60857)
-==> (ALL f::'q_60815 => 'q_60857. iterate x {} f = neutral x) &
-    (ALL (f::'q_60859 => 'q_60857) (xa::'q_60859) s::'q_60859 => bool.
-        finite s -->
-        iterate x (insert xa s) f =
-        (if xa : s then iterate x s f else x (f xa) (iterate x s f)))"
-  by (import hollight ITERATE_CLAUSES)
-
-lemma ITERATE_UNION: "[| monoidal (u_4371::'q_60945 => 'q_60945 => 'q_60945);
-   finite (s::'q_60930 => bool) &
-   finite (x::'q_60930 => bool) & s Int x = {} |]
-==> iterate u_4371 (s Un x) (f::'q_60930 => 'q_60945) =
-    u_4371 (iterate u_4371 s f) (iterate u_4371 x f)"
-  by (import hollight ITERATE_UNION)
-
-lemma ITERATE_UNION_GEN: "[| monoidal (x::'B => 'B => 'B);
-   finite (support x (xa::'A => 'B) (xb::'A => bool)) &
-   finite (support x xa (xc::'A => bool)) &
-   support x xa xb Int support x xa xc = {} |]
-==> iterate x (xb Un xc) xa = x (iterate x xb xa) (iterate x xc xa)"
-  by (import hollight ITERATE_UNION_GEN)
-
-lemma ITERATE_DIFF: "[| monoidal (u::'q_61087 => 'q_61087 => 'q_61087);
-   finite (s::'q_61083 => bool) & (t::'q_61083 => bool) <= s |]
-==> u (iterate u (s - t) (f::'q_61083 => 'q_61087)) (iterate u t f) =
-    iterate u s f"
-  by (import hollight ITERATE_DIFF)
-
-lemma ITERATE_DIFF_GEN: "[| monoidal (x::'B => 'B => 'B);
-   finite (support x (xa::'A => 'B) (xb::'A => bool)) &
-   support x xa (xc::'A => bool) <= support x xa xb |]
-==> x (iterate x (xb - xc) xa) (iterate x xc xa) = iterate x xb xa"
-  by (import hollight ITERATE_DIFF_GEN)
-
-lemma ITERATE_INCL_EXCL: "[| monoidal (u_4371::'q_61316 => 'q_61316 => 'q_61316);
-   finite (s::'q_61298 => bool) & finite (t::'q_61298 => bool) |]
-==> u_4371 (iterate u_4371 s (f::'q_61298 => 'q_61316))
-     (iterate u_4371 t f) =
-    u_4371 (iterate u_4371 (s Un t) f) (iterate u_4371 (s Int t) f)"
-  by (import hollight ITERATE_INCL_EXCL)
-
-lemma ITERATE_CLOSED: "[| monoidal (u_4371::'B => 'B => 'B);
-   (P::'B => bool) (neutral u_4371) &
-   (ALL (x::'B) y::'B. P x & P y --> P (u_4371 x y));
-   !!x::'A.
-      x : (s::'A => bool) & (f::'A => 'B) x ~= neutral u_4371 ==> P (f x) |]
-==> P (iterate u_4371 s f)"
-  by (import hollight ITERATE_CLOSED)
-
-lemma ITERATE_RELATED: "[| monoidal (u_4371::'B => 'B => 'B);
-   (R::'B => 'B => bool) (neutral u_4371) (neutral u_4371) &
-   (ALL (x1::'B) (y1::'B) (x2::'B) y2::'B.
-       R x1 x2 & R y1 y2 --> R (u_4371 x1 y1) (u_4371 x2 y2));
-   finite (x::'A => bool) &
-   (ALL xa::'A. xa : x --> R ((f::'A => 'B) xa) ((g::'A => 'B) xa)) |]
-==> R (iterate u_4371 x f) (iterate u_4371 x g)"
-  by (import hollight ITERATE_RELATED)
-
-lemma ITERATE_EQ_NEUTRAL: "[| monoidal (u_4371::'B => 'B => 'B);
-   !!x::'A. x : (s::'A => bool) ==> (f::'A => 'B) x = neutral u_4371 |]
-==> iterate u_4371 s f = neutral u_4371"
-  by (import hollight ITERATE_EQ_NEUTRAL)
-
-lemma ITERATE_SING: "monoidal (x::'B => 'B => 'B) ==> iterate x {xa::'A} (f::'A => 'B) = f xa"
-  by (import hollight ITERATE_SING)
-
-lemma ITERATE_DELETE: "[| monoidal (u::'B => 'B => 'B); finite (s::'A => bool) & (a::'A) : s |]
-==> u ((f::'A => 'B) a) (iterate u (s - {a}) f) = iterate u s f"
-  by (import hollight ITERATE_DELETE)
-
-lemma ITERATE_DELTA: "monoidal (u_4371::'q_61672 => 'q_61672 => 'q_61672)
-==> iterate u_4371 (xb::'q_61691 => bool)
-     (%xb::'q_61691.
-         if xb = (xa::'q_61691) then (x::'q_61691 => 'q_61672) xb
-         else neutral u_4371) =
-    (if xa : xb then x xa else neutral u_4371)"
-  by (import hollight ITERATE_DELTA)
-
-lemma ITERATE_IMAGE: "[| monoidal (u_4371::'C => 'C => 'C);
-   !!(x::'A) y::'A.
-      x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y ==> x = y |]
-==> iterate u_4371 (f ` s) (g::'B => 'C) = iterate u_4371 s (g o f)"
-  by (import hollight ITERATE_IMAGE)
-
-lemma ITERATE_BIJECTION: "[| monoidal (u_4371::'B => 'B => 'B);
-   (ALL x::'A. x : (s::'A => bool) --> (p::'A => 'A) x : s) &
-   (ALL y::'A. y : s --> (EX! x::'A. x : s & p x = y)) |]
-==> iterate u_4371 s (f::'A => 'B) = iterate u_4371 s (f o p)"
-  by (import hollight ITERATE_BIJECTION)
-
-lemma ITERATE_ITERATE_PRODUCT: "[| monoidal (u_4371::'C => 'C => 'C);
-   finite (x::'A => bool) &
-   (ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i)) |]
-==> iterate u_4371 x
-     (%i::'A. iterate u_4371 (xa i) ((xb::'A => 'B => 'C) i)) =
-    iterate u_4371
-     {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)}
-     (SOME f::'A * 'B => 'C. ALL (i::'A) j::'B. f (i, j) = xb i j)"
-  by (import hollight ITERATE_ITERATE_PRODUCT)
-
-lemma ITERATE_EQ: "[| monoidal (u_4371::'B => 'B => 'B);
-   !!x::'A. x : (s::'A => bool) ==> (f::'A => 'B) x = (g::'A => 'B) x |]
-==> iterate u_4371 s f = iterate u_4371 s g"
-  by (import hollight ITERATE_EQ)
-
-lemma ITERATE_EQ_GENERAL: "[| monoidal (u_4371::'C => 'C => 'C);
-   (ALL y::'B.
-       y : (t::'B => bool) -->
-       (EX! x::'A. x : (s::'A => bool) & (h::'A => 'B) x = y)) &
-   (ALL x::'A. x : s --> h x : t & (g::'B => 'C) (h x) = (f::'A => 'C) x) |]
-==> iterate u_4371 s f = iterate u_4371 t g"
-  by (import hollight ITERATE_EQ_GENERAL)
-
-lemma ITERATE_EQ_GENERAL_INVERSES: "[| monoidal (u_4371::'C => 'C => 'C);
-   (ALL y::'B.
-       y : (t::'B => bool) -->
-       (k::'B => 'A) y : (s::'A => bool) & (h::'A => 'B) (k y) = y) &
-   (ALL x::'A.
-       x : s -->
-       h x : t & k (h x) = x & (g::'B => 'C) (h x) = (f::'A => 'C) x) |]
-==> iterate u_4371 s f = iterate u_4371 t g"
-  by (import hollight ITERATE_EQ_GENERAL_INVERSES)
-
-lemma ITERATE_INJECTION: "[| monoidal (u_4371::'B => 'B => 'B);
-   finite (s::'A => bool) &
-   (ALL x::'A. x : s --> (p::'A => 'A) x : s) &
-   (ALL (x::'A) y::'A. x : s & y : s & p x = p y --> x = y) |]
-==> iterate u_4371 s ((f::'A => 'B) o p) = iterate u_4371 s f"
-  by (import hollight ITERATE_INJECTION)
-
-lemma ITERATE_UNION_NONZERO: "[| monoidal (u_4371::'B => 'B => 'B);
-   finite (s::'A => bool) &
-   finite (t::'A => bool) &
-   (ALL x::'A. x : s Int t --> (f::'A => 'B) x = neutral u_4371) |]
-==> iterate u_4371 (s Un t) f =
-    u_4371 (iterate u_4371 s f) (iterate u_4371 t f)"
-  by (import hollight ITERATE_UNION_NONZERO)
-
-lemma ITERATE_OP: "[| monoidal (u_4371::'q_62649 => 'q_62649 => 'q_62649);
-   finite (s::'q_62648 => bool) |]
-==> iterate u_4371 s
-     (%x::'q_62648.
-         u_4371 ((f::'q_62648 => 'q_62649) x)
-          ((g::'q_62648 => 'q_62649) x)) =
-    u_4371 (iterate u_4371 s f) (iterate u_4371 s g)"
-  by (import hollight ITERATE_OP)
-
-lemma ITERATE_SUPERSET: "[| monoidal (u_4371::'B => 'B => 'B);
-   (u::'A => bool) <= (v::'A => bool) &
-   (ALL x::'A. x : v & x ~: u --> (f::'A => 'B) x = neutral u_4371) |]
-==> iterate u_4371 v f = iterate u_4371 u f"
-  by (import hollight ITERATE_SUPERSET)
-
-lemma ITERATE_IMAGE_NONZERO: "[| monoidal (u_4371::'C => 'C => 'C);
-   finite (x::'A => bool) &
-   (ALL (xa::'A) y::'A.
-       xa : x & y : x & xa ~= y & (f::'A => 'B) xa = f y -->
-       (g::'B => 'C) (f xa) = neutral u_4371) |]
-==> iterate u_4371 (f ` x) g = iterate u_4371 x (g o f)"
-  by (import hollight ITERATE_IMAGE_NONZERO)
-
-lemma ITERATE_CASES: "[| monoidal (u_4371::'B => 'B => 'B); finite (s::'A => bool) |]
-==> iterate u_4371 s
-     (%x::'A.
-         if (P::'A => bool) x then (f::'A => 'B) x else (g::'A => 'B) x) =
-    u_4371 (iterate u_4371 {u::'A. EX x::'A. (x : s & P x) & u = x} f)
-     (iterate u_4371 {u::'A. EX x::'A. (x : s & ~ P x) & u = x} g)"
-  by (import hollight ITERATE_CASES)
-
-lemma ITERATE_OP_GEN: "[| monoidal (u_4371::'B => 'B => 'B);
-   finite (support u_4371 (f::'A => 'B) (s::'A => bool)) &
-   finite (support u_4371 (g::'A => 'B) s) |]
-==> iterate u_4371 s (%x::'A. u_4371 (f x) (g x)) =
-    u_4371 (iterate u_4371 s f) (iterate u_4371 s g)"
-  by (import hollight ITERATE_OP_GEN)
-
-lemma ITERATE_CLAUSES_NUMSEG: "monoidal (x::'q_63246 => 'q_63246 => 'q_63246)
-==> (ALL xa::nat.
-        iterate x {xa..0::nat} (f::nat => 'q_63246) =
-        (if xa = (0::nat) then f (0::nat) else neutral x)) &
-    (ALL (xa::nat) xb::nat.
-        iterate x {xa..Suc xb} f =
-        (if xa <= Suc xb then x (iterate x {xa..xb} f) (f (Suc xb))
-         else iterate x {xa..xb} f))"
-  by (import hollight ITERATE_CLAUSES_NUMSEG)
-
-lemma ITERATE_PAIR: "monoidal (u_4371::'q_63421 => 'q_63421 => 'q_63421)
-==> iterate u_4371 {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)}
-     (f::nat => 'q_63421) =
-    iterate u_4371 {m..n}
-     (%i::nat. u_4371 (f ((2::nat) * i)) (f ((2::nat) * i + (1::nat))))"
-  by (import hollight ITERATE_PAIR)
-
-definition
-  nsum :: "('q_63439 => bool) => ('q_63439 => nat) => nat"  where
-  "(op ==::(('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-        => (('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-           => prop)
- (nsum::('q_63439::type => bool) => ('q_63439::type => nat) => nat)
- ((iterate::(nat => nat => nat)
-            => ('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-   (op +::nat => nat => nat))"
-
-lemma DEF_nsum: "(op =::(('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-       => (('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-          => bool)
- (nsum::('q_63439::type => bool) => ('q_63439::type => nat) => nat)
- ((iterate::(nat => nat => nat)
-            => ('q_63439::type => bool) => ('q_63439::type => nat) => nat)
-   (op +::nat => nat => nat))"
-  by (import hollight DEF_nsum)
-
-lemma NEUTRAL_ADD: "neutral op + = (0::nat)"
-  by (import hollight NEUTRAL_ADD)
-
-lemma NEUTRAL_MUL: "neutral op * = (1::nat)"
-  by (import hollight NEUTRAL_MUL)
-
-lemma MONOIDAL_ADD: "(monoidal::(nat => nat => nat) => bool) (op +::nat => nat => nat)"
-  by (import hollight MONOIDAL_ADD)
-
-lemma MONOIDAL_MUL: "(monoidal::(nat => nat => nat) => bool) (op *::nat => nat => nat)"
-  by (import hollight MONOIDAL_MUL)
-
-lemma NSUM_CLAUSES: "(ALL x::'q_63477 => nat. nsum {} x = (0::nat)) &
-(ALL (x::'q_63516) (xa::'q_63516 => nat) xb::'q_63516 => bool.
-    finite xb -->
-    nsum (insert x xb) xa =
-    (if x : xb then nsum xb xa else xa x + nsum xb xa))"
-  by (import hollight NSUM_CLAUSES)
-
-lemma NSUM_UNION: "finite (xa::'q_63542 => bool) &
-finite (xb::'q_63542 => bool) & xa Int xb = {}
-==> nsum (xa Un xb) (x::'q_63542 => nat) = nsum xa x + nsum xb x"
-  by (import hollight NSUM_UNION)
-
-lemma NSUM_DIFF: "finite (s::'q_63597 => bool) & (t::'q_63597 => bool) <= s
-==> nsum (s - t) (f::'q_63597 => nat) = nsum s f - nsum t f"
-  by (import hollight NSUM_DIFF)
-
-lemma NSUM_INCL_EXCL: "finite (x::'A => bool) & finite (xa::'A => bool)
-==> nsum x (xb::'A => nat) + nsum xa xb =
-    nsum (x Un xa) xb + nsum (x Int xa) xb"
-  by (import hollight NSUM_INCL_EXCL)
-
-lemma NSUM_SUPPORT: "nsum (support op + (x::'q_63686 => nat) (xa::'q_63686 => bool)) x =
-nsum xa x"
-  by (import hollight NSUM_SUPPORT)
-
-lemma NSUM_ADD: "finite (xb::'q_63720 => bool)
-==> nsum xb
-     (%xb::'q_63720. (x::'q_63720 => nat) xb + (xa::'q_63720 => nat) xb) =
-    nsum xb x + nsum xb xa"
-  by (import hollight NSUM_ADD)
-
-lemma NSUM_ADD_GEN: "finite
- {xa::'q_63807.
-  EX xc::'q_63807.
-     (xc : (xb::'q_63807 => bool) & (x::'q_63807 => nat) xc ~= (0::nat)) &
-     xa = xc} &
-finite
- {x::'q_63807.
-  EX xc::'q_63807.
-     (xc : xb & (xa::'q_63807 => nat) xc ~= (0::nat)) & x = xc}
-==> nsum xb (%xb::'q_63807. x xb + xa xb) = nsum xb x + nsum xb xa"
-  by (import hollight NSUM_ADD_GEN)
-
-lemma NSUM_EQ_0: "(!!xb::'A. xb : (xa::'A => bool) ==> (x::'A => nat) xb = (0::nat))
-==> nsum xa x = (0::nat)"
-  by (import hollight NSUM_EQ_0)
-
-lemma NSUM_0: "nsum (x::'A => bool) (%n::'A. 0::nat) = (0::nat)"
-  by (import hollight NSUM_0)
-
-lemma NSUM_LMUL: "nsum (s::'A => bool) (%x::'A. (c::nat) * (f::'A => nat) x) = c * nsum s f"
-  by (import hollight NSUM_LMUL)
-
-lemma NSUM_RMUL: "nsum (xb::'A => bool) (%xb::'A. (x::'A => nat) xb * (xa::nat)) =
-nsum xb x * xa"
-  by (import hollight NSUM_RMUL)
-
-lemma NSUM_LE: "finite (xb::'q_63997 => bool) &
-(ALL xc::'q_63997.
-    xc : xb --> (x::'q_63997 => nat) xc <= (xa::'q_63997 => nat) xc)
-==> nsum xb x <= nsum xb xa"
-  by (import hollight NSUM_LE)
-
-lemma NSUM_LT: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> (f::'A => nat) x <= (g::'A => nat) x) &
-(EX x::'A. x : s & f x < g x)
-==> nsum s f < nsum s g"
-  by (import hollight NSUM_LT)
-
-lemma NSUM_LT_ALL: "finite (s::'q_64119 => bool) &
-s ~= {} &
-(ALL x::'q_64119. x : s --> (f::'q_64119 => nat) x < (g::'q_64119 => nat) x)
-==> nsum s f < nsum s g"
-  by (import hollight NSUM_LT_ALL)
-
-lemma NSUM_EQ: "(!!xc::'q_64157.
-    xc : (xb::'q_64157 => bool)
-    ==> (x::'q_64157 => nat) xc = (xa::'q_64157 => nat) xc)
-==> nsum xb x = nsum xb xa"
-  by (import hollight NSUM_EQ)
-
-lemma NSUM_CONST: "finite (s::'q_64187 => bool) ==> nsum s (%n::'q_64187. c::nat) = CARD s * c"
-  by (import hollight NSUM_CONST)
-
-lemma NSUM_POS_BOUND: "[| finite (x::'A => bool) & nsum x (f::'A => nat) <= (b::nat);
-   (xa::'A) : x |]
-==> f xa <= b"
-  by (import hollight NSUM_POS_BOUND)
-
-lemma NSUM_EQ_0_IFF: "finite (s::'q_64296 => bool)
-==> (nsum s (f::'q_64296 => nat) = (0::nat)) =
-    (ALL x::'q_64296. x : s --> f x = (0::nat))"
-  by (import hollight NSUM_EQ_0_IFF)
-
-lemma NSUM_DELETE: "finite (xa::'q_64325 => bool) & (xb::'q_64325) : xa
-==> (x::'q_64325 => nat) xb + nsum (xa - {xb}) x = nsum xa x"
-  by (import hollight NSUM_DELETE)
-
-lemma NSUM_SING: "nsum {xa::'q_64354} (x::'q_64354 => nat) = x xa"
-  by (import hollight NSUM_SING)
-
-lemma NSUM_DELTA: "nsum (x::'A => bool) (%x::'A. if x = (xa::'A) then b::nat else (0::nat)) =
-(if xa : x then b else (0::nat))"
-  by (import hollight NSUM_DELTA)
-
-lemma NSUM_SWAP: "finite (x::'A => bool) & finite (xa::'B => bool)
-==> nsum x (%i::'A. nsum xa ((f::'A => 'B => nat) i)) =
-    nsum xa (%j::'B. nsum x (%i::'A. f i j))"
-  by (import hollight NSUM_SWAP)
-
-lemma NSUM_IMAGE: "(!!(xa::'q_64490) y::'q_64490.
-    xa : (xb::'q_64490 => bool) &
-    y : xb & (x::'q_64490 => 'q_64466) xa = x y
-    ==> xa = y)
-==> nsum (x ` xb) (xa::'q_64466 => nat) = nsum xb (xa o x)"
-  by (import hollight NSUM_IMAGE)
-
-lemma NSUM_SUPERSET: "(xa::'A => bool) <= (xb::'A => bool) &
-(ALL xc::'A. xc : xb & xc ~: xa --> (x::'A => nat) xc = (0::nat))
-==> nsum xb x = nsum xa x"
-  by (import hollight NSUM_SUPERSET)
-
-lemma NSUM_UNION_RZERO: "finite (u::'A => bool) &
-(ALL x::'A. x : (v::'A => bool) & x ~: u --> (f::'A => nat) x = (0::nat))
-==> nsum (u Un v) f = nsum u f"
-  by (import hollight NSUM_UNION_RZERO)
-
-lemma NSUM_UNION_LZERO: "finite (v::'A => bool) &
-(ALL x::'A. x : (u::'A => bool) & x ~: v --> (f::'A => nat) x = (0::nat))
-==> nsum (u Un v) f = nsum v f"
-  by (import hollight NSUM_UNION_LZERO)
-
-lemma NSUM_RESTRICT: "finite (s::'q_64681 => bool)
-==> nsum s
-     (%x::'q_64681. if x : s then (f::'q_64681 => nat) x else (0::nat)) =
-    nsum s f"
-  by (import hollight NSUM_RESTRICT)
-
-lemma NSUM_BOUND: "finite (x::'A => bool) &
-(ALL xc::'A. xc : x --> (xa::'A => nat) xc <= (xb::nat))
-==> nsum x xa <= CARD x * xb"
-  by (import hollight NSUM_BOUND)
-
-lemma NSUM_BOUND_GEN: "finite (x::'A => bool) &
-x ~= {} & (ALL xa::'A. xa : x --> (f::'A => nat) xa <= (b::nat) div CARD x)
-==> nsum x f <= b"
-  by (import hollight NSUM_BOUND_GEN)
-
-lemma NSUM_BOUND_LT: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> (f::'A => nat) x <= (b::nat)) &
-(EX x::'A. x : s & f x < b)
-==> nsum s f < CARD s * b"
-  by (import hollight NSUM_BOUND_LT)
-
-lemma NSUM_BOUND_LT_ALL: "finite (s::'q_64899 => bool) &
-s ~= {} & (ALL x::'q_64899. x : s --> (f::'q_64899 => nat) x < (b::nat))
-==> nsum s f < CARD s * b"
-  by (import hollight NSUM_BOUND_LT_ALL)
-
-lemma NSUM_BOUND_LT_GEN: "finite (s::'A => bool) &
-s ~= {} & (ALL x::'A. x : s --> (f::'A => nat) x < (b::nat) div CARD s)
-==> nsum s f < b"
-  by (import hollight NSUM_BOUND_LT_GEN)
-
-lemma NSUM_UNION_EQ: "finite (u::'q_65000 => bool) &
-(s::'q_65000 => bool) Int (t::'q_65000 => bool) = {} & s Un t = u
-==> nsum s (f::'q_65000 => nat) + nsum t f = nsum u f"
-  by (import hollight NSUM_UNION_EQ)
-
-lemma NSUM_EQ_SUPERSET: "finite (t::'A => bool) &
-t <= (s::'A => bool) &
-(ALL x::'A. x : t --> (f::'A => nat) x = (g::'A => nat) x) &
-(ALL x::'A. x : s & x ~: t --> f x = (0::nat))
-==> nsum s f = nsum t g"
-  by (import hollight NSUM_EQ_SUPERSET)
-
-lemma NSUM_RESTRICT_SET: "nsum
- {u::'A. EX xb::'A. (xb : (xa::'A => bool) & (x::'A => bool) xb) & u = xb}
- (xb::'A => nat) =
-nsum xa (%xa::'A. if x xa then xb xa else (0::nat))"
-  by (import hollight NSUM_RESTRICT_SET)
-
-lemma NSUM_NSUM_RESTRICT: "finite (s::'q_65257 => bool) & finite (t::'q_65256 => bool)
-==> nsum s
-     (%x::'q_65257.
-         nsum
-          {u::'q_65256.
-           EX y::'q_65256.
-              (y : t & (R::'q_65257 => 'q_65256 => bool) x y) & u = y}
-          ((f::'q_65257 => 'q_65256 => nat) x)) =
-    nsum t
-     (%y::'q_65256.
-         nsum {u::'q_65257. EX x::'q_65257. (x : s & R x y) & u = x}
-          (%x::'q_65257. f x y))"
-  by (import hollight NSUM_NSUM_RESTRICT)
-
-lemma CARD_EQ_NSUM: "finite (x::'q_65276 => bool) ==> CARD x = nsum x (%x::'q_65276. 1::nat)"
-  by (import hollight CARD_EQ_NSUM)
-
-lemma NSUM_MULTICOUNT_GEN: "finite (s::'A => bool) &
-finite (t::'B => bool) &
-(ALL j::'B.
-    j : t -->
-    CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} =
-    (k::'B => nat) j)
-==> nsum s (%i::'A. CARD {u::'B. EX j::'B. (j : t & R i j) & u = j}) =
-    nsum t k"
-  by (import hollight NSUM_MULTICOUNT_GEN)
-
-lemma NSUM_MULTICOUNT: "finite (s::'A => bool) &
-finite (t::'B => bool) &
-(ALL j::'B.
-    j : t -->
-    CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} =
-    (k::nat))
-==> nsum s (%i::'A. CARD {u::'B. EX j::'B. (j : t & R i j) & u = j}) =
-    k * CARD t"
-  by (import hollight NSUM_MULTICOUNT)
-
-lemma NSUM_IMAGE_GEN: "finite (s::'A => bool)
-==> nsum s (g::'A => nat) =
-    nsum ((f::'A => 'B) ` s)
-     (%y::'B. nsum {u::'A. EX x::'A. (x : s & f x = y) & u = x} g)"
-  by (import hollight NSUM_IMAGE_GEN)
-
-lemma NSUM_GROUP: "finite (s::'A => bool) & (f::'A => 'B) ` s <= (t::'B => bool)
-==> nsum t
-     (%y::'B.
-         nsum {u::'A. EX x::'A. (x : s & f x = y) & u = x} (g::'A => nat)) =
-    nsum s g"
-  by (import hollight NSUM_GROUP)
-
-lemma NSUM_SUBSET: "finite (u::'A => bool) &
-finite (v::'A => bool) &
-(ALL x::'A. x : u - v --> (f::'A => nat) x = (0::nat))
-==> nsum u f <= nsum v f"
-  by (import hollight NSUM_SUBSET)
-
-lemma NSUM_SUBSET_SIMPLE: "finite (v::'q_65804 => bool) & (u::'q_65804 => bool) <= v
-==> nsum u (f::'q_65804 => nat) <= nsum v f"
-  by (import hollight NSUM_SUBSET_SIMPLE)
-
-lemma NSUM_IMAGE_NONZERO: "finite (xb::'A => bool) &
-(ALL (xc::'A) xd::'A.
-    xc : xb & xd : xb & xc ~= xd & (xa::'A => 'B) xc = xa xd -->
-    (x::'B => nat) (xa xc) = (0::nat))
-==> nsum (xa ` xb) x = nsum xb (x o xa)"
-  by (import hollight NSUM_IMAGE_NONZERO)
-
-lemma NSUM_BIJECTION: "(ALL x::'A. x : (xb::'A => bool) --> (xa::'A => 'A) x : xb) &
-(ALL y::'A. y : xb --> (EX! x::'A. x : xb & xa x = y))
-==> nsum xb (x::'A => nat) = nsum xb (x o xa)"
-  by (import hollight NSUM_BIJECTION)
-
-lemma NSUM_NSUM_PRODUCT: "finite (x::'A => bool) &
-(ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i))
-==> nsum x (%x::'A. nsum (xa x) ((xb::'A => 'B => nat) x)) =
-    nsum {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)}
-     (SOME f::'A * 'B => nat. ALL (i::'A) j::'B. f (i, j) = xb i j)"
-  by (import hollight NSUM_NSUM_PRODUCT)
-
-lemma NSUM_EQ_GENERAL: "(ALL y::'B.
-    y : (xa::'B => bool) -->
-    (EX! xa::'A. xa : (x::'A => bool) & (xd::'A => 'B) xa = y)) &
-(ALL xe::'A.
-    xe : x --> xd xe : xa & (xc::'B => nat) (xd xe) = (xb::'A => nat) xe)
-==> nsum x xb = nsum xa xc"
-  by (import hollight NSUM_EQ_GENERAL)
-
-lemma NSUM_EQ_GENERAL_INVERSES: "(ALL y::'B.
-    y : (xa::'B => bool) -->
-    (xe::'B => 'A) y : (x::'A => bool) & (xd::'A => 'B) (xe y) = y) &
-(ALL xf::'A.
-    xf : x -->
-    xd xf : xa &
-    xe (xd xf) = xf & (xc::'B => nat) (xd xf) = (xb::'A => nat) xf)
-==> nsum x xb = nsum xa xc"
-  by (import hollight NSUM_EQ_GENERAL_INVERSES)
-
-lemma NSUM_INJECTION: "finite (xb::'q_66274 => bool) &
-(ALL x::'q_66274. x : xb --> (xa::'q_66274 => 'q_66274) x : xb) &
-(ALL (x::'q_66274) y::'q_66274. x : xb & y : xb & xa x = xa y --> x = y)
-==> nsum xb ((x::'q_66274 => nat) o xa) = nsum xb x"
-  by (import hollight NSUM_INJECTION)
-
-lemma NSUM_UNION_NONZERO: "finite (xa::'q_66317 => bool) &
-finite (xb::'q_66317 => bool) &
-(ALL xc::'q_66317. xc : xa Int xb --> (x::'q_66317 => nat) xc = (0::nat))
-==> nsum (xa Un xb) x = nsum xa x + nsum xb x"
-  by (import hollight NSUM_UNION_NONZERO)
-
-lemma NSUM_UNIONS_NONZERO: "finite (x::('A => bool) => bool) &
-(ALL t::'A => bool. t : x --> finite t) &
-(ALL (t1::'A => bool) (t2::'A => bool) xa::'A.
-    t1 : x & t2 : x & t1 ~= t2 & xa : t1 & xa : t2 -->
-    (f::'A => nat) xa = (0::nat))
-==> nsum (Union x) f = nsum x (%t::'A => bool. nsum t f)"
-  by (import hollight NSUM_UNIONS_NONZERO)
-
-lemma NSUM_CASES: "finite (x::'A => bool)
-==> nsum x
-     (%x::'A.
-         if (xa::'A => bool) x then (xb::'A => nat) x
-         else (xc::'A => nat) x) =
-    nsum {u::'A. EX xb::'A. (xb : x & xa xb) & u = xb} xb +
-    nsum {u::'A. EX xb::'A. (xb : x & ~ xa xb) & u = xb} xc"
-  by (import hollight NSUM_CASES)
-
-lemma NSUM_CLOSED: "(P::nat => bool) (0::nat) &
-(ALL (x::nat) y::nat. P x & P y --> P (x + y)) &
-(ALL a::'A. a : (s::'A => bool) --> P ((f::'A => nat) a))
-==> P (nsum s f)"
-  by (import hollight NSUM_CLOSED)
-
-lemma NSUM_ADD_NUMSEG: "nsum {xb::nat..xc::nat} (%i::nat. (x::nat => nat) i + (xa::nat => nat) i) =
-nsum {xb..xc} x + nsum {xb..xc} xa"
-  by (import hollight NSUM_ADD_NUMSEG)
-
-lemma NSUM_LE_NUMSEG: "(!!i::nat.
-    (xb::nat) <= i & i <= (xc::nat)
-    ==> (x::nat => nat) i <= (xa::nat => nat) i)
-==> nsum {xb..xc} x <= nsum {xb..xc} xa"
-  by (import hollight NSUM_LE_NUMSEG)
-
-lemma NSUM_EQ_NUMSEG: "(!!i::nat.
-    (m::nat) <= i & i <= (n::nat) ==> (f::nat => nat) i = (g::nat => nat) i)
-==> nsum {m..n} f = nsum {m..n} g"
-  by (import hollight NSUM_EQ_NUMSEG)
-
-lemma NSUM_CONST_NUMSEG: "nsum {xa..xb} (%n. x) = (xb + 1 - xa) * x"
-  by (import hollight NSUM_CONST_NUMSEG)
-
-lemma NSUM_EQ_0_NUMSEG: "(!!i::nat. (m::nat) <= i & i <= (n::nat) ==> (x::nat => nat) i = (0::nat))
-==> nsum {m..n} x = (0::nat)"
-  by (import hollight NSUM_EQ_0_NUMSEG)
-
-lemma NSUM_EQ_0_IFF_NUMSEG: "(nsum {xa::nat..xb::nat} (x::nat => nat) = (0::nat)) =
-(ALL i::nat. xa <= i & i <= xb --> x i = (0::nat))"
-  by (import hollight NSUM_EQ_0_IFF_NUMSEG)
-
-lemma NSUM_TRIV_NUMSEG: "(n::nat) < (m::nat) ==> nsum {m..n} (f::nat => nat) = (0::nat)"
-  by (import hollight NSUM_TRIV_NUMSEG)
-
-lemma NSUM_SING_NUMSEG: "nsum {xa::nat..xa} (x::nat => nat) = x xa"
-  by (import hollight NSUM_SING_NUMSEG)
-
-lemma NSUM_CLAUSES_NUMSEG: "(ALL m. nsum {m..0} f = (if m = 0 then f 0 else 0)) &
-(ALL m n.
-    nsum {m..Suc n} f =
-    (if m <= Suc n then nsum {m..n} f + f (Suc n) else nsum {m..n} f))"
-  by (import hollight NSUM_CLAUSES_NUMSEG)
-
-lemma NSUM_SWAP_NUMSEG: "nsum {a::nat..b::nat}
- (%i::nat. nsum {c::nat..d::nat} ((f::nat => nat => nat) i)) =
-nsum {c..d} (%j::nat. nsum {a..b} (%i::nat. f i j))"
-  by (import hollight NSUM_SWAP_NUMSEG)
-
-lemma NSUM_ADD_SPLIT: "(xa::nat) <= (xb::nat) + (1::nat)
-==> nsum {xa..xb + (xc::nat)} (x::nat => nat) =
-    nsum {xa..xb} x + nsum {xb + (1::nat)..xb + xc} x"
-  by (import hollight NSUM_ADD_SPLIT)
-
-lemma NSUM_OFFSET: "nsum {(xb::nat) + (x::nat)..(xc::nat) + x} (xa::nat => nat) =
-nsum {xb..xc} (%i::nat. xa (i + x))"
-  by (import hollight NSUM_OFFSET)
-
-lemma NSUM_OFFSET_0: "(xa::nat) <= (xb::nat)
-==> nsum {xa..xb} (x::nat => nat) =
-    nsum {0::nat..xb - xa} (%i::nat. x (i + xa))"
-  by (import hollight NSUM_OFFSET_0)
-
-lemma NSUM_CLAUSES_LEFT: "(xa::nat) <= (xb::nat)
-==> nsum {xa..xb} (x::nat => nat) = x xa + nsum {xa + (1::nat)..xb} x"
-  by (import hollight NSUM_CLAUSES_LEFT)
-
-lemma NSUM_CLAUSES_RIGHT: "(0::nat) < (n::nat) & (m::nat) <= n
-==> nsum {m..n} (f::nat => nat) = nsum {m..n - (1::nat)} f + f n"
-  by (import hollight NSUM_CLAUSES_RIGHT)
-
-lemma NSUM_PAIR: "nsum {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)} (f::nat => nat) =
-nsum {m..n} (%i::nat. f ((2::nat) * i) + f ((2::nat) * i + (1::nat)))"
-  by (import hollight NSUM_PAIR)
-
-lemma CARD_UNIONS: "finite (x::('A => bool) => bool) &
-(ALL t::'A => bool. t : x --> finite t) &
-(ALL (t::'A => bool) u::'A => bool. t : x & u : x & t ~= u --> t Int u = {})
-==> CARD (Union x) = nsum x CARD"
-  by (import hollight CARD_UNIONS)
-
-consts
-  sum :: "('q_67488 => bool) => ('q_67488 => hollight.real) => hollight.real" 
-
-defs
-  sum_def: "(op ==::(('q_67488::type => bool)
-         => ('q_67488::type => hollight.real) => hollight.real)
-        => (('q_67488::type => bool)
-            => ('q_67488::type => hollight.real) => hollight.real)
-           => prop)
- (hollight.sum::('q_67488::type => bool)
-                => ('q_67488::type => hollight.real) => hollight.real)
- ((iterate::(hollight.real => hollight.real => hollight.real)
-            => ('q_67488::type => bool)
-               => ('q_67488::type => hollight.real) => hollight.real)
-   (real_add::hollight.real => hollight.real => hollight.real))"
-
-lemma DEF_sum: "(op =::(('q_67488::type => bool)
-        => ('q_67488::type => hollight.real) => hollight.real)
-       => (('q_67488::type => bool)
-           => ('q_67488::type => hollight.real) => hollight.real)
-          => bool)
- (hollight.sum::('q_67488::type => bool)
-                => ('q_67488::type => hollight.real) => hollight.real)
- ((iterate::(hollight.real => hollight.real => hollight.real)
-            => ('q_67488::type => bool)
-               => ('q_67488::type => hollight.real) => hollight.real)
-   (real_add::hollight.real => hollight.real => hollight.real))"
-  by (import hollight DEF_sum)
-
-lemma NEUTRAL_REAL_ADD: "neutral real_add = real_of_num 0"
-  by (import hollight NEUTRAL_REAL_ADD)
-
-lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num 1"
-  by (import hollight NEUTRAL_REAL_MUL)
-
-lemma MONOIDAL_REAL_ADD: "monoidal real_add"
-  by (import hollight MONOIDAL_REAL_ADD)
-
-lemma MONOIDAL_REAL_MUL: "monoidal real_mul"
-  by (import hollight MONOIDAL_REAL_MUL)
-
-lemma SUM_CLAUSES: "(ALL x::'q_67530 => hollight.real.
-    hollight.sum {} x = real_of_num (0::nat)) &
-(ALL (x::'q_67571) (xa::'q_67571 => hollight.real) xb::'q_67571 => bool.
-    finite xb -->
-    hollight.sum (insert x xb) xa =
-    (if x : xb then hollight.sum xb xa
-     else real_add (xa x) (hollight.sum xb xa)))"
-  by (import hollight SUM_CLAUSES)
-
-lemma SUM_UNION: "finite (xa::'q_67597 => bool) &
-finite (xb::'q_67597 => bool) & xa Int xb = {}
-==> hollight.sum (xa Un xb) (x::'q_67597 => hollight.real) =
-    real_add (hollight.sum xa x) (hollight.sum xb x)"
-  by (import hollight SUM_UNION)
-
-lemma SUM_DIFF: "finite (xa::'q_67637 => bool) & (xb::'q_67637 => bool) <= xa
-==> hollight.sum (xa - xb) (x::'q_67637 => hollight.real) =
-    real_sub (hollight.sum xa x) (hollight.sum xb x)"
-  by (import hollight SUM_DIFF)
-
-lemma SUM_INCL_EXCL: "finite (x::'A => bool) & finite (xa::'A => bool)
-==> real_add (hollight.sum x (xb::'A => hollight.real))
-     (hollight.sum xa xb) =
-    real_add (hollight.sum (x Un xa) xb) (hollight.sum (x Int xa) xb)"
-  by (import hollight SUM_INCL_EXCL)
-
-lemma SUM_SUPPORT: "hollight.sum
- (support real_add (x::'q_67726 => hollight.real) (xa::'q_67726 => bool))
- x =
-hollight.sum xa x"
-  by (import hollight SUM_SUPPORT)
-
-lemma SUM_ADD: "finite (xb::'q_67760 => bool)
-==> hollight.sum xb
-     (%xb::'q_67760.
-         real_add ((x::'q_67760 => hollight.real) xb)
-          ((xa::'q_67760 => hollight.real) xb)) =
-    real_add (hollight.sum xb x) (hollight.sum xb xa)"
-  by (import hollight SUM_ADD)
-
-lemma SUM_ADD_GEN: "finite
- {xa::'q_67851.
-  EX xc::'q_67851.
-     (xc : (xb::'q_67851 => bool) &
-      (x::'q_67851 => hollight.real) xc ~= real_of_num (0::nat)) &
-     xa = xc} &
-finite
- {x::'q_67851.
-  EX xc::'q_67851.
-     (xc : xb &
-      (xa::'q_67851 => hollight.real) xc ~= real_of_num (0::nat)) &
-     x = xc}
-==> hollight.sum xb (%xb::'q_67851. real_add (x xb) (xa xb)) =
-    real_add (hollight.sum xb x) (hollight.sum xb xa)"
-  by (import hollight SUM_ADD_GEN)
-
-lemma SUM_EQ_0: "(!!xb::'A.
-    xb : (xa::'A => bool)
-    ==> (x::'A => hollight.real) xb = real_of_num (0::nat))
-==> hollight.sum xa x = real_of_num (0::nat)"
-  by (import hollight SUM_EQ_0)
-
-lemma SUM_0: "hollight.sum (x::'A => bool) (%n::'A. real_of_num (0::nat)) =
-real_of_num (0::nat)"
-  by (import hollight SUM_0)
-
-lemma SUM_LMUL: "hollight.sum (s::'A => bool)
- (%x::'A. real_mul (c::hollight.real) ((f::'A => hollight.real) x)) =
-real_mul c (hollight.sum s f)"
-  by (import hollight SUM_LMUL)
-
-lemma SUM_RMUL: "hollight.sum (xb::'A => bool)
- (%xb::'A. real_mul ((x::'A => hollight.real) xb) (xa::hollight.real)) =
-real_mul (hollight.sum xb x) xa"
-  by (import hollight SUM_RMUL)
-
-lemma SUM_NEG: "hollight.sum (xa::'q_68051 => bool)
- (%xa::'q_68051. real_neg ((x::'q_68051 => hollight.real) xa)) =
-real_neg (hollight.sum xa x)"
-  by (import hollight SUM_NEG)
-
-lemma SUM_SUB: "finite (xb::'q_68086 => bool)
-==> hollight.sum xb
-     (%xb::'q_68086.
-         real_sub ((x::'q_68086 => hollight.real) xb)
-          ((xa::'q_68086 => hollight.real) xb)) =
-    real_sub (hollight.sum xb x) (hollight.sum xb xa)"
-  by (import hollight SUM_SUB)
-
-lemma SUM_LE: "finite (xb::'q_68128 => bool) &
-(ALL xc::'q_68128.
-    xc : xb -->
-    real_le ((x::'q_68128 => hollight.real) xc)
-     ((xa::'q_68128 => hollight.real) xc))
-==> real_le (hollight.sum xb x) (hollight.sum xb xa)"
-  by (import hollight SUM_LE)
-
-lemma SUM_LT: "finite (s::'A => bool) &
-(ALL x::'A.
-    x : s -->
-    real_le ((f::'A => hollight.real) x) ((g::'A => hollight.real) x)) &
-(EX x::'A. x : s & real_lt (f x) (g x))
-==> real_lt (hollight.sum s f) (hollight.sum s g)"
-  by (import hollight SUM_LT)
-
-lemma SUM_LT_ALL: "finite (s::'q_68250 => bool) &
-s ~= {} &
-(ALL x::'q_68250.
-    x : s -->
-    real_lt ((f::'q_68250 => hollight.real) x)
-     ((g::'q_68250 => hollight.real) x))
-==> real_lt (hollight.sum s f) (hollight.sum s g)"
-  by (import hollight SUM_LT_ALL)
-
-lemma SUM_EQ: "(!!xc::'q_68288.
-    xc : (xb::'q_68288 => bool)
-    ==> (x::'q_68288 => hollight.real) xc =
-        (xa::'q_68288 => hollight.real) xc)
-==> hollight.sum xb x = hollight.sum xb xa"
-  by (import hollight SUM_EQ)
-
-lemma SUM_ABS: "finite (s::'q_68347 => bool)
-==> real_le (real_abs (hollight.sum s (f::'q_68347 => hollight.real)))
-     (hollight.sum s (%x::'q_68347. real_abs (f x)))"
-  by (import hollight SUM_ABS)
-
-lemma SUM_ABS_LE: "finite (s::'A => bool) &
-(ALL x::'A.
-    x : s -->
-    real_le (real_abs ((f::'A => hollight.real) x))
-     ((g::'A => hollight.real) x))
-==> real_le (real_abs (hollight.sum s f)) (hollight.sum s g)"
-  by (import hollight SUM_ABS_LE)
-
-lemma SUM_CONST: "finite (s::'q_68423 => bool)
-==> hollight.sum s (%n::'q_68423. c::hollight.real) =
-    real_mul (real_of_num (CARD s)) c"
-  by (import hollight SUM_CONST)
-
-lemma SUM_POS_LE: "finite (xa::'q_68465 => bool) &
-(ALL xb::'q_68465.
-    xb : xa -->
-    real_le (real_of_num (0::nat)) ((x::'q_68465 => hollight.real) xb))
-==> real_le (real_of_num (0::nat)) (hollight.sum xa x)"
-  by (import hollight SUM_POS_LE)
-
-lemma SUM_POS_BOUND: "[| finite (x::'A => bool) &
-   (ALL xa::'A.
-       xa : x -->
-       real_le (real_of_num (0::nat)) ((f::'A => hollight.real) xa)) &
-   real_le (hollight.sum x f) (b::hollight.real);
-   (xa::'A) : x |]
-==> real_le (f xa) b"
-  by (import hollight SUM_POS_BOUND)
-
-lemma SUM_POS_EQ_0: "[| finite (xa::'q_68612 => bool) &
-   (ALL xb::'q_68612.
-       xb : xa -->
-       real_le (real_of_num (0::nat)) ((x::'q_68612 => hollight.real) xb)) &
-   hollight.sum xa x = real_of_num (0::nat);
-   (xb::'q_68612) : xa |]
-==> x xb = real_of_num (0::nat)"
-  by (import hollight SUM_POS_EQ_0)
-
-lemma SUM_ZERO_EXISTS: "finite (s::'A => bool) &
-hollight.sum s (u::'A => hollight.real) = real_of_num (0::nat)
-==> (ALL i::'A. i : s --> u i = real_of_num (0::nat)) |
-    (EX (j::'A) k::'A.
-        j : s &
-        real_lt (u j) (real_of_num (0::nat)) &
-        k : s & real_gt (u k) (real_of_num (0::nat)))"
-  by (import hollight SUM_ZERO_EXISTS)
-
-lemma SUM_DELETE: "finite (xa::'q_68854 => bool) & (xb::'q_68854) : xa
-==> hollight.sum (xa - {xb}) (x::'q_68854 => hollight.real) =
-    real_sub (hollight.sum xa x) (x xb)"
-  by (import hollight SUM_DELETE)
-
-lemma SUM_DELETE_CASES: "finite (s::'q_68907 => bool)
-==> hollight.sum (s - {a::'q_68907}) (f::'q_68907 => hollight.real) =
-    (if a : s then real_sub (hollight.sum s f) (f a) else hollight.sum s f)"
-  by (import hollight SUM_DELETE_CASES)
-
-lemma SUM_SING: "hollight.sum {xa::'q_68930} (x::'q_68930 => hollight.real) = x xa"
-  by (import hollight SUM_SING)
-
-lemma SUM_DELTA: "hollight.sum (x::'A => bool)
- (%x::'A. if x = (xa::'A) then b::hollight.real else real_of_num (0::nat)) =
-(if xa : x then b else real_of_num (0::nat))"
-  by (import hollight SUM_DELTA)
-
-lemma SUM_SWAP: "finite (x::'A => bool) & finite (xa::'B => bool)
-==> hollight.sum x
-     (%i::'A. hollight.sum xa ((f::'A => 'B => hollight.real) i)) =
-    hollight.sum xa (%j::'B. hollight.sum x (%i::'A. f i j))"
-  by (import hollight SUM_SWAP)
-
-lemma SUM_IMAGE: "(!!(xa::'q_69070) y::'q_69070.
-    xa : (xb::'q_69070 => bool) &
-    y : xb & (x::'q_69070 => 'q_69046) xa = x y
-    ==> xa = y)
-==> hollight.sum (x ` xb) (xa::'q_69046 => hollight.real) =
-    hollight.sum xb (xa o x)"
-  by (import hollight SUM_IMAGE)
-
-lemma SUM_SUPERSET: "(xa::'A => bool) <= (xb::'A => bool) &
-(ALL xc::'A.
-    xc : xb & xc ~: xa -->
-    (x::'A => hollight.real) xc = real_of_num (0::nat))
-==> hollight.sum xb x = hollight.sum xa x"
-  by (import hollight SUM_SUPERSET)
-
-lemma SUM_UNION_RZERO: "finite (u::'A => bool) &
-(ALL x::'A.
-    x : (v::'A => bool) & x ~: u -->
-    (f::'A => hollight.real) x = real_of_num (0::nat))
-==> hollight.sum (u Un v) f = hollight.sum u f"
-  by (import hollight SUM_UNION_RZERO)
-
-lemma SUM_UNION_LZERO: "finite (v::'A => bool) &
-(ALL x::'A.
-    x : (u::'A => bool) & x ~: v -->
-    (f::'A => hollight.real) x = real_of_num (0::nat))
-==> hollight.sum (u Un v) f = hollight.sum v f"
-  by (import hollight SUM_UNION_LZERO)
-
-lemma SUM_RESTRICT: "finite (s::'q_69267 => bool)
-==> hollight.sum s
-     (%x::'q_69267.
-         if x : s then (f::'q_69267 => hollight.real) x
-         else real_of_num (0::nat)) =
-    hollight.sum s f"
-  by (import hollight SUM_RESTRICT)
-
-lemma SUM_BOUND: "finite (x::'A => bool) &
-(ALL xc::'A.
-    xc : x --> real_le ((xa::'A => hollight.real) xc) (xb::hollight.real))
-==> real_le (hollight.sum x xa) (real_mul (real_of_num (CARD x)) xb)"
-  by (import hollight SUM_BOUND)
-
-lemma SUM_BOUND_GEN: "finite (s::'A => bool) &
-s ~= {} &
-(ALL x::'A.
-    x : s -->
-    real_le ((f::'A => hollight.real) x)
-     (real_div (b::hollight.real) (real_of_num (CARD s))))
-==> real_le (hollight.sum s f) b"
-  by (import hollight SUM_BOUND_GEN)
-
-lemma SUM_ABS_BOUND: "finite (s::'A => bool) &
-(ALL x::'A.
-    x : s -->
-    real_le (real_abs ((f::'A => hollight.real) x)) (b::hollight.real))
-==> real_le (real_abs (hollight.sum s f))
-     (real_mul (real_of_num (CARD s)) b)"
-  by (import hollight SUM_ABS_BOUND)
-
-lemma SUM_BOUND_LT: "finite (s::'A => bool) &
-(ALL x::'A.
-    x : s --> real_le ((f::'A => hollight.real) x) (b::hollight.real)) &
-(EX x::'A. x : s & real_lt (f x) b)
-==> real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)"
-  by (import hollight SUM_BOUND_LT)
-
-lemma SUM_BOUND_LT_ALL: "finite (s::'q_69531 => bool) &
-s ~= {} &
-(ALL x::'q_69531.
-    x : s --> real_lt ((f::'q_69531 => hollight.real) x) (b::hollight.real))
-==> real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)"
-  by (import hollight SUM_BOUND_LT_ALL)
-
-lemma SUM_BOUND_LT_GEN: "finite (s::'A => bool) &
-s ~= {} &
-(ALL x::'A.
-    x : s -->
-    real_lt ((f::'A => hollight.real) x)
-     (real_div (b::hollight.real) (real_of_num (CARD s))))
-==> real_lt (hollight.sum s f) b"
-  by (import hollight SUM_BOUND_LT_GEN)
-
-lemma SUM_UNION_EQ: "finite (u::'q_69614 => bool) &
-(s::'q_69614 => bool) Int (t::'q_69614 => bool) = {} & s Un t = u
-==> real_add (hollight.sum s (f::'q_69614 => hollight.real))
-     (hollight.sum t f) =
-    hollight.sum u f"
-  by (import hollight SUM_UNION_EQ)
-
-lemma SUM_EQ_SUPERSET: "finite (t::'A => bool) &
-t <= (s::'A => bool) &
-(ALL x::'A.
-    x : t --> (f::'A => hollight.real) x = (g::'A => hollight.real) x) &
-(ALL x::'A. x : s & x ~: t --> f x = real_of_num (0::nat))
-==> hollight.sum s f = hollight.sum t g"
-  by (import hollight SUM_EQ_SUPERSET)
-
-lemma SUM_RESTRICT_SET: "hollight.sum
- {u::'q_69783.
-  EX xb::'q_69783.
-     (xb : (xa::'q_69783 => bool) & (x::'q_69783 => bool) xb) & u = xb}
- (xb::'q_69783 => hollight.real) =
-hollight.sum xa
- (%xa::'q_69783. if x xa then xb xa else real_of_num (0::nat))"
-  by (import hollight SUM_RESTRICT_SET)
-
-lemma SUM_SUM_RESTRICT: "finite (s::'q_69875 => bool) & finite (t::'q_69874 => bool)
-==> hollight.sum s
-     (%x::'q_69875.
-         hollight.sum
-          {u::'q_69874.
-           EX y::'q_69874.
-              (y : t & (R::'q_69875 => 'q_69874 => bool) x y) & u = y}
-          ((f::'q_69875 => 'q_69874 => hollight.real) x)) =
-    hollight.sum t
-     (%y::'q_69874.
-         hollight.sum {u::'q_69875. EX x::'q_69875. (x : s & R x y) & u = x}
-          (%x::'q_69875. f x y))"
-  by (import hollight SUM_SUM_RESTRICT)
-
-lemma CARD_EQ_SUM: "finite (x::'q_69896 => bool)
-==> real_of_num (CARD x) =
-    hollight.sum x (%x::'q_69896. real_of_num (1::nat))"
-  by (import hollight CARD_EQ_SUM)
-
-lemma SUM_MULTICOUNT_GEN: "finite (s::'A => bool) &
-finite (t::'B => bool) &
-(ALL j::'B.
-    j : t -->
-    CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} =
-    (k::'B => nat) j)
-==> hollight.sum s
-     (%i::'A.
-         real_of_num (CARD {u::'B. EX j::'B. (j : t & R i j) & u = j})) =
-    hollight.sum t (%i::'B. real_of_num (k i))"
-  by (import hollight SUM_MULTICOUNT_GEN)
-
-lemma SUM_MULTICOUNT: "finite (s::'A => bool) &
-finite (t::'B => bool) &
-(ALL j::'B.
-    j : t -->
-    CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} =
-    (k::nat))
-==> hollight.sum s
-     (%i::'A.
-         real_of_num (CARD {u::'B. EX j::'B. (j : t & R i j) & u = j})) =
-    real_of_num (k * CARD t)"
-  by (import hollight SUM_MULTICOUNT)
-
-lemma SUM_IMAGE_GEN: "finite (s::'A => bool)
-==> hollight.sum s (g::'A => hollight.real) =
-    hollight.sum ((f::'A => 'B) ` s)
-     (%y::'B. hollight.sum {u::'A. EX x::'A. (x : s & f x = y) & u = x} g)"
-  by (import hollight SUM_IMAGE_GEN)
-
-lemma SUM_GROUP: "finite (s::'A => bool) & (f::'A => 'B) ` s <= (t::'B => bool)
-==> hollight.sum t
-     (%y::'B.
-         hollight.sum {u::'A. EX x::'A. (x : s & f x = y) & u = x}
-          (g::'A => hollight.real)) =
-    hollight.sum s g"
-  by (import hollight SUM_GROUP)
-
-lemma REAL_OF_NUM_SUM: "finite (s::'q_70361 => bool)
-==> real_of_num (nsum s (f::'q_70361 => nat)) =
-    hollight.sum s (%x::'q_70361. real_of_num (f x))"
-  by (import hollight REAL_OF_NUM_SUM)
-
-lemma SUM_SUBSET: "finite (u::'A => bool) &
-finite (v::'A => bool) &
-(ALL x::'A.
-    x : u - v -->
-    real_le ((f::'A => hollight.real) x) (real_of_num (0::nat))) &
-(ALL x::'A. x : v - u --> real_le (real_of_num (0::nat)) (f x))
-==> real_le (hollight.sum u f) (hollight.sum v f)"
-  by (import hollight SUM_SUBSET)
-
-lemma SUM_SUBSET_SIMPLE: "finite (v::'A => bool) &
-(u::'A => bool) <= v &
-(ALL x::'A.
-    x : v - u -->
-    real_le (real_of_num (0::nat)) ((f::'A => hollight.real) x))
-==> real_le (hollight.sum u f) (hollight.sum v f)"
-  by (import hollight SUM_SUBSET_SIMPLE)
-
-lemma SUM_IMAGE_NONZERO: "finite (xb::'A => bool) &
-(ALL (xc::'A) xd::'A.
-    xc : xb & xd : xb & xc ~= xd & (xa::'A => 'B) xc = xa xd -->
-    (x::'B => hollight.real) (xa xc) = real_of_num (0::nat))
-==> hollight.sum (xa ` xb) x = hollight.sum xb (x o xa)"
-  by (import hollight SUM_IMAGE_NONZERO)
-
-lemma SUM_BIJECTION: "(ALL x::'A. x : (xb::'A => bool) --> (xa::'A => 'A) x : xb) &
-(ALL y::'A. y : xb --> (EX! x::'A. x : xb & xa x = y))
-==> hollight.sum xb (x::'A => hollight.real) = hollight.sum xb (x o xa)"
-  by (import hollight SUM_BIJECTION)
-
-lemma SUM_SUM_PRODUCT: "finite (x::'A => bool) &
-(ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i))
-==> hollight.sum x
-     (%x::'A. hollight.sum (xa x) ((xb::'A => 'B => hollight.real) x)) =
-    hollight.sum
-     {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)}
-     (SOME f::'A * 'B => hollight.real.
-         ALL (i::'A) j::'B. f (i, j) = xb i j)"
-  by (import hollight SUM_SUM_PRODUCT)
-
-lemma SUM_EQ_GENERAL: "(ALL y::'B.
-    y : (xa::'B => bool) -->
-    (EX! xa::'A. xa : (x::'A => bool) & (xd::'A => 'B) xa = y)) &
-(ALL xe::'A.
-    xe : x -->
-    xd xe : xa &
-    (xc::'B => hollight.real) (xd xe) = (xb::'A => hollight.real) xe)
-==> hollight.sum x xb = hollight.sum xa xc"
-  by (import hollight SUM_EQ_GENERAL)
-
-lemma SUM_EQ_GENERAL_INVERSES: "(ALL y::'B.
-    y : (xa::'B => bool) -->
-    (xe::'B => 'A) y : (x::'A => bool) & (xd::'A => 'B) (xe y) = y) &
-(ALL xf::'A.
-    xf : x -->
-    xd xf : xa &
-    xe (xd xf) = xf &
-    (xc::'B => hollight.real) (xd xf) = (xb::'A => hollight.real) xf)
-==> hollight.sum x xb = hollight.sum xa xc"
-  by (import hollight SUM_EQ_GENERAL_INVERSES)
-
-lemma SUM_INJECTION: "finite (xb::'q_71007 => bool) &
-(ALL x::'q_71007. x : xb --> (xa::'q_71007 => 'q_71007) x : xb) &
-(ALL (x::'q_71007) y::'q_71007. x : xb & y : xb & xa x = xa y --> x = y)
-==> hollight.sum xb ((x::'q_71007 => hollight.real) o xa) =
-    hollight.sum xb x"
-  by (import hollight SUM_INJECTION)
-
-lemma SUM_UNION_NONZERO: "finite (xa::'q_71050 => bool) &
-finite (xb::'q_71050 => bool) &
-(ALL xc::'q_71050.
-    xc : xa Int xb -->
-    (x::'q_71050 => hollight.real) xc = real_of_num (0::nat))
-==> hollight.sum (xa Un xb) x =
-    real_add (hollight.sum xa x) (hollight.sum xb x)"
-  by (import hollight SUM_UNION_NONZERO)
-
-lemma SUM_UNIONS_NONZERO: "finite (x::('A => bool) => bool) &
-(ALL t::'A => bool. t : x --> finite t) &
-(ALL (t1::'A => bool) (t2::'A => bool) xa::'A.
-    t1 : x & t2 : x & t1 ~= t2 & xa : t1 & xa : t2 -->
-    (f::'A => hollight.real) xa = real_of_num (0::nat))
-==> hollight.sum (Union x) f =
-    hollight.sum x (%t::'A => bool. hollight.sum t f)"
-  by (import hollight SUM_UNIONS_NONZERO)
-
-lemma SUM_CASES: "finite (x::'A => bool)
-==> hollight.sum x
-     (%x::'A.
-         if (xa::'A => bool) x then (xb::'A => hollight.real) x
-         else (xc::'A => hollight.real) x) =
-    real_add (hollight.sum {u::'A. EX xb::'A. (xb : x & xa xb) & u = xb} xb)
-     (hollight.sum {u::'A. EX xb::'A. (xb : x & ~ xa xb) & u = xb} xc)"
-  by (import hollight SUM_CASES)
-
-lemma SUM_CASES_1: "finite (s::'q_71319 => bool) & (a::'q_71319) : s
-==> hollight.sum s
-     (%x::'q_71319.
-         if x = a then y::hollight.real
-         else (f::'q_71319 => hollight.real) x) =
-    real_add (hollight.sum s f) (real_sub y (f a))"
-  by (import hollight SUM_CASES_1)
-
-lemma SUM_LE_INCLUDED: "finite (s::'A => bool) &
-finite (t::'B => bool) &
-(ALL y::'B.
-    y : t --> real_le (real_of_num (0::nat)) ((g::'B => hollight.real) y)) &
-(ALL x::'A.
-    x : s -->
-    (EX y::'B.
-        y : t &
-        (i::'B => 'A) y = x & real_le ((f::'A => hollight.real) x) (g y)))
-==> real_le (hollight.sum s f) (hollight.sum t g)"
-  by (import hollight SUM_LE_INCLUDED)
-
-lemma SUM_IMAGE_LE: "finite (s::'A => bool) &
-(ALL x::'A.
-    x : s -->
-    real_le (real_of_num (0::nat))
-     ((g::'B => hollight.real) ((f::'A => 'B) x)))
-==> real_le (hollight.sum (f ` s) g) (hollight.sum s (g o f))"
-  by (import hollight SUM_IMAGE_LE)
-
-lemma SUM_CLOSED: "(P::hollight.real => bool) (real_of_num (0::nat)) &
-(ALL (x::hollight.real) y::hollight.real. P x & P y --> P (real_add x y)) &
-(ALL a::'A. a : (s::'A => bool) --> P ((f::'A => hollight.real) a))
-==> P (hollight.sum s f)"
-  by (import hollight SUM_CLOSED)
-
-lemma SUM_ADD_NUMSEG: "hollight.sum {xb::nat..xc::nat}
- (%i::nat.
-     real_add ((x::nat => hollight.real) i)
-      ((xa::nat => hollight.real) i)) =
-real_add (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)"
-  by (import hollight SUM_ADD_NUMSEG)
-
-lemma SUM_SUB_NUMSEG: "hollight.sum {xb::nat..xc::nat}
- (%i::nat.
-     real_sub ((x::nat => hollight.real) i)
-      ((xa::nat => hollight.real) i)) =
-real_sub (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)"
-  by (import hollight SUM_SUB_NUMSEG)
-
-lemma SUM_LE_NUMSEG: "(!!i::nat.
-    (xb::nat) <= i & i <= (xc::nat)
-    ==> real_le ((x::nat => hollight.real) i)
-         ((xa::nat => hollight.real) i))
-==> real_le (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)"
-  by (import hollight SUM_LE_NUMSEG)
-
-lemma SUM_EQ_NUMSEG: "(!!i::nat.
-    (m::nat) <= i & i <= (n::nat)
-    ==> (f::nat => hollight.real) i = (g::nat => hollight.real) i)
-==> hollight.sum {m..n} f = hollight.sum {m..n} g"
-  by (import hollight SUM_EQ_NUMSEG)
-
-lemma SUM_ABS_NUMSEG: "real_le
- (real_abs (hollight.sum {xa::nat..xb::nat} (x::nat => hollight.real)))
- (hollight.sum {xa..xb} (%i::nat. real_abs (x i)))"
-  by (import hollight SUM_ABS_NUMSEG)
-
-lemma SUM_CONST_NUMSEG: "hollight.sum {xa..xb} (%n. x) = real_mul (real_of_num (xb + 1 - xa)) x"
-  by (import hollight SUM_CONST_NUMSEG)
-
-lemma SUM_EQ_0_NUMSEG: "(!!i::nat.
-    (m::nat) <= i & i <= (n::nat)
-    ==> (x::nat => hollight.real) i = real_of_num (0::nat))
-==> hollight.sum {m..n} x = real_of_num (0::nat)"
-  by (import hollight SUM_EQ_0_NUMSEG)
-
-lemma SUM_TRIV_NUMSEG: "(n::nat) < (m::nat)
-==> hollight.sum {m..n} (f::nat => hollight.real) = real_of_num (0::nat)"
-  by (import hollight SUM_TRIV_NUMSEG)
-
-lemma SUM_POS_LE_NUMSEG: "(!!p::nat.
-    (x::nat) <= p & p <= (xa::nat)
-    ==> real_le (real_of_num (0::nat)) ((xb::nat => hollight.real) p))
-==> real_le (real_of_num (0::nat)) (hollight.sum {x..xa} xb)"
-  by (import hollight SUM_POS_LE_NUMSEG)
-
-lemma SUM_POS_EQ_0_NUMSEG: "[| (ALL p::nat.
-       (m::nat) <= p & p <= (n::nat) -->
-       real_le (real_of_num (0::nat)) ((f::nat => hollight.real) p)) &
-   hollight.sum {m..n} f = real_of_num (0::nat);
-   m <= (p::nat) & p <= n |]
-==> f p = real_of_num (0::nat)"
-  by (import hollight SUM_POS_EQ_0_NUMSEG)
-
-lemma SUM_SING_NUMSEG: "hollight.sum {xa::nat..xa} (x::nat => hollight.real) = x xa"
-  by (import hollight SUM_SING_NUMSEG)
-
-lemma SUM_CLAUSES_NUMSEG: "(ALL m. hollight.sum {m..0} f = (if m = 0 then f 0 else real_of_num 0)) &
-(ALL m n.
-    hollight.sum {m..Suc n} f =
-    (if m <= Suc n then real_add (hollight.sum {m..n} f) (f (Suc n))
-     else hollight.sum {m..n} f))"
-  by (import hollight SUM_CLAUSES_NUMSEG)
-
-lemma SUM_SWAP_NUMSEG: "hollight.sum {a::nat..b::nat}
- (%i::nat.
-     hollight.sum {c::nat..d::nat} ((f::nat => nat => hollight.real) i)) =
-hollight.sum {c..d} (%j::nat. hollight.sum {a..b} (%i::nat. f i j))"
-  by (import hollight SUM_SWAP_NUMSEG)
-
-lemma SUM_ADD_SPLIT: "(xa::nat) <= (xb::nat) + (1::nat)
-==> hollight.sum {xa..xb + (xc::nat)} (x::nat => hollight.real) =
-    real_add (hollight.sum {xa..xb} x)
-     (hollight.sum {xb + (1::nat)..xb + xc} x)"
-  by (import hollight SUM_ADD_SPLIT)
-
-lemma SUM_OFFSET: "hollight.sum {(xb::nat) + (x::nat)..(xc::nat) + x}
- (xa::nat => hollight.real) =
-hollight.sum {xb..xc} (%i::nat. xa (i + x))"
-  by (import hollight SUM_OFFSET)
-
-lemma SUM_OFFSET_0: "(xa::nat) <= (xb::nat)
-==> hollight.sum {xa..xb} (x::nat => hollight.real) =
-    hollight.sum {0::nat..xb - xa} (%i::nat. x (i + xa))"
-  by (import hollight SUM_OFFSET_0)
-
-lemma SUM_CLAUSES_LEFT: "(xa::nat) <= (xb::nat)
-==> hollight.sum {xa..xb} (x::nat => hollight.real) =
-    real_add (x xa) (hollight.sum {xa + (1::nat)..xb} x)"
-  by (import hollight SUM_CLAUSES_LEFT)
-
-lemma SUM_CLAUSES_RIGHT: "(0::nat) < (n::nat) & (m::nat) <= n
-==> hollight.sum {m..n} (f::nat => hollight.real) =
-    real_add (hollight.sum {m..n - (1::nat)} f) (f n)"
-  by (import hollight SUM_CLAUSES_RIGHT)
-
-lemma SUM_PAIR: "hollight.sum {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)}
- (f::nat => hollight.real) =
-hollight.sum {m..n}
- (%i::nat. real_add (f ((2::nat) * i)) (f ((2::nat) * i + (1::nat))))"
-  by (import hollight SUM_PAIR)
-
-lemma REAL_OF_NUM_SUM_NUMSEG: "real_of_num (nsum {xa::nat..xb::nat} (x::nat => nat)) =
-hollight.sum {xa..xb} (%i::nat. real_of_num (x i))"
-  by (import hollight REAL_OF_NUM_SUM_NUMSEG)
-
-lemma SUM_PARTIAL_SUC: "hollight.sum {m::nat..n::nat}
- (%k::nat.
-     real_mul ((f::nat => hollight.real) k)
-      (real_sub ((g::nat => hollight.real) (k + (1::nat))) (g k))) =
-(if m <= n
- then real_sub
-       (real_sub (real_mul (f (n + (1::nat))) (g (n + (1::nat))))
-         (real_mul (f m) (g m)))
-       (hollight.sum {m..n}
-         (%k::nat.
-             real_mul (g (k + (1::nat)))
-              (real_sub (f (k + (1::nat))) (f k))))
- else real_of_num (0::nat))"
-  by (import hollight SUM_PARTIAL_SUC)
-
-lemma SUM_PARTIAL_PRE: "hollight.sum {m::nat..n::nat}
- (%k::nat.
-     real_mul ((f::nat => hollight.real) k)
-      (real_sub ((g::nat => hollight.real) k) (g (k - (1::nat))))) =
-(if m <= n
- then real_sub
-       (real_sub (real_mul (f (n + (1::nat))) (g n))
-         (real_mul (f m) (g (m - (1::nat)))))
-       (hollight.sum {m..n}
-         (%k::nat. real_mul (g k) (real_sub (f (k + (1::nat))) (f k))))
- else real_of_num (0::nat))"
-  by (import hollight SUM_PARTIAL_PRE)
-
-lemma SUM_DIFFS: "hollight.sum {x::nat..xa::nat}
- (%x::nat. real_sub ((f::nat => hollight.real) x) (f (x + (1::nat)))) =
-(if x <= xa then real_sub (f x) (f (xa + (1::nat)))
- else real_of_num (0::nat))"
-  by (import hollight SUM_DIFFS)
-
-lemma SUM_DIFFS_ALT: "hollight.sum {m::nat..n::nat}
- (%x::nat. real_sub ((f::nat => hollight.real) (x + (1::nat))) (f x)) =
-(if m <= n then real_sub (f (n + (1::nat))) (f m) else real_of_num (0::nat))"
-  by (import hollight SUM_DIFFS_ALT)
-
-lemma SUM_COMBINE_R: "(m::nat) <= (n::nat) + (1::nat) & n <= (p::nat)
-==> real_add (hollight.sum {m..n} (f::nat => hollight.real))
-     (hollight.sum {n + (1::nat)..p} f) =
-    hollight.sum {m..p} f"
-  by (import hollight SUM_COMBINE_R)
-
-lemma SUM_COMBINE_L: "(0::nat) < (n::nat) & (m::nat) <= n & n <= (p::nat) + (1::nat)
-==> real_add (hollight.sum {m..n - (1::nat)} (f::nat => hollight.real))
-     (hollight.sum {n..p} f) =
-    hollight.sum {m..p} f"
-  by (import hollight SUM_COMBINE_L)
-
-lemma REAL_SUB_POW: "1 <= xb
-==> real_sub (real_pow x xb) (real_pow xa xb) =
-    real_mul (real_sub x xa)
-     (hollight.sum {0..xb - 1}
-       (%i. real_mul (real_pow x i) (real_pow xa (xb - 1 - i))))"
-  by (import hollight REAL_SUB_POW)
-
-lemma REAL_SUB_POW_R1: "1 <= n
-==> real_sub (real_pow x n) (real_of_num 1) =
-    real_mul (real_sub x (real_of_num 1))
-     (hollight.sum {0..n - 1} (real_pow x))"
-  by (import hollight REAL_SUB_POW_R1)
-
-lemma REAL_SUB_POW_L1: "1 <= xa
-==> real_sub (real_of_num 1) (real_pow x xa) =
-    real_mul (real_sub (real_of_num 1) x)
-     (hollight.sum {0..xa - 1} (real_pow x))"
-  by (import hollight REAL_SUB_POW_L1)
-
-definition
-  dimindex :: "('A => bool) => nat"  where
-  "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop)
- (dimindex::('A::type => bool) => nat)
- (%u::'A::type => bool.
-     (If::bool => nat => nat => nat)
-      ((finite::('A::type => bool) => bool) (UNIV::'A::type => bool))
-      ((CARD::('A::type => bool) => nat) (UNIV::'A::type => bool)) (1::nat))"
-
-lemma DEF_dimindex: "(op =::(('A::type => bool) => nat) => (('A::type => bool) => nat) => bool)
- (dimindex::('A::type => bool) => nat)
- (%u::'A::type => bool.
-     (If::bool => nat => nat => nat)
-      ((finite::('A::type => bool) => bool) (UNIV::'A::type => bool))
-      ((CARD::('A::type => bool) => nat) (UNIV::'A::type => bool)) (1::nat))"
-  by (import hollight DEF_dimindex)
-
-lemma DIMINDEX_NONZERO: "dimindex (s::'A => bool) ~= (0::nat)"
-  by (import hollight DIMINDEX_NONZERO)
-
-lemma DIMINDEX_GE_1: "(1::nat) <= dimindex (x::'A => bool)"
-  by (import hollight DIMINDEX_GE_1)
-
-lemma DIMINDEX_UNIV: "(op =::nat => nat => bool)
- ((dimindex::('A::type => bool) => nat) (x::'A::type => bool))
- ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))"
-  by (import hollight DIMINDEX_UNIV)
-
-lemma DIMINDEX_UNIQUE: "(op ==>::prop => prop => prop)
- ((Trueprop::bool => prop)
-   ((HAS_SIZE::('A::type => bool) => nat => bool) (UNIV::'A::type => bool)
-     (n::nat)))
- ((Trueprop::bool => prop)
-   ((op =::nat => nat => bool)
-     ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool)) n))"
-  by (import hollight DIMINDEX_UNIQUE)
-
-typedef (open) ('A) finite_image = "{x::nat. x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) (dimindex UNIV)}"  morphisms "dest_finite_image" "finite_index"
-  apply (rule light_ex_imp_nonempty[where t="SOME x::nat. x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) (dimindex UNIV)"])
-  by (import hollight TYDEF_finite_image)
-
-syntax
-  dest_finite_image :: _ 
-
-syntax
-  finite_index :: _ 
-
-lemmas "TYDEF_finite_image_@intern" = typedef_hol2hollight 
-  [where a="a :: 'A finite_image" and r=r ,
-   OF type_definition_finite_image]
-
-lemma FINITE_IMAGE_IMAGE: "(op =::('A::type finite_image => bool)
-       => ('A::type finite_image => bool) => bool)
- (UNIV::'A::type finite_image => bool)
- ((op `::(nat => 'A::type finite_image)
-         => (nat => bool) => 'A::type finite_image => bool)
-   (finite_index::nat => 'A::type finite_image)
-   ((atLeastAtMost::nat => nat => nat => bool) (1::nat)
-     ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))))"
-  by (import hollight FINITE_IMAGE_IMAGE)
-
-lemma HAS_SIZE_FINITE_IMAGE: "(HAS_SIZE::('A::type finite_image => bool) => nat => bool)
- (UNIV::'A::type finite_image => bool)
- ((dimindex::('A::type => bool) => nat) (s::'A::type => bool))"
-  by (import hollight HAS_SIZE_FINITE_IMAGE)
-
-lemma CARD_FINITE_IMAGE: "(op =::nat => nat => bool)
- ((CARD::('A::type finite_image => bool) => nat)
-   (UNIV::'A::type finite_image => bool))
- ((dimindex::('A::type => bool) => nat) (s::'A::type => bool))"
-  by (import hollight CARD_FINITE_IMAGE)
-
-lemma FINITE_FINITE_IMAGE: "(finite::('A::type finite_image => bool) => bool)
- (UNIV::'A::type finite_image => bool)"
-  by (import hollight FINITE_FINITE_IMAGE)
-
-lemma DIMINDEX_FINITE_IMAGE: "dimindex (s::'A finite_image => bool) = dimindex (t::'A => bool)"
-  by (import hollight DIMINDEX_FINITE_IMAGE)
-
-lemma FINITE_INDEX_WORKS: "(Ex1::(nat => bool) => bool)
- (%xa::nat.
-     (op &::bool => bool => bool) ((op <=::nat => nat => bool) (1::nat) xa)
-      ((op &::bool => bool => bool)
-        ((op <=::nat => nat => bool) xa
-          ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool)))
-        ((op =::'A::type finite_image => 'A::type finite_image => bool)
-          ((finite_index::nat => 'A::type finite_image) xa)
-          (x::'A::type finite_image))))"
-  by (import hollight FINITE_INDEX_WORKS)
-
-lemma FINITE_INDEX_INJ: "(op ==>::prop => prop => prop)
- ((Trueprop::bool => prop)
-   ((op &::bool => bool => bool)
-     ((op <=::nat => nat => bool) (1::nat) (i::nat))
-     ((op &::bool => bool => bool)
-       ((op <=::nat => nat => bool) i
-         ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool)))
-       ((op &::bool => bool => bool)
-         ((op <=::nat => nat => bool) (1::nat) (j::nat))
-         ((op <=::nat => nat => bool) j
-           ((dimindex::('A::type => bool) => nat)
-             (UNIV::'A::type => bool)))))))
- ((Trueprop::bool => prop)
-   ((op =::bool => bool => bool)
-     ((op =::'A::type finite_image => 'A::type finite_image => bool)
-       ((finite_index::nat => 'A::type finite_image) i)
-       ((finite_index::nat => 'A::type finite_image) j))
-     ((op =::nat => nat => bool) i j)))"
-  by (import hollight FINITE_INDEX_INJ)
-
-lemma FORALL_FINITE_INDEX: "(op =::bool => bool => bool)
- ((All::('N::type finite_image => bool) => bool)
-   (P::'N::type finite_image => bool))
- ((All::(nat => bool) => bool)
-   (%i::nat.
-       (op -->::bool => bool => bool)
-        ((op &::bool => bool => bool)
-          ((op <=::nat => nat => bool) (1::nat) i)
-          ((op <=::nat => nat => bool) i
-            ((dimindex::('N::type => bool) => nat)
-              (UNIV::'N::type => bool))))
-        (P ((finite_index::nat => 'N::type finite_image) i))))"
-  by (import hollight FORALL_FINITE_INDEX)
-
-typedef (open) ('A, 'B) cart = "{f. True}"  morphisms "dest_cart" "mk_cart"
-  apply (rule light_ex_imp_nonempty[where t="SOME f. True"])
-  by (import hollight TYDEF_cart)
-
-syntax
-  dest_cart :: _ 
-
-syntax
-  mk_cart :: _ 
-
-lemmas "TYDEF_cart_@intern" = typedef_hol2hollight 
-  [where a="a :: ('A, 'B) cart" and r=r ,
-   OF type_definition_cart]
-
-consts
-  "$" :: "('q_73536, 'q_73546) cart => nat => 'q_73536" ("$")
-
-defs
-  "$_def": "$ == %(u::('q_73536, 'q_73546) cart) ua::nat. dest_cart u (finite_index ua)"
-
-lemma "DEF_$": "$ = (%(u::('q_73536, 'q_73546) cart) ua::nat. dest_cart u (finite_index ua))"
-  by (import hollight "DEF_$")
-
-lemma CART_EQ: "(op =::bool => bool => bool)
- ((op =::('A::type, 'B::type) cart => ('A::type, 'B::type) cart => bool)
-   (x::('A::type, 'B::type) cart) (y::('A::type, 'B::type) cart))
- ((All::(nat => bool) => bool)
-   (%xa::nat.
-       (op -->::bool => bool => bool)
-        ((op &::bool => bool => bool)
-          ((op <=::nat => nat => bool) (1::nat) xa)
-          ((op <=::nat => nat => bool) xa
-            ((dimindex::('B::type => bool) => nat)
-              (UNIV::'B::type => bool))))
-        ((op =::'A::type => 'A::type => bool)
-          (($::('A::type, 'B::type) cart => nat => 'A::type) x xa)
-          (($::('A::type, 'B::type) cart => nat => 'A::type) y xa))))"
-  by (import hollight CART_EQ)
-
-definition
-  lambda :: "(nat => 'A) => ('A, 'B) cart"  where
-  "(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart)
-        => ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop)
- (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
- (%u::nat => 'A::type.
-     (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
-      (%f::('A::type, 'B::type) cart.
-          (All::(nat => bool) => bool)
-           (%i::nat.
-               (op -->::bool => bool => bool)
-                ((op &::bool => bool => bool)
-                  ((op <=::nat => nat => bool) (1::nat) i)
-                  ((op <=::nat => nat => bool) i
-                    ((dimindex::('B::type => bool) => nat)
-                      (UNIV::'B::type => bool))))
-                ((op =::'A::type => 'A::type => bool)
-                  (($::('A::type, 'B::type) cart => nat => 'A::type) f i)
-                  (u i)))))"
-
-lemma DEF_lambda: "(op =::((nat => 'A::type) => ('A::type, 'B::type) cart)
-       => ((nat => 'A::type) => ('A::type, 'B::type) cart) => bool)
- (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
- (%u::nat => 'A::type.
-     (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
-      (%f::('A::type, 'B::type) cart.
-          (All::(nat => bool) => bool)
-           (%i::nat.
-               (op -->::bool => bool => bool)
-                ((op &::bool => bool => bool)
-                  ((op <=::nat => nat => bool) (1::nat) i)
-                  ((op <=::nat => nat => bool) i
-                    ((dimindex::('B::type => bool) => nat)
-                      (UNIV::'B::type => bool))))
-                ((op =::'A::type => 'A::type => bool)
-                  (($::('A::type, 'B::type) cart => nat => 'A::type) f i)
-                  (u i)))))"
-  by (import hollight DEF_lambda)
-
-lemma LAMBDA_BETA: "(op ==>::prop => prop => prop)
- ((Trueprop::bool => prop)
-   ((op &::bool => bool => bool)
-     ((op <=::nat => nat => bool) (1::nat) (x::nat))
-     ((op <=::nat => nat => bool) x
-       ((dimindex::('B::type => bool) => nat) (UNIV::'B::type => bool)))))
- ((Trueprop::bool => prop)
-   ((op =::'A::type => 'A::type => bool)
-     (($::('A::type, 'B::type) cart => nat => 'A::type)
-       ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
-         (g::nat => 'A::type))
-       x)
-     (g x)))"
-  by (import hollight LAMBDA_BETA)
-
-lemma LAMBDA_UNIQUE: "(op =::bool => bool => bool)
- ((All::(nat => bool) => bool)
-   (%i::nat.
-       (op -->::bool => bool => bool)
-        ((op &::bool => bool => bool)
-          ((op <=::nat => nat => bool) (1::nat) i)
-          ((op <=::nat => nat => bool) i
-            ((dimindex::('B::type => bool) => nat)
-              (UNIV::'B::type => bool))))
-        ((op =::'A::type => 'A::type => bool)
-          (($::('A::type, 'B::type) cart => nat => 'A::type)
-            (x::('A::type, 'B::type) cart) i)
-          ((xa::nat => 'A::type) i))))
- ((op =::('A::type, 'B::type) cart => ('A::type, 'B::type) cart => bool)
-   ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) xa) x)"
-  by (import hollight LAMBDA_UNIQUE)
-
-lemma LAMBDA_ETA: "lambda ($ (x::('q_73734, 'q_73738) cart)) = x"
-  by (import hollight LAMBDA_ETA)
-
-lemma FINITE_INDEX_INRANGE: "(Ex::(nat => bool) => bool)
- (%xa::nat.
-     (op &::bool => bool => bool) ((op <=::nat => nat => bool) (1::nat) xa)
-      ((op &::bool => bool => bool)
-        ((op <=::nat => nat => bool) xa
-          ((dimindex::('N::type => bool) => nat) (UNIV::'N::type => bool)))
-        ((All::(('A::type, 'N::type) cart => bool) => bool)
-          (%xb::('A::type, 'N::type) cart.
-              (op =::'A::type => 'A::type => bool)
-               (($::('A::type, 'N::type) cart => nat => 'A::type) xb
-                 (x::nat))
-               (($::('A::type, 'N::type) cart => nat => 'A::type) xb xa)))))"
-  by (import hollight FINITE_INDEX_INRANGE)
-
-lemma CART_EQ_FULL: "((x::('A, 'N) cart) = (y::('A, 'N) cart)) = (ALL i::nat. $ x i = $ y i)"
-  by (import hollight CART_EQ_FULL)
-
-typedef (open) ('A, 'B) finite_sum = "{x::nat.
- x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat)))
-      (dimindex UNIV + dimindex UNIV)}"  morphisms "dest_finite_sum" "mk_finite_sum"
-  apply (rule light_ex_imp_nonempty[where t="SOME x::nat.
-   x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat)))
-        (dimindex UNIV + dimindex UNIV)"])
-  by (import hollight TYDEF_finite_sum)
-
-syntax
-  dest_finite_sum :: _ 
-
-syntax
-  mk_finite_sum :: _ 
-
-lemmas "TYDEF_finite_sum_@intern" = typedef_hol2hollight 
-  [where a="a :: ('A, 'B) finite_sum" and r=r ,
-   OF type_definition_finite_sum]
-
-definition
-  pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart"  where
-  "(op ==::(('A::type, 'M::type) cart
-         => ('A::type, 'N::type) cart
-            => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-        => (('A::type, 'M::type) cart
-            => ('A::type, 'N::type) cart
-               => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-           => prop)
- (pastecart::('A::type, 'M::type) cart
-             => ('A::type, 'N::type) cart
-                => ('A::type, ('M::type, 'N::type) finite_sum) cart)
- (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
-     (lambda::(nat => 'A::type)
-              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-      (%i::nat.
-          (If::bool => 'A::type => 'A::type => 'A::type)
-           ((op <=::nat => nat => bool) i
-             ((dimindex::('M::type => bool) => nat)
-               (UNIV::'M::type => bool)))
-           (($::('A::type, 'M::type) cart => nat => 'A::type) u i)
-           (($::('A::type, 'N::type) cart => nat => 'A::type) ua
-             ((op -::nat => nat => nat) i
-               ((dimindex::('M::type => bool) => nat)
-                 (UNIV::'M::type => bool))))))"
-
-lemma DEF_pastecart: "(op =::(('A::type, 'M::type) cart
-        => ('A::type, 'N::type) cart
-           => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-       => (('A::type, 'M::type) cart
-           => ('A::type, 'N::type) cart
-              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-          => bool)
- (pastecart::('A::type, 'M::type) cart
-             => ('A::type, 'N::type) cart
-                => ('A::type, ('M::type, 'N::type) finite_sum) cart)
- (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
-     (lambda::(nat => 'A::type)
-              => ('A::type, ('M::type, 'N::type) finite_sum) cart)
-      (%i::nat.
-          (If::bool => 'A::type => 'A::type => 'A::type)
-           ((op <=::nat => nat => bool) i
-             ((dimindex::('M::type => bool) => nat)
-               (UNIV::'M::type => bool)))
-           (($::('A::type, 'M::type) cart => nat => 'A::type) u i)
-           (($::('A::type, 'N::type) cart => nat => 'A::type) ua
-             ((op -::nat => nat => nat) i
-               ((dimindex::('M::type => bool) => nat)
-                 (UNIV::'M::type => bool))))))"
-  by (import hollight DEF_pastecart)
-
-definition
-  fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart"  where
-  "fstcart == %u::('A, ('M, 'N) finite_sum) cart. lambda ($ u)"
-
-lemma DEF_fstcart: "fstcart = (%u::('A, ('M, 'N) finite_sum) cart. lambda ($ u))"
-  by (import hollight DEF_fstcart)
-
-definition
-  sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart"  where
-  "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart
-         => ('A::type, 'N::type) cart)
-        => (('A::type, ('M::type, 'N::type) finite_sum) cart
-            => ('A::type, 'N::type) cart)
-           => prop)
- (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
-           => ('A::type, 'N::type) cart)
- (%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
-     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
-      (%i::nat.
-          ($::('A::type, ('M::type, 'N::type) finite_sum) cart
-              => nat => 'A::type)
-           u ((op +::nat => nat => nat) i
-               ((dimindex::('M::type => bool) => nat)
-                 (UNIV::'M::type => bool)))))"
-
-lemma DEF_sndcart: "(op =::(('A::type, ('M::type, 'N::type) finite_sum) cart
-        => ('A::type, 'N::type) cart)
-       => (('A::type, ('M::type, 'N::type) finite_sum) cart
-           => ('A::type, 'N::type) cart)
-          => bool)
- (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
-           => ('A::type, 'N::type) cart)
- (%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
-     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
-      (%i::nat.
-          ($::('A::type, ('M::type, 'N::type) finite_sum) cart
-              => nat => 'A::type)
-           u ((op +::nat => nat => nat) i
-               ((dimindex::('M::type => bool) => nat)
-                 (UNIV::'M::type => bool)))))"
-  by (import hollight DEF_sndcart)
-
-lemma FINITE_SUM_IMAGE: "(op =::(('A::type, 'B::type) finite_sum => bool)
-       => (('A::type, 'B::type) finite_sum => bool) => bool)
- (UNIV::('A::type, 'B::type) finite_sum => bool)
- ((op `::(nat => ('A::type, 'B::type) finite_sum)
-         => (nat => bool) => ('A::type, 'B::type) finite_sum => bool)
-   (mk_finite_sum::nat => ('A::type, 'B::type) finite_sum)
-   ((atLeastAtMost::nat => nat => nat => bool) (1::nat)
-     ((op +::nat => nat => nat)
-       ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))
-       ((dimindex::('B::type => bool) => nat) (UNIV::'B::type => bool)))))"
-  by (import hollight FINITE_SUM_IMAGE)
-
-lemma HAS_SIZE_1: "(HAS_SIZE::(unit => bool) => nat => bool) (UNIV::unit => bool) (1::nat)"
-  by (import hollight HAS_SIZE_1)
-
-typedef (open) N_2 = "{x. x : dotdot (NUMERAL (NUMERAL_BIT1 0))
-         (NUMERAL (NUMERAL_BIT0 (NUMERAL_BIT1 0)))}"  morphisms "dest_auto_define_finite_type_2" "mk_auto_define_finite_type_2"
-  apply (rule light_ex_imp_nonempty[where t="SOME x.
-   x : dotdot (NUMERAL (NUMERAL_BIT1 0))
-        (NUMERAL (NUMERAL_BIT0 (NUMERAL_BIT1 0)))"])
-  by (import hollight TYDEF_2)
-
-syntax
-  dest_auto_define_finite_type_2 :: _ 
-
-syntax
-  mk_auto_define_finite_type_2 :: _ 
-
-lemmas "TYDEF_2_@intern" = typedef_hol2hollight 
-  [where a="a :: N_2" and r=r ,
-   OF type_definition_N_2]
-
-typedef (open) N_3 = "{x. x : dotdot (NUMERAL (NUMERAL_BIT1 0))
-         (NUMERAL (NUMERAL_BIT1 (NUMERAL_BIT1 0)))}"  morphisms "dest_auto_define_finite_type_3" "mk_auto_define_finite_type_3"
-  apply (rule light_ex_imp_nonempty[where t="SOME x.
-   x : dotdot (NUMERAL (NUMERAL_BIT1 0))
-        (NUMERAL (NUMERAL_BIT1 (NUMERAL_BIT1 0)))"])
-  by (import hollight TYDEF_3)
-
-syntax
-  dest_auto_define_finite_type_3 :: _ 
-
-syntax
-  mk_auto_define_finite_type_3 :: _ 
-
-lemmas "TYDEF_3_@intern" = typedef_hol2hollight 
-  [where a="a :: N_3" and r=r ,
-   OF type_definition_N_3]
-
-lemma FINITE_CART: "(op ==>::prop => prop => prop)
- ((all::(nat => prop) => prop)
-   (%i::nat.
-       (op ==>::prop => prop => prop)
-        ((Trueprop::bool => prop)
-          ((op &::bool => bool => bool)
-            ((op <=::nat => nat => bool) (1::nat) i)
-            ((op <=::nat => nat => bool) i
-              ((dimindex::('N::type => bool) => nat)
-                (UNIV::'N::type => bool)))))
-        ((Trueprop::bool => prop)
-          ((finite::('A::type => bool) => bool)
-            ((Collect::('A::type => bool) => 'A::type => bool)
-              (%u::'A::type.
-                  (Ex::('A::type => bool) => bool)
-                   (%x::'A::type.
-                       (op &::bool => bool => bool)
-                        ((P::nat => 'A::type => bool) i x)
-                        ((op =::'A::type => 'A::type => bool) u x))))))))
- ((Trueprop::bool => prop)
-   ((finite::(('A::type, 'N::type) cart => bool) => bool)
-     ((Collect::(('A::type, 'N::type) cart => bool)
-                => ('A::type, 'N::type) cart => bool)
-       (%u::('A::type, 'N::type) cart.
-           (Ex::(('A::type, 'N::type) cart => bool) => bool)
-            (%v::('A::type, 'N::type) cart.
-                (op &::bool => bool => bool)
-                 ((All::(nat => bool) => bool)
-                   (%i::nat.
-                       (op -->::bool => bool => bool)
-                        ((op &::bool => bool => bool)
-                          ((op <=::nat => nat => bool) (1::nat) i)
-                          ((op <=::nat => nat => bool) i
-                            ((dimindex::('N::type => bool) => nat)
-                              (UNIV::'N::type => bool))))
-                        (P i (($::('A::type, 'N::type) cart
-                                  => nat => 'A::type)
-                               v i))))
-                 ((op =::('A::type, 'N::type) cart
-                         => ('A::type, 'N::type) cart => bool)
-                   u v))))))"
-  by (import hollight FINITE_CART)
-
-definition
-  vector :: "'A list => ('A, 'N) cart"  where
-  "(op ==::('A::type list => ('A::type, 'N::type) cart)
-        => ('A::type list => ('A::type, 'N::type) cart) => prop)
- (vector::'A::type list => ('A::type, 'N::type) cart)
- (%u::'A::type list.
-     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
-      (%i::nat.
-          (op !::'A::type list => nat => 'A::type) u
-           ((op -::nat => nat => nat) i (1::nat))))"
-
-lemma DEF_vector: "(op =::('A::type list => ('A::type, 'N::type) cart)
-       => ('A::type list => ('A::type, 'N::type) cart) => bool)
- (vector::'A::type list => ('A::type, 'N::type) cart)
- (%u::'A::type list.
-     (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
-      (%i::nat.
-          (op !::'A::type list => nat => 'A::type) u
-           ((op -::nat => nat => nat) i (1::nat))))"
-  by (import hollight DEF_vector)
-
-definition
-  CASEWISE :: "(('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799)) list
-=> 'q_74840 => 'q_74839 => 'q_74799"  where
-  "CASEWISE ==
-SOME CASEWISE::(('q_74835 => 'q_74839) *
-                ('q_74840 => 'q_74835 => 'q_74799)) list
-               => 'q_74840 => 'q_74839 => 'q_74799.
-   (ALL (f::'q_74840) x::'q_74839.
-       CASEWISE [] f x = (SOME y::'q_74799. True)) &
-   (ALL (h::('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799))
-       (t::(('q_74835 => 'q_74839) *
-            ('q_74840 => 'q_74835 => 'q_74799)) list)
-       (f::'q_74840) x::'q_74839.
-       CASEWISE (h # t) f x =
-       (if EX y::'q_74835. fst h y = x
-        then snd h f (SOME y::'q_74835. fst h y = x) else CASEWISE t f x))"
-
-lemma DEF_CASEWISE: "CASEWISE =
-(SOME CASEWISE::(('q_74835 => 'q_74839) *
-                 ('q_74840 => 'q_74835 => 'q_74799)) list
-                => 'q_74840 => 'q_74839 => 'q_74799.
-    (ALL (f::'q_74840) x::'q_74839.
-        CASEWISE [] f x = (SOME y::'q_74799. True)) &
-    (ALL (h::('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799))
-        (t::(('q_74835 => 'q_74839) *
-             ('q_74840 => 'q_74835 => 'q_74799)) list)
-        (f::'q_74840) x::'q_74839.
-        CASEWISE (h # t) f x =
-        (if EX y::'q_74835. fst h y = x
-         then snd h f (SOME y::'q_74835. fst h y = x) else CASEWISE t f x)))"
-  by (import hollight DEF_CASEWISE)
-
-definition
-  admissible :: "('q_75137 => 'q_75130 => bool)
-=> (('q_75137 => 'q_75133) => 'q_75143 => bool)
-   => ('q_75143 => 'q_75130)
-      => (('q_75137 => 'q_75133) => 'q_75143 => 'q_75138) => bool"  where
-  "admissible ==
-%(u::'q_75137 => 'q_75130 => bool)
-   (ua::('q_75137 => 'q_75133) => 'q_75143 => bool)
-   (ub::'q_75143 => 'q_75130)
-   uc::('q_75137 => 'q_75133) => 'q_75143 => 'q_75138.
-   ALL (f::'q_75137 => 'q_75133) (g::'q_75137 => 'q_75133) a::'q_75143.
-      ua f a & ua g a & (ALL z::'q_75137. u z (ub a) --> f z = g z) -->
-      uc f a = uc g a"
-
-lemma DEF_admissible: "admissible =
-(%(u::'q_75137 => 'q_75130 => bool)
-    (ua::('q_75137 => 'q_75133) => 'q_75143 => bool)
-    (ub::'q_75143 => 'q_75130)
-    uc::('q_75137 => 'q_75133) => 'q_75143 => 'q_75138.
-    ALL (f::'q_75137 => 'q_75133) (g::'q_75137 => 'q_75133) a::'q_75143.
-       ua f a & ua g a & (ALL z::'q_75137. u z (ub a) --> f z = g z) -->
-       uc f a = uc g a)"
-  by (import hollight DEF_admissible)
-
-definition
-  tailadmissible :: "('A => 'A => bool)
-=> (('A => 'B) => 'P => bool)
-   => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool"  where
-  "tailadmissible ==
-%(u::'A => 'A => bool) (ua::('A => 'B) => 'P => bool) (ub::'P => 'A)
-   uc::('A => 'B) => 'P => 'B.
-   EX (P::('A => 'B) => 'P => bool) (G::('A => 'B) => 'P => 'A)
-      H::('A => 'B) => 'P => 'B.
-      (ALL (f::'A => 'B) (a::'P) y::'A.
-          P f a & u y (G f a) --> u y (ub a)) &
-      (ALL (f::'A => 'B) (g::'A => 'B) a::'P.
-          (ALL z::'A. u z (ub a) --> f z = g z) -->
-          P f a = P g a & G f a = G g a & H f a = H g a) &
-      (ALL (f::'A => 'B) a::'P.
-          ua f a --> uc f a = (if P f a then f (G f a) else H f a))"
-
-lemma DEF_tailadmissible: "tailadmissible =
-(%(u::'A => 'A => bool) (ua::('A => 'B) => 'P => bool) (ub::'P => 'A)
-    uc::('A => 'B) => 'P => 'B.
-    EX (P::('A => 'B) => 'P => bool) (G::('A => 'B) => 'P => 'A)
-       H::('A => 'B) => 'P => 'B.
-       (ALL (f::'A => 'B) (a::'P) y::'A.
-           P f a & u y (G f a) --> u y (ub a)) &
-       (ALL (f::'A => 'B) (g::'A => 'B) a::'P.
-           (ALL z::'A. u z (ub a) --> f z = g z) -->
-           P f a = P g a & G f a = G g a & H f a = H g a) &
-       (ALL (f::'A => 'B) a::'P.
-           ua f a --> uc f a = (if P f a then f (G f a) else H f a)))"
-  by (import hollight DEF_tailadmissible)
-
-definition
-  superadmissible :: "('q_75287 => 'q_75287 => bool)
-=> (('q_75287 => 'q_75289) => 'q_75295 => bool)
-   => ('q_75295 => 'q_75287)
-      => (('q_75287 => 'q_75289) => 'q_75295 => 'q_75289) => bool"  where
-  "superadmissible ==
-%(u::'q_75287 => 'q_75287 => bool)
-   (ua::('q_75287 => 'q_75289) => 'q_75295 => bool)
-   (ub::'q_75295 => 'q_75287)
-   uc::('q_75287 => 'q_75289) => 'q_75295 => 'q_75289.
-   admissible u (%(f::'q_75287 => 'q_75289) a::'q_75295. True) ub ua -->
-   tailadmissible u ua ub uc"
-
-lemma DEF_superadmissible: "superadmissible =
-(%(u::'q_75287 => 'q_75287 => bool)
-    (ua::('q_75287 => 'q_75289) => 'q_75295 => bool)
-    (ub::'q_75295 => 'q_75287)
-    uc::('q_75287 => 'q_75289) => 'q_75295 => 'q_75289.
-    admissible u (%(f::'q_75287 => 'q_75289) a::'q_75295. True) ub ua -->
-    tailadmissible u ua ub uc)"
-  by (import hollight DEF_superadmissible)
-
-lemma MATCH_SEQPATTERN: "_MATCH (x::'q_75330)
- (_SEQPATTERN (r::'q_75330 => 'q_75323 => bool)
-   (s::'q_75330 => 'q_75323 => bool)) =
-(if Ex (r x) then _MATCH x r else _MATCH x s)"
-  by (import hollight MATCH_SEQPATTERN)
-
-lemma ADMISSIBLE_CONST: "admissible (u_556::'q_75351 => 'q_75350 => bool)
- (x::('q_75351 => 'q_75352) => 'q_75353 => bool) (xa::'q_75353 => 'q_75350)
- (%f::'q_75351 => 'q_75352. xb::'q_75353 => 'q_75354)"
-  by (import hollight ADMISSIBLE_CONST)
-
-lemma ADMISSIBLE_BASE: "(!!(f::'A => 'B) a::'P.
-    (xa::('A => 'B) => 'P => bool) f a
-    ==> (x::'A => 'A => bool) ((xc::'P => 'A) a) ((xb::'P => 'A) a))
-==> admissible x xa xb (%(f::'A => 'B) x::'P. f (xc x))"
-  by (import hollight ADMISSIBLE_BASE)
-
-lemma ADMISSIBLE_COMB: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'A) (xc::('A => 'B) => 'P => 'C => 'D) &
-admissible x xa xb (xd::('A => 'B) => 'P => 'C)
-==> admissible x xa xb (%(f::'A => 'B) x::'P. xc f x (xd f x))"
-  by (import hollight ADMISSIBLE_COMB)
-
-lemma ADMISSIBLE_RAND: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'A) (xd::('A => 'B) => 'P => 'C)
-==> admissible x xa xb
-     (%(f::'A => 'B) x::'P. (xc::'P => 'C => 'D) x (xd f x))"
-  by (import hollight ADMISSIBLE_RAND)
-
-lemma ADMISSIBLE_LAMBDA: "admissible (x::'A => 'A => bool)
- (%f::'A => 'B.
-     SOME fa::'C * 'P => bool.
-        ALL (u::'C) x::'P. fa (u, x) = (xa::('A => 'B) => 'P => bool) f x)
- (SOME f::'C * 'P => 'A. ALL (u::'C) x::'P. f (u, x) = (xb::'P => 'A) x)
- (%f::'A => 'B.
-     SOME fa::'C * 'P => bool.
-        ALL (u::'C) x::'P.
-           fa (u, x) = (xc::('A => 'B) => 'C => 'P => bool) f u x)
-==> admissible x xa xb (%(f::'A => 'B) (x::'P) u::'C. xc f u x)"
-  by (import hollight ADMISSIBLE_LAMBDA)
-
-lemma ADMISSIBLE_NEST: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'A) (xc::('A => 'B) => 'P => 'A) &
-(ALL (f::'A => 'B) a::'P. xa f a --> x (xc f a) (xb a))
-==> admissible x xa xb (%(f::'A => 'B) x::'P. f (xc f x))"
-  by (import hollight ADMISSIBLE_NEST)
-
-lemma ADMISSIBLE_COND: "admissible (u_556::'q_75688 => 'q_75687 => bool)
- (p::('q_75688 => 'q_75719) => 'P => bool) (s::'P => 'q_75687)
- (P::('q_75688 => 'q_75719) => 'P => bool) &
-admissible u_556 (%(f::'q_75688 => 'q_75719) x::'P. p f x & P f x) s
- (h::('q_75688 => 'q_75719) => 'P => 'q_75744) &
-admissible u_556 (%(f::'q_75688 => 'q_75719) x::'P. p f x & ~ P f x) s
- (k::('q_75688 => 'q_75719) => 'P => 'q_75744)
-==> admissible u_556 p s
-     (%(f::'q_75688 => 'q_75719) x::'P. if P f x then h f x else k f x)"
-  by (import hollight ADMISSIBLE_COND)
-
-lemma ADMISSIBLE_MATCH: "admissible (x::'q_75790 => 'q_75789 => bool)
- (xa::('q_75790 => 'q_75791) => 'P => bool) (xb::'P => 'q_75789)
- (xc::('q_75790 => 'q_75791) => 'P => 'q_75826) &
-admissible x xa xb
- (%(f::'q_75790 => 'q_75791) x::'P.
-     (c::('q_75790 => 'q_75791) => 'P => 'q_75826 => 'q_75823 => bool) f x
-      (xc f x))
-==> admissible x xa xb
-     (%(f::'q_75790 => 'q_75791) x::'P. _MATCH (xc f x) (c f x))"
-  by (import hollight ADMISSIBLE_MATCH)
-
-lemma ADMISSIBLE_SEQPATTERN: "admissible (x::'q_75867 => 'q_75866 => bool)
- (xa::('q_75867 => 'q_75929) => 'P => bool) (xb::'P => 'q_75866)
- (%(f::'q_75867 => 'q_75929) x::'P.
-     Ex ((xc::('q_75867 => 'q_75929) => 'P => 'q_75955 => 'q_75945 => bool)
-          f x ((xe::('q_75867 => 'q_75929) => 'P => 'q_75955) f x))) &
-admissible x
- (%(f::'q_75867 => 'q_75929) x::'P. xa f x & Ex (xc f x (xe f x))) xb
- (%(f::'q_75867 => 'q_75929) x::'P. xc f x (xe f x)) &
-admissible x
- (%(f::'q_75867 => 'q_75929) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb
- (%(f::'q_75867 => 'q_75929) x::'P.
-     (xd::('q_75867 => 'q_75929) => 'P => 'q_75955 => 'q_75945 => bool) f x
-      (xe f x))
-==> admissible x xa xb
-     (%(f::'q_75867 => 'q_75929) x::'P.
-         _SEQPATTERN (xc f x) (xd f x) (xe f x))"
-  by (import hollight ADMISSIBLE_SEQPATTERN)
-
-lemma ADMISSIBLE_UNGUARDED_PATTERN: "admissible (u_556::'q_76041 => 'q_76040 => bool)
- (p::('q_76041 => 'q_76088) => 'P => bool) (s::'P => 'q_76040)
- (pat::('q_76041 => 'q_76088) => 'P => 'q_76121) &
-admissible u_556 p s (e::('q_76041 => 'q_76088) => 'P => 'q_76121) &
-admissible u_556 (%(f::'q_76041 => 'q_76088) x::'P. p f x & pat f x = e f x)
- s (t::('q_76041 => 'q_76088) => 'P => 'q_76128) &
-admissible u_556 (%(f::'q_76041 => 'q_76088) x::'P. p f x & pat f x = e f x)
- s (y::('q_76041 => 'q_76088) => 'P => 'q_76128)
-==> admissible u_556 p s
-     (%(f::'q_76041 => 'q_76088) x::'P.
-         _UNGUARDED_PATTERN (pat f x = e f x) (t f x = y f x))"
-  by (import hollight ADMISSIBLE_UNGUARDED_PATTERN)
-
-lemma ADMISSIBLE_GUARDED_PATTERN: "admissible (u_556::'q_76215 => 'q_76214 => bool)
- (p::('q_76215 => 'q_76292) => 'P => bool) (s::'P => 'q_76214)
- (pat::('q_76215 => 'q_76292) => 'P => 'q_76330) &
-admissible u_556 p s (e::('q_76215 => 'q_76292) => 'P => 'q_76330) &
-admissible u_556
- (%(f::'q_76215 => 'q_76292) x::'P.
-     p f x &
-     pat f x = e f x & (q::('q_76215 => 'q_76292) => 'P => bool) f x)
- s (t::('q_76215 => 'q_76292) => 'P => 'q_76339) &
-admissible u_556 (%(f::'q_76215 => 'q_76292) x::'P. p f x & pat f x = e f x)
- s q &
-admissible u_556
- (%(f::'q_76215 => 'q_76292) x::'P. p f x & pat f x = e f x & q f x) s
- (y::('q_76215 => 'q_76292) => 'P => 'q_76339)
-==> admissible u_556 p s
-     (%(f::'q_76215 => 'q_76292) x::'P.
-         _GUARDED_PATTERN (pat f x = e f x) (q f x) (t f x = y f x))"
-  by (import hollight ADMISSIBLE_GUARDED_PATTERN)
-
-lemma ADMISSIBLE_NSUM: "admissible (x::'B => 'A => bool)
- (%f::'B => 'C.
-     SOME fa::nat * 'P => bool.
-        ALL (k::nat) x::'P.
-           fa (k, x) =
-           ((xd::'P => nat) x <= k &
-            k <= (xe::'P => nat) x & (xa::('B => 'C) => 'P => bool) f x))
- (SOME f::nat * 'P => 'A. ALL (k::nat) x::'P. f (k, x) = (xb::'P => 'A) x)
- (%f::'B => 'C.
-     SOME fa::nat * 'P => nat.
-        ALL (k::nat) x::'P.
-           fa (k, x) = (xc::('B => 'C) => 'P => nat => nat) f x k)
-==> admissible x xa xb (%(f::'B => 'C) x::'P. nsum {xd x..xe x} (xc f x))"
-  by (import hollight ADMISSIBLE_NSUM)
-
-lemma ADMISSIBLE_SUM: "admissible (x::'B => 'A => bool)
- (%f::'B => 'C.
-     SOME fa::nat * 'P => bool.
-        ALL (k::nat) x::'P.
-           fa (k, x) =
-           ((xd::'P => nat) x <= k &
-            k <= (xe::'P => nat) x & (xa::('B => 'C) => 'P => bool) f x))
- (SOME f::nat * 'P => 'A. ALL (k::nat) x::'P. f (k, x) = (xb::'P => 'A) x)
- (%f::'B => 'C.
-     SOME fa::nat * 'P => hollight.real.
-        ALL (k::nat) x::'P.
-           fa (k, x) = (xc::('B => 'C) => 'P => nat => hollight.real) f x k)
-==> admissible x xa xb
-     (%(f::'B => 'C) x::'P. hollight.sum {xd x..xe x} (xc f x))"
-  by (import hollight ADMISSIBLE_SUM)
-
-lemma ADMISSIBLE_MAP: "admissible (x::'A => 'q_76632 => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'q_76632) (xd::('A => 'B) => 'P => 'q_76647 list) &
-admissible x
- (%f::'A => 'B.
-     SOME fa::'q_76647 * 'P => bool.
-        ALL (y::'q_76647) x::'P. fa (y, x) = (xa f x & y : set (xd f x)))
- (SOME f::'q_76647 * 'P => 'q_76632.
-     ALL (y::'q_76647) x::'P. f (y, x) = xb x)
- (%f::'A => 'B.
-     SOME fa::'q_76647 * 'P => 'q_76641.
-        ALL (y::'q_76647) x::'P.
-           fa (y, x) = (xc::('A => 'B) => 'P => 'q_76647 => 'q_76641) f x y)
-==> admissible x xa xb (%(f::'A => 'B) x::'P. map (xc f x) (xd f x))"
-  by (import hollight ADMISSIBLE_MAP)
-
-lemma ADMISSIBLE_MATCH_SEQPATTERN: "admissible (x::'q_76705 => 'q_76704 => bool)
- (xa::('q_76705 => 'q_76770) => 'P => bool) (xb::'P => 'q_76704)
- (%(f::'q_76705 => 'q_76770) x::'P.
-     Ex ((xc::('q_76705 => 'q_76770) => 'P => 'q_76825 => 'q_76794 => bool)
-          f x ((xe::('q_76705 => 'q_76770) => 'P => 'q_76825) f x))) &
-admissible x
- (%(f::'q_76705 => 'q_76770) x::'P. xa f x & Ex (xc f x (xe f x))) xb
- (%(f::'q_76705 => 'q_76770) x::'P. _MATCH (xe f x) (xc f x)) &
-admissible x
- (%(f::'q_76705 => 'q_76770) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb
- (%(f::'q_76705 => 'q_76770) x::'P.
-     _MATCH (xe f x)
-      ((xd::('q_76705 => 'q_76770) => 'P => 'q_76825 => 'q_76794 => bool) f
-        x))
-==> admissible x xa xb
-     (%(x::'q_76705 => 'q_76770) xa::'P.
-         _MATCH (xe x xa) (_SEQPATTERN (xc x xa) (xd x xa)))"
-  by (import hollight ADMISSIBLE_MATCH_SEQPATTERN)
-
-lemma ADMISSIBLE_IMP_SUPERADMISSIBLE: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'A) (xc::('A => 'B) => 'P => 'B)
-==> superadmissible x xa xb xc"
-  by (import hollight ADMISSIBLE_IMP_SUPERADMISSIBLE)
-
-lemma SUPERADMISSIBLE_CONST: "superadmissible (u_556::'q_76904 => 'q_76904 => bool)
- (p::('q_76904 => 'q_76906) => 'q_76905 => bool) (s::'q_76905 => 'q_76904)
- (%f::'q_76904 => 'q_76906. c::'q_76905 => 'q_76906)"
-  by (import hollight SUPERADMISSIBLE_CONST)
-
-lemma SUPERADMISSIBLE_TAIL: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xb::'P => 'A) (xc::('A => 'B) => 'P => 'A) &
-(ALL (f::'A => 'B) a::'P.
-    xa f a --> (ALL y::'A. x y (xc f a) --> x y (xb a)))
-==> superadmissible x xa xb (%(f::'A => 'B) x::'P. f (xc f x))"
-  by (import hollight SUPERADMISSIBLE_TAIL)
-
-lemma SUPERADMISSIBLE_COND: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool)
- (xc::'P => 'A) (xb::('A => 'B) => 'P => bool) &
-superadmissible x (%(f::'A => 'B) x::'P. xa f x & xb f x) xc
- (xd::('A => 'B) => 'P => 'B) &
-superadmissible x (%(f::'A => 'B) x::'P. xa f x & ~ xb f x) xc
- (xe::('A => 'B) => 'P => 'B)
-==> superadmissible x xa xc
-     (%(f::'A => 'B) x::'P. if xb f x then xd f x else xe f x)"
-  by (import hollight SUPERADMISSIBLE_COND)
-
-lemma SUPERADMISSIBLE_MATCH_SEQPATTERN: "admissible (x::'q_77225 => 'q_77225 => bool)
- (xa::('q_77225 => 'q_77341) => 'P => bool) (xb::'P => 'q_77225)
- (%(f::'q_77225 => 'q_77341) x::'P.
-     Ex ((xc::('q_77225 => 'q_77341) => 'P => 'q_77340 => 'q_77341 => bool)
-          f x ((xe::('q_77225 => 'q_77341) => 'P => 'q_77340) f x))) &
-superadmissible x
- (%(f::'q_77225 => 'q_77341) x::'P. xa f x & Ex (xc f x (xe f x))) xb
- (%(f::'q_77225 => 'q_77341) x::'P. _MATCH (xe f x) (xc f x)) &
-superadmissible x
- (%(f::'q_77225 => 'q_77341) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb
- (%(f::'q_77225 => 'q_77341) x::'P.
-     _MATCH (xe f x)
-      ((xd::('q_77225 => 'q_77341) => 'P => 'q_77340 => 'q_77341 => bool) f
-        x))
-==> superadmissible x xa xb
-     (%(x::'q_77225 => 'q_77341) xa::'P.
-         _MATCH (xe x xa) (_SEQPATTERN (xc x xa) (xd x xa)))"
-  by (import hollight SUPERADMISSIBLE_MATCH_SEQPATTERN)
-
-lemma SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN: "(ALL (f::'A => 'B) (a::'P) (t::'Q) u::'Q.
-    (p::('A => 'B) => 'P => bool) f a &
-    (pat::'Q => 'D) t = (e::'P => 'D) a & pat u = e a -->
-    (arg::'P => 'Q => 'A) a t = arg a u) &
-(ALL (f::'A => 'B) (a::'P) t::'Q.
-    p f a & pat t = e a -->
-    (ALL y::'A.
-        (u_556::'A => 'A => bool) y (arg a t) -->
-        u_556 y ((s::'P => 'A) a)))
-==> superadmissible u_556 p s
-     (%(f::'A => 'B) x::'P.
-         _MATCH (e x)
-          (%(u::'D) v::'B.
-              EX t::'Q. _UNGUARDED_PATTERN (pat t = u) (f (arg x t) = v)))"
-  by (import hollight SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN)
-
-lemma SUPERADMISSIBLE_MATCH_GUARDED_PATTERN: "(ALL (f::'A => 'B) (a::'P) (t::'Q) u::'Q.
-    (p::('A => 'B) => 'P => bool) f a &
-    (pat::'Q => 'D) t = (e::'P => 'D) a &
-    (q::'P => 'Q => bool) a t & pat u = e a & q a u -->
-    (arg::'P => 'Q => 'A) a t = arg a u) &
-(ALL (f::'A => 'B) (a::'P) t::'Q.
-    p f a & q a t & pat t = e a -->
-    (ALL y::'A.
-        (u_556::'A => 'A => bool) y (arg a t) -->
-        u_556 y ((s::'P => 'A) a)))
-==> superadmissible u_556 p s
-     (%(f::'A => 'B) x::'P.
-         _MATCH (e x)
-          (%(u::'D) v::'B.
-              EX t::'Q.
-                 _GUARDED_PATTERN (pat t = u) (q x t) (f (arg x t) = v)))"
-  by (import hollight SUPERADMISSIBLE_MATCH_GUARDED_PATTERN)
-
-lemma cth: "[| !!(c::'q_78547) (x::'A) y::'A.
-      (p1::'A => 'q_78536) x = (p1'::'A => 'q_78536) y
-      ==> (p2::'q_78547 => 'A => 'q_78541) c x =
-          (p2'::'q_78547 => 'A => 'q_78541) c y;
-   p1' (x::'A) = p1 (y::'A) |]
-==> p2' (c::'q_78547) x = p2 c y"
-  by (import hollight cth)
-
-lemma SUPERADMISSIBLE_T: "superadmissible (u_556::'q_78694 => 'q_78694 => bool)
- (%(f::'q_78694 => 'q_78696) x::'q_78700. True) (s::'q_78700 => 'q_78694)
- (t::('q_78694 => 'q_78696) => 'q_78700 => 'q_78696) =
-tailadmissible u_556 (%(f::'q_78694 => 'q_78696) x::'q_78700. True) s t"
-  by (import hollight SUPERADMISSIBLE_T)
-
-lemma elemma0: "(ALL z::'q_78985 * 'q_78984.
-    (f::'q_78985 * 'q_78984 => 'q_78975) z =
-    (g::'q_78985 * 'q_78984 => 'q_78975) z) =
-(ALL (x::'q_78985) y::'q_78984. f (x, y) = g (x, y)) &
-(P::'q_79002 * 'q_79001 => 'q_78994) =
-(SOME f::'q_79002 * 'q_79001 => 'q_78994.
-    ALL (x::'q_79002) y::'q_79001. f (x, y) = P (x, y))"
-  by (import hollight elemma0)
-
-;end_setup
-
-end
-
--- a/src/HOL/Import/HOL_Light/Generated/hollight.imp	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2335 +0,0 @@
-import
-
-import_segment "hollight"
-
-def_maps
-  "vector" > "vector_def"
-  "treal_of_num" > "treal_of_num_def"
-  "treal_neg" > "treal_neg_def"
-  "treal_mul" > "treal_mul_def"
-  "treal_le" > "treal_le_def"
-  "treal_inv" > "treal_inv_def"
-  "treal_eq" > "treal_eq_def"
-  "treal_add" > "treal_add_def"
-  "tailadmissible" > "tailadmissible_def"
-  "support" > "support_def"
-  "superadmissible" > "superadmissible_def"
-  "sum" > "sum_def"
-  "sndcart" > "sndcart_def"
-  "rem" > "rem_def"
-  "real_sub" > "real_sub_def"
-  "real_sgn" > "real_sgn_def"
-  "real_pow" > "real_pow_def"
-  "real_of_num" > "real_of_num_def"
-  "real_neg" > "real_neg_def"
-  "real_mul" > "real_mul_def"
-  "real_mod" > "real_mod_def"
-  "real_min" > "real_min_def"
-  "real_max" > "real_max_def"
-  "real_lt" > "real_lt_def"
-  "real_le" > "real_le_def"
-  "real_inv" > "real_inv_def"
-  "real_gt" > "real_gt_def"
-  "real_ge" > "real_ge_def"
-  "real_div" > "real_div_def"
-  "real_add" > "real_add_def"
-  "real_abs" > "real_abs_def"
-  "pastecart" > "pastecart_def"
-  "pairwise" > "pairwise_def"
-  "num_of_int" > "num_of_int_def"
-  "num_mod" > "num_mod_def"
-  "num_gcd" > "num_gcd_def"
-  "num_divides" > "num_divides_def"
-  "num_coprime" > "num_coprime_def"
-  "nsum" > "nsum_def"
-  "neutral" > "neutral_def"
-  "nadd_rinv" > "nadd_rinv_def"
-  "nadd_of_num" > "nadd_of_num_def"
-  "nadd_mul" > "nadd_mul_def"
-  "nadd_le" > "nadd_le_def"
-  "nadd_inv" > "nadd_inv_def"
-  "nadd_eq" > "nadd_eq_def"
-  "nadd_add" > "nadd_add_def"
-  "monoidal" > "monoidal_def"
-  "minimal" > "minimal_def"
-  "lambda" > "lambda_def"
-  "iterate" > "iterate_def"
-  "is_nadd" > "is_nadd_def"
-  "integer" > "integer_def"
-  "int_sub" > "int_sub_def"
-  "int_sgn" > "int_sgn_def"
-  "int_pow" > "int_pow_def"
-  "int_of_num" > "int_of_num_def"
-  "int_neg" > "int_neg_def"
-  "int_mul" > "int_mul_def"
-  "int_mod" > "int_mod_def"
-  "int_min" > "int_min_def"
-  "int_max" > "int_max_def"
-  "int_lt" > "int_lt_def"
-  "int_le" > "int_le_def"
-  "int_gt" > "int_gt_def"
-  "int_ge" > "int_ge_def"
-  "int_gcd" > "int_gcd_def"
-  "int_divides" > "int_divides_def"
-  "int_coprime" > "int_coprime_def"
-  "int_add" > "int_add_def"
-  "int_abs" > "int_abs_def"
-  "hreal_of_num" > "hreal_of_num_def"
-  "hreal_mul" > "hreal_mul_def"
-  "hreal_le" > "hreal_le_def"
-  "hreal_inv" > "hreal_inv_def"
-  "hreal_add" > "hreal_add_def"
-  "fstcart" > "fstcart_def"
-  "eqeq" > "eqeq_def"
-  "div" > "div_def"
-  "dist" > "dist_def"
-  "dimindex" > "dimindex_def"
-  "admissible" > "admissible_def"
-  "_UNGUARDED_PATTERN" > "_UNGUARDED_PATTERN_def"
-  "_SEQPATTERN" > "_SEQPATTERN_def"
-  "_MATCH" > "_MATCH_def"
-  "_GUARDED_PATTERN" > "_GUARDED_PATTERN_def"
-  "_FUNCTION" > "_FUNCTION_def"
-  "_FALSITY_" > "_FALSITY__def"
-  "_11937" > "_11937_def"
-  "ZRECSPACE" > "ZRECSPACE_def"
-  "ZCONSTR" > "ZCONSTR_def"
-  "ZBOT" > "ZBOT_def"
-  "UNCURRY" > "UNCURRY_def"
-  "SURJ" > "SURJ_def"
-  "SING" > "SING_def"
-  "REST" > "REST_def"
-  "PASSOC" > "PASSOC_def"
-  "PAIRWISE" > "PAIRWISE_def"
-  "NUM_REP" > "NUM_REP_def"
-  "NUMSUM" > "NUMSUM_def"
-  "NUMSND" > "NUMSND_def"
-  "NUMRIGHT" > "NUMRIGHT_def"
-  "NUMPAIR" > "NUMPAIR_def"
-  "NUMLEFT" > "NUMLEFT_def"
-  "NUMFST" > "NUMFST_def"
-  "LET_END" > "LET_END_def"
-  "ITSET" > "ITSET_def"
-  "ISO" > "ISO_def"
-  "INJP" > "INJP_def"
-  "INJN" > "INJN_def"
-  "INJF" > "INJF_def"
-  "INJA" > "INJA_def"
-  "INJ" > "INJ_def"
-  "IND_SUC" > "IND_SUC_def"
-  "IND_0" > "IND_0_def"
-  "HAS_SIZE" > "HAS_SIZE_def"
-  "FNIL" > "FNIL_def"
-  "FINREC" > "FINREC_def"
-  "FCONS" > "FCONS_def"
-  "DECIMAL" > "DECIMAL_def"
-  "CROSS" > "CROSS_def"
-  "COUNTABLE" > "COUNTABLE_def"
-  "CONSTR" > "CONSTR_def"
-  "CASEWISE" > "CASEWISE_def"
-  "CARD" > "CARD_def"
-  "BOTTOM" > "BOTTOM_def"
-  "BIJ" > "BIJ_def"
-  "ASCII" > "ASCII_def"
-  ">_c" > ">_c_def"
-  ">=_c" > ">=_c_def"
-  "=_c" > "=_c_def"
-  "<_c" > "<_c_def"
-  "<=_c" > "<=_c_def"
-  "$" > "$_def"
-
-type_maps
-  "sum" > "Sum_Type.sum"
-  "recspace" > "HOLLight.hollight.recspace"
-  "real" > "HOLLight.hollight.real"
-  "prod" > "Product_Type.prod"
-  "option" > "Datatype.option"
-  "num" > "Nat.nat"
-  "nadd" > "HOLLight.hollight.nadd"
-  "list" > "List.list"
-  "int" > "HOLLight.hollight.int"
-  "ind" > "Nat.ind"
-  "hreal" > "HOLLight.hollight.hreal"
-  "fun" > "fun"
-  "finite_sum" > "HOLLight.hollight.finite_sum"
-  "finite_image" > "HOLLight.hollight.finite_image"
-  "char" > "HOLLight.hollight.char"
-  "cart" > "HOLLight.hollight.cart"
-  "bool" > "HOL.bool"
-  "N_3" > "HOLLight.hollight.N_3"
-  "N_2" > "HOLLight.hollight.N_2"
-  "N_1" > "Product_Type.unit"
-
-const_maps
-  "~" > "HOL.Not"
-  "vector" > "HOLLight.hollight.vector"
-  "treal_of_num" > "HOLLight.hollight.treal_of_num"
-  "treal_neg" > "HOLLight.hollight.treal_neg"
-  "treal_mul" > "HOLLight.hollight.treal_mul"
-  "treal_le" > "HOLLight.hollight.treal_le"
-  "treal_inv" > "HOLLight.hollight.treal_inv"
-  "treal_eq" > "HOLLight.hollight.treal_eq"
-  "treal_add" > "HOLLight.hollight.treal_add"
-  "tailadmissible" > "HOLLight.hollight.tailadmissible"
-  "support" > "HOLLight.hollight.support"
-  "superadmissible" > "HOLLight.hollight.superadmissible"
-  "sum" > "HOLLight.hollight.sum"
-  "sndcart" > "HOLLight.hollight.sndcart"
-  "set_of_list" > "List.set"
-  "rem" > "HOLLight.hollight.rem"
-  "real_sub" > "HOLLight.hollight.real_sub"
-  "real_sgn" > "HOLLight.hollight.real_sgn"
-  "real_pow" > "HOLLight.hollight.real_pow"
-  "real_of_num" > "HOLLight.hollight.real_of_num"
-  "real_neg" > "HOLLight.hollight.real_neg"
-  "real_mul" > "HOLLight.hollight.real_mul"
-  "real_mod" > "HOLLight.hollight.real_mod"
-  "real_min" > "HOLLight.hollight.real_min"
-  "real_max" > "HOLLight.hollight.real_max"
-  "real_lt" > "HOLLight.hollight.real_lt"
-  "real_le" > "HOLLight.hollight.real_le"
-  "real_inv" > "HOLLight.hollight.real_inv"
-  "real_gt" > "HOLLight.hollight.real_gt"
-  "real_ge" > "HOLLight.hollight.real_ge"
-  "real_div" > "HOLLight.hollight.real_div"
-  "real_add" > "HOLLight.hollight.real_add"
-  "real_abs" > "HOLLight.hollight.real_abs"
-  "pastecart" > "HOLLight.hollight.pastecart"
-  "pairwise" > "HOLLight.hollight.pairwise"
-  "one" > "Product_Type.Unity"
-  "o" > "Fun.comp"
-  "num_of_int" > "HOLLight.hollight.num_of_int"
-  "num_mod" > "HOLLight.hollight.num_mod"
-  "num_gcd" > "HOLLight.hollight.num_gcd"
-  "num_divides" > "HOLLight.hollight.num_divides"
-  "num_coprime" > "HOLLight.hollight.num_coprime"
-  "nsum" > "HOLLight.hollight.nsum"
-  "neutral" > "HOLLight.hollight.neutral"
-  "nadd_rinv" > "HOLLight.hollight.nadd_rinv"
-  "nadd_of_num" > "HOLLight.hollight.nadd_of_num"
-  "nadd_mul" > "HOLLight.hollight.nadd_mul"
-  "nadd_le" > "HOLLight.hollight.nadd_le"
-  "nadd_inv" > "HOLLight.hollight.nadd_inv"
-  "nadd_eq" > "HOLLight.hollight.nadd_eq"
-  "nadd_add" > "HOLLight.hollight.nadd_add"
-  "monoidal" > "HOLLight.hollight.monoidal"
-  "mk_pair" > "Product_Type.Pair_Rep"
-  "mk_num" > "Fun.id"
-  "minimal" > "HOLLight.hollight.minimal"
-  "list_EX" > "List.list_ex"
-  "list_ALL" > "List.list_all"
-  "lambda" > "HOLLight.hollight.lambda"
-  "iterate" > "HOLLight.hollight.iterate"
-  "is_nadd" > "HOLLight.hollight.is_nadd"
-  "integer" > "HOLLight.hollight.integer"
-  "int_sub" > "HOLLight.hollight.int_sub"
-  "int_sgn" > "HOLLight.hollight.int_sgn"
-  "int_pow" > "HOLLight.hollight.int_pow"
-  "int_of_num" > "HOLLight.hollight.int_of_num"
-  "int_neg" > "HOLLight.hollight.int_neg"
-  "int_mul" > "HOLLight.hollight.int_mul"
-  "int_mod" > "HOLLight.hollight.int_mod"
-  "int_min" > "HOLLight.hollight.int_min"
-  "int_max" > "HOLLight.hollight.int_max"
-  "int_lt" > "HOLLight.hollight.int_lt"
-  "int_le" > "HOLLight.hollight.int_le"
-  "int_gt" > "HOLLight.hollight.int_gt"
-  "int_ge" > "HOLLight.hollight.int_ge"
-  "int_gcd" > "HOLLight.hollight.int_gcd"
-  "int_divides" > "HOLLight.hollight.int_divides"
-  "int_coprime" > "HOLLight.hollight.int_coprime"
-  "int_add" > "HOLLight.hollight.int_add"
-  "int_abs" > "HOLLight.hollight.int_abs"
-  "hreal_of_num" > "HOLLight.hollight.hreal_of_num"
-  "hreal_mul" > "HOLLight.hollight.hreal_mul"
-  "hreal_le" > "HOLLight.hollight.hreal_le"
-  "hreal_inv" > "HOLLight.hollight.hreal_inv"
-  "hreal_add" > "HOLLight.hollight.hreal_add"
-  "fstcart" > "HOLLight.hollight.fstcart"
-  "eqeq" > "HOLLight.hollight.eqeq"
-  "div" > "HOLLight.hollight.div"
-  "dist" > "HOLLight.hollight.dist"
-  "dimindex" > "HOLLight.hollight.dimindex"
-  "admissible" > "HOLLight.hollight.admissible"
-  "_UNGUARDED_PATTERN" > "HOLLight.hollight._UNGUARDED_PATTERN"
-  "_SEQPATTERN" > "HOLLight.hollight._SEQPATTERN"
-  "_MATCH" > "HOLLight.hollight._MATCH"
-  "_GUARDED_PATTERN" > "HOLLight.hollight._GUARDED_PATTERN"
-  "_FUNCTION" > "HOLLight.hollight._FUNCTION"
-  "_FALSITY_" > "HOLLight.hollight._FALSITY_"
-  "_11937" > "HOLLight.hollight._11937"
-  "_0" > "Groups.zero_class.zero" :: "nat"
-  "\\/" > "HOL.disj"
-  "ZRECSPACE" > "HOLLight.hollight.ZRECSPACE"
-  "ZIP" > "List.zip"
-  "ZCONSTR" > "HOLLight.hollight.ZCONSTR"
-  "ZBOT" > "HOLLight.hollight.ZBOT"
-  "WF" > "Wellfounded.wfP"
-  "UNIV" > "Orderings.top_class.top" :: "'a => bool"
-  "UNIONS" > "Complete_Lattices.Sup_class.Sup" :: "(('a => bool) => bool) => 'a => bool"
-  "UNION" > "Lattices.sup_class.sup" :: "('a => bool) => ('a => bool) => 'a => bool"
-  "UNCURRY" > "HOLLight.hollight.UNCURRY"
-  "TL" > "List.tl"
-  "T" > "HOL.True"
-  "SURJ" > "HOLLight.hollight.SURJ"
-  "SUC" > "Nat.Suc"
-  "SUBSET" > "Orderings.ord_class.less_eq" :: "('a => bool) => ('a => bool) => bool"
-  "SOME" > "Datatype.Some"
-  "SND" > "Product_Type.snd"
-  "SING" > "HOLLight.hollight.SING"
-  "SETSPEC" > "HOLLightCompat.SETSPEC"
-  "REVERSE" > "List.rev"
-  "REST" > "HOLLight.hollight.REST"
-  "REPLICATE" > "List.replicate"
-  "PSUBSET" > "Orderings.ord_class.less" :: "('a => bool) => ('a => bool) => bool"
-  "PRE" > "HOLLightCompat.Pred"
-  "PASSOC" > "HOLLight.hollight.PASSOC"
-  "PAIRWISE" > "HOLLight.hollight.PAIRWISE"
-  "OUTR" > "HOLLightCompat.OUTR"
-  "OUTL" > "HOLLightCompat.OUTL"
-  "ONTO" > "Fun.surj"
-  "ONE_ONE" > "Fun.inj"
-  "ODD" > "HOLLightCompat.ODD"
-  "NUM_REP" > "HOLLight.hollight.NUM_REP"
-  "NUMSUM" > "HOLLight.hollight.NUMSUM"
-  "NUMSND" > "HOLLight.hollight.NUMSND"
-  "NUMRIGHT" > "HOLLight.hollight.NUMRIGHT"
-  "NUMPAIR" > "HOLLight.hollight.NUMPAIR"
-  "NUMLEFT" > "HOLLight.hollight.NUMLEFT"
-  "NUMFST" > "HOLLight.hollight.NUMFST"
-  "NUMERAL" > "HOLLightCompat.NUMERAL"
-  "NULL" > "List.null"
-  "NONE" > "Datatype.None"
-  "NIL" > "List.list.Nil"
-  "MOD" > "Divides.div_class.mod" :: "nat => nat => nat"
-  "MIN" > "Orderings.ord_class.min" :: "nat => nat => nat"
-  "MEM" > "HOLLightList.list_mem"
-  "MEASURE" > "HOLLightCompat.MEASURE"
-  "MAX" > "Orderings.ord_class.max" :: "nat => nat => nat"
-  "MAP2" > "HOLLightList.map2"
-  "MAP" > "List.map"
-  "LET_END" > "HOLLight.hollight.LET_END"
-  "LET" > "HOLLightCompat.LET"
-  "LENGTH" > "List.length"
-  "LAST" > "List.last"
-  "ITSET" > "HOLLight.hollight.ITSET"
-  "ITLIST2" > "HOLLightList.fold2"
-  "ITLIST" > "List.foldr"
-  "ISO" > "HOLLight.hollight.ISO"
-  "INTERS" > "Complete_Lattices.Inf_class.Inf" :: "(('a => bool) => bool) => 'a => bool"
-  "INTER" > "Lattices.inf_class.inf" :: "('a => bool) => ('a => bool) => 'a => bool"
-  "INSERT" > "Set.insert"
-  "INR" > "Sum_Type.Inr"
-  "INL" > "Sum_Type.Inl"
-  "INJP" > "HOLLight.hollight.INJP"
-  "INJN" > "HOLLight.hollight.INJN"
-  "INJF" > "HOLLight.hollight.INJF"
-  "INJA" > "HOLLight.hollight.INJA"
-  "INJ" > "HOLLight.hollight.INJ"
-  "INFINITE" > "HOLLightCompat.INFINITE"
-  "IND_SUC" > "HOLLight.hollight.IND_SUC"
-  "IND_0" > "HOLLight.hollight.IND_0"
-  "IN" > "Set.member"
-  "IMAGE" > "Set.image"
-  "I" > "Fun.id"
-  "HD" > "List.hd"
-  "HAS_SIZE" > "HOLLight.hollight.HAS_SIZE"
-  "GSPEC" > "Set.Collect"
-  "GEQ" > "HOL.eq"
-  "GABS" > "Hilbert_Choice.Eps"
-  "FST" > "Product_Type.fst"
-  "FNIL" > "HOLLight.hollight.FNIL"
-  "FINREC" > "HOLLight.hollight.FINREC"
-  "FINITE" > "Finite_Set.finite"
-  "FILTER" > "List.filter"
-  "FCONS" > "HOLLight.hollight.FCONS"
-  "FACT" > "Fact.fact_class.fact" :: "nat => nat"
-  "F" > "HOL.False"
-  "EXP" > "Power.power_class.power" :: "nat => nat => nat"
-  "EVEN" > "Parity.even_odd_class.even" :: "nat => bool"
-  "EMPTY" > "Orderings.bot_class.bot" :: "'a => bool"
-  "EL" > "HOLLightList.list_el"
-  "DIV" > "Divides.div_class.div" :: "nat => nat => nat"
-  "DISJOINT" > "HOLLightCompat.DISJOINT"
-  "DIFF" > "Groups.minus_class.minus" :: "('a => bool) => ('a => bool) => 'a => bool"
-  "DELETE" > "HOLLightCompat.DELETE"
-  "DECIMAL" > "HOLLight.hollight.DECIMAL"
-  "CURRY" > "Product_Type.curry"
-  "CROSS" > "HOLLight.hollight.CROSS"
-  "COUNTABLE" > "HOLLight.hollight.COUNTABLE"
-  "CONSTR" > "HOLLight.hollight.CONSTR"
-  "CONS" > "List.list.Cons"
-  "COND" > "HOL.If"
-  "CHOICE" > "Hilbert_Choice.Eps"
-  "CASEWISE" > "HOLLight.hollight.CASEWISE"
-  "CARD" > "HOLLight.hollight.CARD"
-  "BUTLAST" > "List.butlast"
-  "BOTTOM" > "HOLLight.hollight.BOTTOM"
-  "BIT1" > "HOLLightCompat.NUMERAL_BIT1"
-  "BIT0" > "HOLLightCompat.NUMERAL_BIT0"
-  "BIJ" > "HOLLight.hollight.BIJ"
-  "ASCII" > "HOLLight.hollight.ASCII"
-  "APPEND" > "List.append"
-  "ALL2" > "List.list_all2"
-  "@" > "Hilbert_Choice.Eps"
-  "?!" > "HOL.Ex1"
-  "?" > "HOL.Ex"
-  ">_c" > "HOLLight.hollight.>_c"
-  ">=_c" > "HOLLight.hollight.>=_c"
-  ">=" > "Orderings.ord_class.greater_eq" :: "nat => nat => bool"
-  ">" > "Orderings.ord_class.greater" :: "nat => nat => bool"
-  "=_c" > "HOLLight.hollight.=_c"
-  "==>" > "HOL.implies"
-  "=" > "HOL.eq"
-  "<_c" > "HOLLight.hollight.<_c"
-  "<=_c" > "HOLLight.hollight.<=_c"
-  "<=" > "Orderings.ord_class.less_eq" :: "nat => nat => bool"
-  "<" > "Orderings.ord_class.less" :: "nat => nat => bool"
-  "/\\" > "HOL.conj"
-  ".." > "HOLLightCompat.dotdot"
-  "-" > "Groups.minus_class.minus" :: "nat => nat => nat"
-  "," > "Product_Type.Pair"
-  "+" > "Groups.plus_class.plus" :: "nat => nat => nat"
-  "*" > "Groups.times_class.times" :: "nat => nat => nat"
-  "$" > "HOLLight.hollight.$"
-  "!" > "HOL.All"
-
-const_renames
-  "EX" > "list_EX"
-  "ALL" > "list_ALL"
-  "==" > "eqeq"
-
-thm_maps
-  "vector_def" > "HOLLight.hollight.vector_def"
-  "treal_of_num_def" > "HOLLight.hollight.treal_of_num_def"
-  "treal_neg_def" > "HOLLight.hollight.treal_neg_def"
-  "treal_mul_def" > "HOLLight.hollight.treal_mul_def"
-  "treal_le_def" > "HOLLight.hollight.treal_le_def"
-  "treal_inv_def" > "HOLLight.hollight.treal_inv_def"
-  "treal_eq_def" > "HOLLight.hollight.treal_eq_def"
-  "treal_add_def" > "HOLLight.hollight.treal_add_def"
-  "th_cond" > "HOLLight.hollight.th_cond"
-  "th" > "HOL.eta_contract_eq"
-  "tailadmissible_def" > "HOLLight.hollight.tailadmissible_def"
-  "support_def" > "HOLLight.hollight.support_def"
-  "superadmissible_def" > "HOLLight.hollight.superadmissible_def"
-  "sum_def" > "HOLLight.hollight.sum_def"
-  "string_INFINITE" > "List.infinite_UNIV_listI"
-  "sth" > "HOLLight.hollight.sth"
-  "sndcart_def" > "HOLLight.hollight.sndcart_def"
-  "right_th" > "HOLLight.hollight.right_th"
-  "rem_def" > "HOLLight.hollight.rem_def"
-  "real_sub_def" > "HOLLight.hollight.real_sub_def"
-  "real_sgn_def" > "HOLLight.hollight.real_sgn_def"
-  "real_pow_def" > "HOLLight.hollight.real_pow_def"
-  "real_of_num_def" > "HOLLight.hollight.real_of_num_def"
-  "real_neg_def" > "HOLLight.hollight.real_neg_def"
-  "real_mul_def" > "HOLLight.hollight.real_mul_def"
-  "real_mod_def" > "HOLLight.hollight.real_mod_def"
-  "real_min_def" > "HOLLight.hollight.real_min_def"
-  "real_max_def" > "HOLLight.hollight.real_max_def"
-  "real_lt_def" > "HOLLight.hollight.real_lt_def"
-  "real_le_def" > "HOLLight.hollight.real_le_def"
-  "real_inv_def" > "HOLLight.hollight.real_inv_def"
-  "real_gt_def" > "HOLLight.hollight.real_gt_def"
-  "real_ge_def" > "HOLLight.hollight.real_ge_def"
-  "real_div_def" > "HOLLight.hollight.real_div_def"
-  "real_add_def" > "HOLLight.hollight.real_add_def"
-  "real_abs_def" > "HOLLight.hollight.real_abs_def"
-  "real_INFINITE" > "HOLLight.hollight.real_INFINITE"
-  "pastecart_def" > "HOLLight.hollight.pastecart_def"
-  "pairwise_def" > "HOLLight.hollight.pairwise_def"
-  "pair_RECURSION" > "HOLLight.hollight.pair_RECURSION"
-  "pair_INDUCT" > "Product_Type.prod.induct"
-  "one_axiom" > "HOLLight.hollight.one_axiom"
-  "one_RECURSION" > "HOLLight.hollight.one_RECURSION"
-  "one_INDUCT" > "Product_Type.unit.induct"
-  "one_Axiom" > "HOLLight.hollight.one_Axiom"
-  "one" > "HOLLightCompat.one"
-  "o_THM" > "Fun.comp_def"
-  "o_ASSOC" > "HOLLight.hollight.o_ASSOC"
-  "num_of_int_def" > "HOLLight.hollight.num_of_int_def"
-  "num_mod_def" > "HOLLight.hollight.num_mod_def"
-  "num_gcd_def" > "HOLLight.hollight.num_gcd_def"
-  "num_divides_def" > "HOLLight.hollight.num_divides_def"
-  "num_coprime_def" > "HOLLight.hollight.num_coprime_def"
-  "num_congruent" > "HOLLight.hollight.num_congruent"
-  "num_WOP" > "HOLLight.hollight.num_WOP"
-  "num_WF" > "Nat.nat_less_induct"
-  "num_RECURSION_STD" > "HOLLight.hollight.num_RECURSION_STD"
-  "num_MAX" > "HOLLight.hollight.num_MAX"
-  "num_INFINITE" > "Finite_Set.infinite_UNIV_char_0"
-  "num_INDUCTION" > "Fact.fact_nat.induct"
-  "num_FINITE_AVOID" > "HOLLight.hollight.num_FINITE_AVOID"
-  "num_FINITE" > "HOLLight.hollight.num_FINITE"
-  "num_CASES" > "Nat.nat.nchotomy"
-  "num_Axiom" > "HOLLightCompat.num_Axiom"
-  "nsum_def" > "HOLLight.hollight.nsum_def"
-  "neutral_def" > "HOLLight.hollight.neutral_def"
-  "nadd_rinv_def" > "HOLLight.hollight.nadd_rinv_def"
-  "nadd_of_num_def" > "HOLLight.hollight.nadd_of_num_def"
-  "nadd_mul_def" > "HOLLight.hollight.nadd_mul_def"
-  "nadd_le_def" > "HOLLight.hollight.nadd_le_def"
-  "nadd_inv_def" > "HOLLight.hollight.nadd_inv_def"
-  "nadd_eq_def" > "HOLLight.hollight.nadd_eq_def"
-  "nadd_add_def" > "HOLLight.hollight.nadd_add_def"
-  "monoidal_def" > "HOLLight.hollight.monoidal_def"
-  "minimal_def" > "HOLLight.hollight.minimal_def"
-  "list_INDUCT" > "List.list.induct"
-  "list_CASES" > "List.list.nchotomy"
-  "lambda_def" > "HOLLight.hollight.lambda_def"
-  "iterate_def" > "HOLLight.hollight.iterate_def"
-  "is_nadd_def" > "HOLLight.hollight.is_nadd_def"
-  "is_nadd_0" > "HOLLight.hollight.is_nadd_0"
-  "is_int" > "HOLLight.hollight.is_int"
-  "integer_def" > "HOLLight.hollight.integer_def"
-  "int_sub_th" > "HOLLight.hollight.int_sub_th"
-  "int_sub_def" > "HOLLight.hollight.int_sub_def"
-  "int_sgn_th" > "HOLLight.hollight.int_sgn_th"
-  "int_sgn_def" > "HOLLight.hollight.int_sgn_def"
-  "int_pow_th" > "HOLLight.hollight.int_pow_th"
-  "int_pow_def" > "HOLLight.hollight.int_pow_def"
-  "int_of_num_th" > "HOLLight.hollight.int_of_num_th"
-  "int_of_num_def" > "HOLLight.hollight.int_of_num_def"
-  "int_neg_th" > "HOLLight.hollight.int_neg_th"
-  "int_neg_def" > "HOLLight.hollight.int_neg_def"
-  "int_mul_th" > "HOLLight.hollight.int_mul_th"
-  "int_mul_def" > "HOLLight.hollight.int_mul_def"
-  "int_mod_def" > "HOLLight.hollight.int_mod_def"
-  "int_min_th" > "HOLLight.hollight.int_min_th"
-  "int_min_def" > "HOLLight.hollight.int_min_def"
-  "int_max_th" > "HOLLight.hollight.int_max_th"
-  "int_max_def" > "HOLLight.hollight.int_max_def"
-  "int_lt_def" > "HOLLight.hollight.int_lt_def"
-  "int_le_def" > "HOLLight.hollight.int_le_def"
-  "int_gt_def" > "HOLLight.hollight.int_gt_def"
-  "int_ge_def" > "HOLLight.hollight.int_ge_def"
-  "int_gcd_def" > "HOLLight.hollight.int_gcd_def"
-  "int_eq" > "HOLLight.hollight.int.real_of_int_inject"
-  "int_divides_def" > "HOLLight.hollight.int_divides_def"
-  "int_coprime_def" > "HOLLight.hollight.int_coprime_def"
-  "int_congruent" > "HOLLight.hollight.int_congruent"
-  "int_add_th" > "HOLLight.hollight.int_add_th"
-  "int_add_def" > "HOLLight.hollight.int_add_def"
-  "int_abs_th" > "HOLLight.hollight.int_abs_th"
-  "int_abs_def" > "HOLLight.hollight.int_abs_def"
-  "hreal_of_num_def" > "HOLLight.hollight.hreal_of_num_def"
-  "hreal_mul_def" > "HOLLight.hollight.hreal_mul_def"
-  "hreal_le_def" > "HOLLight.hollight.hreal_le_def"
-  "hreal_inv_def" > "HOLLight.hollight.hreal_inv_def"
-  "hreal_add_def" > "HOLLight.hollight.hreal_add_def"
-  "fstcart_def" > "HOLLight.hollight.fstcart_def"
-  "eqeq_def" > "HOLLight.hollight.eqeq_def"
-  "elemma0" > "HOLLight.hollight.elemma0"
-  "div_def" > "HOLLight.hollight.div_def"
-  "dist_def" > "HOLLight.hollight.dist_def"
-  "dimindex_def" > "HOLLight.hollight.dimindex_def"
-  "dest_int_rep" > "HOLLight.hollight.dest_int_rep"
-  "cth" > "HOLLight.hollight.cth"
-  "bool_RECURSION" > "HOLLight.hollight.bool_RECURSION"
-  "bool_INDUCT" > "Product_Type.bool.induct"
-  "ax__3" > "Importer.INFINITY_AX"
-  "ax__2" > "Hilbert_Choice.someI"
-  "ax__1" > "HOL.eta_contract_eq"
-  "admissible_def" > "HOLLight.hollight.admissible_def"
-  "_UNGUARDED_PATTERN_def" > "HOLLight.hollight._UNGUARDED_PATTERN_def"
-  "_SEQPATTERN_def" > "HOLLight.hollight._SEQPATTERN_def"
-  "_MATCH_def" > "HOLLight.hollight._MATCH_def"
-  "_GUARDED_PATTERN_def" > "HOLLight.hollight._GUARDED_PATTERN_def"
-  "_FUNCTION_def" > "HOLLight.hollight._FUNCTION_def"
-  "_FALSITY__def" > "HOLLight.hollight._FALSITY__def"
-  "_11937_def" > "HOLLight.hollight._11937_def"
-  "ZRECSPACE_def" > "HOLLight.hollight.ZRECSPACE_def"
-  "ZIP" > "HOLLightList.ZIP"
-  "ZCONSTR_def" > "HOLLight.hollight.ZCONSTR_def"
-  "ZCONSTR_ZBOT" > "HOLLight.hollight.ZCONSTR_ZBOT"
-  "ZBOT_def" > "HOLLight.hollight.ZBOT_def"
-  "WLOG_LT" > "HOLLight.hollight.WLOG_LT"
-  "WLOG_LE" > "HOLLight.hollight.WLOG_LE"
-  "WF_num" > "HOLLight.hollight.WF_num"
-  "WF_UREC_WF" > "HOLLight.hollight.WF_UREC_WF"
-  "WF_UREC" > "HOLLight.hollight.WF_UREC"
-  "WF_SUBSET" > "HOLLight.hollight.WF_SUBSET"
-  "WF_REFL" > "HOLLight.hollight.WF_REFL"
-  "WF_REC_num" > "HOLLight.hollight.WF_REC_num"
-  "WF_REC_WF" > "HOLLight.hollight.WF_REC_WF"
-  "WF_REC_TAIL_GENERAL" > "HOLLight.hollight.WF_REC_TAIL_GENERAL"
-  "WF_REC_TAIL" > "HOLLight.hollight.WF_REC_TAIL"
-  "WF_REC_INVARIANT" > "HOLLight.hollight.WF_REC_INVARIANT"
-  "WF_REC" > "HOLLight.hollight.WF_REC"
-  "WF_POINTWISE" > "HOLLight.hollight.WF_POINTWISE"
-  "WF_MEASURE_GEN" > "HOLLight.hollight.WF_MEASURE_GEN"
-  "WF_MEASURE" > "HOLLight.hollight.WF_MEASURE"
-  "WF_LEX_DEPENDENT" > "HOLLight.hollight.WF_LEX_DEPENDENT"
-  "WF_LEX" > "HOLLight.hollight.WF_LEX"
-  "WF_INT_MEASURE_2" > "HOLLight.hollight.WF_INT_MEASURE_2"
-  "WF_INT_MEASURE" > "HOLLight.hollight.WF_INT_MEASURE"
-  "WF_IND" > "HOLLight.hollight.WF_IND"
-  "WF_FINITE" > "HOLLight.hollight.WF_FINITE"
-  "WF_FALSE" > "Wellfounded.wfP_empty"
-  "WF_EREC" > "HOLLight.hollight.WF_EREC"
-  "WF_EQ" > "HOLLight.hollight.WF_EQ"
-  "WF_DCHAIN" > "HOLLight.hollight.WF_DCHAIN"
-  "UNWIND_THM2" > "HOL.simp_thms_39"
-  "UNWIND_THM1" > "HOL.simp_thms_40"
-  "UNIV_SUBSET" > "Orderings.top_class.top_unique"
-  "UNIV_NOT_EMPTY" > "Set.UNIV_not_empty"
-  "UNIQUE_SKOLEM_THM" > "HOL.choice_eq"
-  "UNIQUE_SKOLEM_ALT" > "HOLLight.hollight.UNIQUE_SKOLEM_ALT"
-  "UNION_UNIV" > "HOLLight.hollight.UNION_UNIV"
-  "UNION_SUBSET" > "Lattices.semilattice_sup_class.le_sup_iff"
-  "UNION_OVER_INTER" > "Lattices.distrib_lattice_class.distrib_3"
-  "UNION_IDEMPOT" > "Big_Operators.lattice_class.Sup_fin.idem"
-  "UNION_EMPTY" > "HOLLight.hollight.UNION_EMPTY"
-  "UNION_COMM" > "Lattices.lattice_class.inf_sup_aci_5"
-  "UNION_ASSOC" > "Lattices.lattice_class.inf_sup_aci_6"
-  "UNION_ACI" > "HOLLight.hollight.UNION_ACI"
-  "UNIONS_UNION" > "Complete_Lattices.Union_Un_distrib"
-  "UNIONS_SUBSET" > "HOLLight.hollight.UNIONS_SUBSET"
-  "UNIONS_INTERS" > "HOLLight.hollight.UNIONS_INTERS"
-  "UNIONS_INSERT" > "Complete_Lattices.Union_insert"
-  "UNIONS_IMAGE" > "HOLLight.hollight.UNIONS_IMAGE"
-  "UNIONS_GSPEC" > "HOLLight.hollight.UNIONS_GSPEC"
-  "UNIONS_2" > "Complete_Lattices.Un_eq_Union"
-  "UNIONS_1" > "Complete_Lattices.complete_lattice_class.Sup_singleton"
-  "UNIONS_0" > "Complete_Lattices.Union_empty"
-  "UNCURRY_def" > "HOLLight.hollight.UNCURRY_def"
-  "TYDEF_recspace" > "HOLLight.hollight.TYDEF_recspace"
-  "TYDEF_real" > "HOLLight.hollight.TYDEF_real"
-  "TYDEF_nadd" > "HOLLight.hollight.TYDEF_nadd"
-  "TYDEF_int" > "HOLLight.hollight.TYDEF_int"
-  "TYDEF_hreal" > "HOLLight.hollight.TYDEF_hreal"
-  "TYDEF_finite_sum" > "HOLLight.hollight.TYDEF_finite_sum"
-  "TYDEF_finite_image" > "HOLLight.hollight.TYDEF_finite_image"
-  "TYDEF_char" > "HOLLight.hollight.TYDEF_char"
-  "TYDEF_cart" > "HOLLight.hollight.TYDEF_cart"
-  "TYDEF_3" > "HOLLight.hollight.TYDEF_3"
-  "TYDEF_2" > "HOLLight.hollight.TYDEF_2"
-  "TWO" > "HOLLight.hollight.TWO"
-  "TRIV_OR_FORALL_THM" > "HOLLight.hollight.TRIV_OR_FORALL_THM"
-  "TRIV_FORALL_OR_THM" > "HOLLight.hollight.TRIV_FORALL_OR_THM"
-  "TRIV_FORALL_IMP_THM" > "HOLLight.hollight.TRIV_FORALL_IMP_THM"
-  "TRIV_EXISTS_IMP_THM" > "HOLLight.hollight.TRIV_EXISTS_IMP_THM"
-  "TRIV_EXISTS_AND_THM" > "HOLLight.hollight.TRIV_EXISTS_AND_THM"
-  "TRIV_AND_EXISTS_THM" > "HOLLight.hollight.TRIV_AND_EXISTS_THM"
-  "TREAL_OF_NUM_WELLDEF" > "HOLLight.hollight.TREAL_OF_NUM_WELLDEF"
-  "TREAL_OF_NUM_MUL" > "HOLLight.hollight.TREAL_OF_NUM_MUL"
-  "TREAL_OF_NUM_LE" > "HOLLight.hollight.TREAL_OF_NUM_LE"
-  "TREAL_OF_NUM_EQ" > "HOLLight.hollight.TREAL_OF_NUM_EQ"
-  "TREAL_OF_NUM_ADD" > "HOLLight.hollight.TREAL_OF_NUM_ADD"
-  "TREAL_NEG_WELLDEF" > "HOLLight.hollight.TREAL_NEG_WELLDEF"
-  "TREAL_MUL_WELLDEFR" > "HOLLight.hollight.TREAL_MUL_WELLDEFR"
-  "TREAL_MUL_WELLDEF" > "HOLLight.hollight.TREAL_MUL_WELLDEF"
-  "TREAL_MUL_SYM_EQ" > "HOLLight.hollight.TREAL_MUL_SYM_EQ"
-  "TREAL_MUL_SYM" > "HOLLight.hollight.TREAL_MUL_SYM"
-  "TREAL_MUL_LINV" > "HOLLight.hollight.TREAL_MUL_LINV"
-  "TREAL_MUL_LID" > "HOLLight.hollight.TREAL_MUL_LID"
-  "TREAL_MUL_ASSOC" > "HOLLight.hollight.TREAL_MUL_ASSOC"
-  "TREAL_LE_WELLDEF" > "HOLLight.hollight.TREAL_LE_WELLDEF"
-  "TREAL_LE_TRANS" > "HOLLight.hollight.TREAL_LE_TRANS"
-  "TREAL_LE_TOTAL" > "HOLLight.hollight.TREAL_LE_TOTAL"
-  "TREAL_LE_REFL" > "HOLLight.hollight.TREAL_LE_REFL"
-  "TREAL_LE_MUL" > "HOLLight.hollight.TREAL_LE_MUL"
-  "TREAL_LE_LADD_IMP" > "HOLLight.hollight.TREAL_LE_LADD_IMP"
-  "TREAL_LE_ANTISYM" > "HOLLight.hollight.TREAL_LE_ANTISYM"
-  "TREAL_INV_WELLDEF" > "HOLLight.hollight.TREAL_INV_WELLDEF"
-  "TREAL_INV_0" > "HOLLight.hollight.TREAL_INV_0"
-  "TREAL_EQ_TRANS" > "HOLLight.hollight.TREAL_EQ_TRANS"
-  "TREAL_EQ_SYM" > "HOLLight.hollight.TREAL_EQ_SYM"
-  "TREAL_EQ_REFL" > "HOLLight.hollight.TREAL_EQ_REFL"
-  "TREAL_EQ_IMP_LE" > "HOLLight.hollight.TREAL_EQ_IMP_LE"
-  "TREAL_EQ_AP" > "HOLLight.hollight.TREAL_EQ_AP"
-  "TREAL_ADD_WELLDEFR" > "HOLLight.hollight.TREAL_ADD_WELLDEFR"
-  "TREAL_ADD_WELLDEF" > "HOLLight.hollight.TREAL_ADD_WELLDEF"
-  "TREAL_ADD_SYM_EQ" > "HOLLight.hollight.TREAL_ADD_SYM_EQ"
-  "TREAL_ADD_SYM" > "HOLLight.hollight.TREAL_ADD_SYM"
-  "TREAL_ADD_LINV" > "HOLLight.hollight.TREAL_ADD_LINV"
-  "TREAL_ADD_LID" > "HOLLight.hollight.TREAL_ADD_LID"
-  "TREAL_ADD_LDISTRIB" > "HOLLight.hollight.TREAL_ADD_LDISTRIB"
-  "TREAL_ADD_ASSOC" > "HOLLight.hollight.TREAL_ADD_ASSOC"
-  "TRANSITIVE_STEPWISE_LT_EQ" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LT_EQ"
-  "TRANSITIVE_STEPWISE_LT" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LT"
-  "TRANSITIVE_STEPWISE_LE_EQ" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LE_EQ"
-  "TRANSITIVE_STEPWISE_LE" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LE"
-  "TOPOLOGICAL_SORT" > "HOLLight.hollight.TOPOLOGICAL_SORT"
-  "SWAP_FORALL_THM" > "HOL.all_comm"
-  "SWAP_EXISTS_THM" > "HOL.ex_comm"
-  "SURJ_def" > "HOLLight.hollight.SURJ_def"
-  "SURJECTIVE_RIGHT_INVERSE" > "HOLLight.hollight.SURJECTIVE_RIGHT_INVERSE"
-  "SURJECTIVE_ON_RIGHT_INVERSE" > "HOLLight.hollight.SURJECTIVE_ON_RIGHT_INVERSE"
-  "SURJECTIVE_ON_IMAGE" > "HOLLight.hollight.SURJECTIVE_ON_IMAGE"
-  "SURJECTIVE_MAP" > "HOLLightList.SURJECTIVE_MAP"
-  "SURJECTIVE_IMAGE_THM" > "HOLLight.hollight.SURJECTIVE_IMAGE_THM"
-  "SURJECTIVE_IMAGE_EQ" > "HOLLight.hollight.SURJECTIVE_IMAGE_EQ"
-  "SURJECTIVE_IMAGE" > "HOLLight.hollight.SURJECTIVE_IMAGE"
-  "SURJECTIVE_IFF_INJECTIVE_GEN" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE_GEN"
-  "SURJECTIVE_IFF_INJECTIVE" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE"
-  "SURJECTIVE_FORALL_THM" > "HOLLight.hollight.SURJECTIVE_FORALL_THM"
-  "SURJECTIVE_EXISTS_THM" > "HOLLight.hollight.SURJECTIVE_EXISTS_THM"
-  "SUPPORT_SUPPORT" > "HOLLight.hollight.SUPPORT_SUPPORT"
-  "SUPPORT_SUBSET" > "HOLLight.hollight.SUPPORT_SUBSET"
-  "SUPPORT_EMPTY" > "HOLLight.hollight.SUPPORT_EMPTY"
-  "SUPPORT_DELTA" > "HOLLight.hollight.SUPPORT_DELTA"
-  "SUPPORT_CLAUSES" > "HOLLight.hollight.SUPPORT_CLAUSES"
-  "SUPERADMISSIBLE_TAIL" > "HOLLight.hollight.SUPERADMISSIBLE_TAIL"
-  "SUPERADMISSIBLE_T" > "HOLLight.hollight.SUPERADMISSIBLE_T"
-  "SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN"
-  "SUPERADMISSIBLE_MATCH_SEQPATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_SEQPATTERN"
-  "SUPERADMISSIBLE_MATCH_GUARDED_PATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_GUARDED_PATTERN"
-  "SUPERADMISSIBLE_CONST" > "HOLLight.hollight.SUPERADMISSIBLE_CONST"
-  "SUPERADMISSIBLE_COND" > "HOLLight.hollight.SUPERADMISSIBLE_COND"
-  "SUM_ZERO_EXISTS" > "HOLLight.hollight.SUM_ZERO_EXISTS"
-  "SUM_UNION_RZERO" > "HOLLight.hollight.SUM_UNION_RZERO"
-  "SUM_UNION_NONZERO" > "HOLLight.hollight.SUM_UNION_NONZERO"
-  "SUM_UNION_LZERO" > "HOLLight.hollight.SUM_UNION_LZERO"
-  "SUM_UNION_EQ" > "HOLLight.hollight.SUM_UNION_EQ"
-  "SUM_UNIONS_NONZERO" > "HOLLight.hollight.SUM_UNIONS_NONZERO"
-  "SUM_UNION" > "HOLLight.hollight.SUM_UNION"
-  "SUM_TRIV_NUMSEG" > "HOLLight.hollight.SUM_TRIV_NUMSEG"
-  "SUM_SWAP_NUMSEG" > "HOLLight.hollight.SUM_SWAP_NUMSEG"
-  "SUM_SWAP" > "HOLLight.hollight.SUM_SWAP"
-  "SUM_SUPPORT" > "HOLLight.hollight.SUM_SUPPORT"
-  "SUM_SUPERSET" > "HOLLight.hollight.SUM_SUPERSET"
-  "SUM_SUM_RESTRICT" > "HOLLight.hollight.SUM_SUM_RESTRICT"
-  "SUM_SUM_PRODUCT" > "HOLLight.hollight.SUM_SUM_PRODUCT"
-  "SUM_SUB_NUMSEG" > "HOLLight.hollight.SUM_SUB_NUMSEG"
-  "SUM_SUBSET_SIMPLE" > "HOLLight.hollight.SUM_SUBSET_SIMPLE"
-  "SUM_SUBSET" > "HOLLight.hollight.SUM_SUBSET"
-  "SUM_SUB" > "HOLLight.hollight.SUM_SUB"
-  "SUM_SING_NUMSEG" > "HOLLight.hollight.SUM_SING_NUMSEG"
-  "SUM_SING" > "HOLLight.hollight.SUM_SING"
-  "SUM_RMUL" > "HOLLight.hollight.SUM_RMUL"
-  "SUM_RESTRICT_SET" > "HOLLight.hollight.SUM_RESTRICT_SET"
-  "SUM_RESTRICT" > "HOLLight.hollight.SUM_RESTRICT"
-  "SUM_POS_LE_NUMSEG" > "HOLLight.hollight.SUM_POS_LE_NUMSEG"
-  "SUM_POS_LE" > "HOLLight.hollight.SUM_POS_LE"
-  "SUM_POS_EQ_0_NUMSEG" > "HOLLight.hollight.SUM_POS_EQ_0_NUMSEG"
-  "SUM_POS_EQ_0" > "HOLLight.hollight.SUM_POS_EQ_0"
-  "SUM_POS_BOUND" > "HOLLight.hollight.SUM_POS_BOUND"
-  "SUM_PARTIAL_SUC" > "HOLLight.hollight.SUM_PARTIAL_SUC"
-  "SUM_PARTIAL_PRE" > "HOLLight.hollight.SUM_PARTIAL_PRE"
-  "SUM_PAIR" > "HOLLight.hollight.SUM_PAIR"
-  "SUM_OFFSET_0" > "HOLLight.hollight.SUM_OFFSET_0"
-  "SUM_OFFSET" > "HOLLight.hollight.SUM_OFFSET"
-  "SUM_NEG" > "HOLLight.hollight.SUM_NEG"
-  "SUM_MULTICOUNT_GEN" > "HOLLight.hollight.SUM_MULTICOUNT_GEN"
-  "SUM_MULTICOUNT" > "HOLLight.hollight.SUM_MULTICOUNT"
-  "SUM_LT_ALL" > "HOLLight.hollight.SUM_LT_ALL"
-  "SUM_LT" > "HOLLight.hollight.SUM_LT"
-  "SUM_LMUL" > "HOLLight.hollight.SUM_LMUL"
-  "SUM_LE_NUMSEG" > "HOLLight.hollight.SUM_LE_NUMSEG"
-  "SUM_LE_INCLUDED" > "HOLLight.hollight.SUM_LE_INCLUDED"
-  "SUM_LE" > "HOLLight.hollight.SUM_LE"
-  "SUM_INJECTION" > "HOLLight.hollight.SUM_INJECTION"
-  "SUM_INCL_EXCL" > "HOLLight.hollight.SUM_INCL_EXCL"
-  "SUM_IMAGE_NONZERO" > "HOLLight.hollight.SUM_IMAGE_NONZERO"
-  "SUM_IMAGE_LE" > "HOLLight.hollight.SUM_IMAGE_LE"
-  "SUM_IMAGE_GEN" > "HOLLight.hollight.SUM_IMAGE_GEN"
-  "SUM_IMAGE" > "HOLLight.hollight.SUM_IMAGE"
-  "SUM_GROUP" > "HOLLight.hollight.SUM_GROUP"
-  "SUM_EQ_SUPERSET" > "HOLLight.hollight.SUM_EQ_SUPERSET"
-  "SUM_EQ_NUMSEG" > "HOLLight.hollight.SUM_EQ_NUMSEG"
-  "SUM_EQ_GENERAL_INVERSES" > "HOLLight.hollight.SUM_EQ_GENERAL_INVERSES"
-  "SUM_EQ_GENERAL" > "HOLLight.hollight.SUM_EQ_GENERAL"
-  "SUM_EQ_0_NUMSEG" > "HOLLight.hollight.SUM_EQ_0_NUMSEG"
-  "SUM_EQ_0" > "HOLLight.hollight.SUM_EQ_0"
-  "SUM_EQ" > "HOLLight.hollight.SUM_EQ"
-  "SUM_DIFFS_ALT" > "HOLLight.hollight.SUM_DIFFS_ALT"
-  "SUM_DIFFS" > "HOLLight.hollight.SUM_DIFFS"
-  "SUM_DIFF" > "HOLLight.hollight.SUM_DIFF"
-  "SUM_DELTA" > "HOLLight.hollight.SUM_DELTA"
-  "SUM_DELETE_CASES" > "HOLLight.hollight.SUM_DELETE_CASES"
-  "SUM_DELETE" > "HOLLight.hollight.SUM_DELETE"
-  "SUM_CONST_NUMSEG" > "HOLLight.hollight.SUM_CONST_NUMSEG"
-  "SUM_CONST" > "HOLLight.hollight.SUM_CONST"
-  "SUM_COMBINE_R" > "HOLLight.hollight.SUM_COMBINE_R"
-  "SUM_COMBINE_L" > "HOLLight.hollight.SUM_COMBINE_L"
-  "SUM_CLOSED" > "HOLLight.hollight.SUM_CLOSED"
-  "SUM_CLAUSES_RIGHT" > "HOLLight.hollight.SUM_CLAUSES_RIGHT"
-  "SUM_CLAUSES_NUMSEG" > "HOLLight.hollight.SUM_CLAUSES_NUMSEG"
-  "SUM_CLAUSES_LEFT" > "HOLLight.hollight.SUM_CLAUSES_LEFT"
-  "SUM_CLAUSES" > "HOLLight.hollight.SUM_CLAUSES"
-  "SUM_CASES_1" > "HOLLight.hollight.SUM_CASES_1"
-  "SUM_CASES" > "HOLLight.hollight.SUM_CASES"
-  "SUM_BOUND_LT_GEN" > "HOLLight.hollight.SUM_BOUND_LT_GEN"
-  "SUM_BOUND_LT_ALL" > "HOLLight.hollight.SUM_BOUND_LT_ALL"
-  "SUM_BOUND_LT" > "HOLLight.hollight.SUM_BOUND_LT"
-  "SUM_BOUND_GEN" > "HOLLight.hollight.SUM_BOUND_GEN"
-  "SUM_BOUND" > "HOLLight.hollight.SUM_BOUND"
-  "SUM_BIJECTION" > "HOLLight.hollight.SUM_BIJECTION"
-  "SUM_ADD_SPLIT" > "HOLLight.hollight.SUM_ADD_SPLIT"
-  "SUM_ADD_NUMSEG" > "HOLLight.hollight.SUM_ADD_NUMSEG"
-  "SUM_ADD_GEN" > "HOLLight.hollight.SUM_ADD_GEN"
-  "SUM_ADD" > "HOLLight.hollight.SUM_ADD"
-  "SUM_ABS_NUMSEG" > "HOLLight.hollight.SUM_ABS_NUMSEG"
-  "SUM_ABS_LE" > "HOLLight.hollight.SUM_ABS_LE"
-  "SUM_ABS_BOUND" > "HOLLight.hollight.SUM_ABS_BOUND"
-  "SUM_ABS" > "HOLLight.hollight.SUM_ABS"
-  "SUM_0" > "HOLLight.hollight.SUM_0"
-  "SUC_SUB1" > "Nat.diff_Suc_1"
-  "SUC_INJ" > "HOLLightCompat.SUC_INJ"
-  "SUB_SUC" > "Nat.diff_Suc_Suc"
-  "SUB_REFL" > "Nat.diff_self_eq_0"
-  "SUB_PRESUC" > "HOLLight.hollight.SUB_PRESUC"
-  "SUB_EQ_0" > "Nat.diff_is_0_eq"
-  "SUB_ELIM_THM" > "HOLLight.hollight.SUB_ELIM_THM"
-  "SUB_ADD_RCANCEL" > "Nat.diff_cancel2"
-  "SUB_ADD_LCANCEL" > "Nat.diff_cancel"
-  "SUB_ADD" > "Nat.le_add_diff_inverse2"
-  "SUB_0" > "HOLLight.hollight.SUB_0"
-  "SUBSET_UNIV" > "Orderings.top_class.top_greatest"
-  "SUBSET_UNION_ABSORPTION" > "Lattices.semilattice_sup_class.le_iff_sup"
-  "SUBSET_UNIONS" > "Complete_Lattices.Union_mono"
-  "SUBSET_UNION" > "HOLLight.hollight.SUBSET_UNION"
-  "SUBSET_TRANS" > "Orderings.order_trans_rules_23"
-  "SUBSET_RESTRICT" > "HOLLight.hollight.SUBSET_RESTRICT"
-  "SUBSET_REFL" > "Inductive.basic_monos_1"
-  "SUBSET_PSUBSET_TRANS" > "Orderings.order_le_less_trans"
-  "SUBSET_NUMSEG" > "HOLLight.hollight.SUBSET_NUMSEG"
-  "SUBSET_INTER_ABSORPTION" > "Lattices.semilattice_inf_class.le_iff_inf"
-  "SUBSET_INTER" > "Lattices.semilattice_inf_class.le_inf_iff"
-  "SUBSET_INSERT_DELETE" > "HOLLight.hollight.SUBSET_INSERT_DELETE"
-  "SUBSET_INSERT" > "Set.subset_insert"
-  "SUBSET_IMAGE" > "Set.subset_image_iff"
-  "SUBSET_EMPTY" > "Orderings.bot_class.bot_unique"
-  "SUBSET_DIFF" > "Set.Diff_subset"
-  "SUBSET_DELETE" > "HOLLight.hollight.SUBSET_DELETE"
-  "SUBSET_CARD_EQ" > "HOLLight.hollight.SUBSET_CARD_EQ"
-  "SUBSET_ANTISYM_EQ" > "Orderings.order_class.eq_iff"
-  "SUBSET_ANTISYM" > "Orderings.order_antisym"
-  "SND" > "Product_Type.snd_conv"
-  "SKOLEM_THM" > "HOLLight.hollight.SKOLEM_THM"
-  "SING_def" > "HOLLight.hollight.SING_def"
-  "SING_SUBSET" > "HOLLight.hollight.SING_SUBSET"
-  "SING_GSPEC" > "HOLLight.hollight.SING_GSPEC"
-  "SIMPLE_IMAGE_GEN" > "HOLLight.hollight.SIMPLE_IMAGE_GEN"
-  "SIMPLE_IMAGE" > "HOLLight.hollight.SIMPLE_IMAGE"
-  "SET_RECURSION_LEMMA" > "HOLLight.hollight.SET_RECURSION_LEMMA"
-  "SET_PROVE_CASES" > "HOLLight.hollight.SET_PROVE_CASES"
-  "SET_PAIR_THM" > "HOLLight.hollight.SET_PAIR_THM"
-  "SET_OF_LIST_MAP" > "List.set_map"
-  "SET_OF_LIST_EQ_EMPTY" > "List.set_empty"
-  "SET_OF_LIST_APPEND" > "List.set_append"
-  "SET_CASES" > "HOLLight.hollight.SET_CASES"
-  "SEMIRING_PTHS" > "HOLLight.hollight.SEMIRING_PTHS"
-  "SELECT_UNIQUE" > "HOLLight.hollight.SELECT_UNIQUE"
-  "SELECT_REFL" > "Hilbert_Choice.some_eq_trivial"
-  "SELECT_LEMMA" > "Hilbert_Choice.some_sym_eq_trivial"
-  "RIGHT_SUB_DISTRIB" > "Nat.diff_mult_distrib"
-  "RIGHT_OR_FORALL_THM" > "HOL.all_simps_4"
-  "RIGHT_OR_EXISTS_THM" > "HOL.ex_simps_4"
-  "RIGHT_OR_DISTRIB" > "Groebner_Basis.dnf_2"
-  "RIGHT_IMP_FORALL_THM" > "HOL.all_simps_6"
-  "RIGHT_IMP_EXISTS_THM" > "HOL.ex_simps_6"
-  "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_FORALL_THM" > "HOL.all_simps_2"
-  "RIGHT_AND_EXISTS_THM" > "HOL.ex_simps_2"
-  "RIGHT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_26"
-  "REVERSE_REVERSE" > "List.rev_rev_ident"
-  "REVERSE_APPEND" > "List.rev_append"
-  "REST_def" > "HOLLight.hollight.REST_def"
-  "REFL_CLAUSE" > "Groebner_Basis.bool_simps_6"
-  "REAL_WLOG_LT" > "HOLLight.hollight.REAL_WLOG_LT"
-  "REAL_WLOG_LE" > "HOLLight.hollight.REAL_WLOG_LE"
-  "REAL_SUB_TRIANGLE" > "HOLLight.hollight.REAL_SUB_TRIANGLE"
-  "REAL_SUB_SUB2" > "HOLLight.hollight.REAL_SUB_SUB2"
-  "REAL_SUB_SUB" > "HOLLight.hollight.REAL_SUB_SUB"
-  "REAL_SUB_RZERO" > "HOLLight.hollight.REAL_SUB_RZERO"
-  "REAL_SUB_RNEG" > "HOLLight.hollight.REAL_SUB_RNEG"
-  "REAL_SUB_REFL" > "HOLLight.hollight.REAL_SUB_REFL"
-  "REAL_SUB_RDISTRIB" > "HOLLight.hollight.REAL_SUB_RDISTRIB"
-  "REAL_SUB_POW_R1" > "HOLLight.hollight.REAL_SUB_POW_R1"
-  "REAL_SUB_POW_L1" > "HOLLight.hollight.REAL_SUB_POW_L1"
-  "REAL_SUB_POW" > "HOLLight.hollight.REAL_SUB_POW"
-  "REAL_SUB_NEG2" > "HOLLight.hollight.REAL_SUB_NEG2"
-  "REAL_SUB_LZERO" > "HOLLight.hollight.REAL_SUB_LZERO"
-  "REAL_SUB_LT" > "HOLLight.hollight.REAL_SUB_LT"
-  "REAL_SUB_LNEG" > "HOLLight.hollight.REAL_SUB_LNEG"
-  "REAL_SUB_LE" > "HOLLight.hollight.REAL_SUB_LE"
-  "REAL_SUB_LDISTRIB" > "HOLLight.hollight.REAL_SUB_LDISTRIB"
-  "REAL_SUB_INV" > "HOLLight.hollight.REAL_SUB_INV"
-  "REAL_SUB_ADD2" > "HOLLight.hollight.REAL_SUB_ADD2"
-  "REAL_SUB_ADD" > "HOLLight.hollight.REAL_SUB_ADD"
-  "REAL_SUB_ABS" > "HOLLight.hollight.REAL_SUB_ABS"
-  "REAL_SUB_0" > "HOLLight.hollight.REAL_SUB_0"
-  "REAL_SOS_EQ_0" > "HOLLight.hollight.REAL_SOS_EQ_0"
-  "REAL_SGN_NEG" > "HOLLight.hollight.REAL_SGN_NEG"
-  "REAL_SGN_MUL" > "HOLLight.hollight.REAL_SGN_MUL"
-  "REAL_SGN_INV" > "HOLLight.hollight.REAL_SGN_INV"
-  "REAL_SGN_DIV" > "HOLLight.hollight.REAL_SGN_DIV"
-  "REAL_SGN_ABS" > "HOLLight.hollight.REAL_SGN_ABS"
-  "REAL_SGN_0" > "HOLLight.hollight.REAL_SGN_0"
-  "REAL_SGN" > "HOLLight.hollight.REAL_SGN"
-  "REAL_RNEG_UNIQ" > "HOLLight.hollight.REAL_RNEG_UNIQ"
-  "REAL_POW_ZERO" > "HOLLight.hollight.REAL_POW_ZERO"
-  "REAL_POW_SUB" > "HOLLight.hollight.REAL_POW_SUB"
-  "REAL_POW_POW" > "HOLLight.hollight.REAL_POW_POW"
-  "REAL_POW_ONE" > "HOLLight.hollight.REAL_POW_ONE"
-  "REAL_POW_NZ" > "HOLLight.hollight.REAL_POW_NZ"
-  "REAL_POW_NEG" > "HOLLight.hollight.REAL_POW_NEG"
-  "REAL_POW_MUL" > "HOLLight.hollight.REAL_POW_MUL"
-  "REAL_POW_MONO_LT" > "HOLLight.hollight.REAL_POW_MONO_LT"
-  "REAL_POW_MONO_INV" > "HOLLight.hollight.REAL_POW_MONO_INV"
-  "REAL_POW_MONO" > "HOLLight.hollight.REAL_POW_MONO"
-  "REAL_POW_LT_1" > "HOLLight.hollight.REAL_POW_LT_1"
-  "REAL_POW_LT2_REV" > "HOLLight.hollight.REAL_POW_LT2_REV"
-  "REAL_POW_LT2_ODD_EQ" > "HOLLight.hollight.REAL_POW_LT2_ODD_EQ"
-  "REAL_POW_LT2_ODD" > "HOLLight.hollight.REAL_POW_LT2_ODD"
-  "REAL_POW_LT2" > "HOLLight.hollight.REAL_POW_LT2"
-  "REAL_POW_LT" > "HOLLight.hollight.REAL_POW_LT"
-  "REAL_POW_LE_1" > "HOLLight.hollight.REAL_POW_LE_1"
-  "REAL_POW_LE2_REV" > "HOLLight.hollight.REAL_POW_LE2_REV"
-  "REAL_POW_LE2_ODD_EQ" > "HOLLight.hollight.REAL_POW_LE2_ODD_EQ"
-  "REAL_POW_LE2_ODD" > "HOLLight.hollight.REAL_POW_LE2_ODD"
-  "REAL_POW_LE2" > "HOLLight.hollight.REAL_POW_LE2"
-  "REAL_POW_LE" > "HOLLight.hollight.REAL_POW_LE"
-  "REAL_POW_INV" > "HOLLight.hollight.REAL_POW_INV"
-  "REAL_POW_EQ_ODD_EQ" > "HOLLight.hollight.REAL_POW_EQ_ODD_EQ"
-  "REAL_POW_EQ_ODD" > "HOLLight.hollight.REAL_POW_EQ_ODD"
-  "REAL_POW_EQ_EQ" > "HOLLight.hollight.REAL_POW_EQ_EQ"
-  "REAL_POW_EQ_ABS" > "HOLLight.hollight.REAL_POW_EQ_ABS"
-  "REAL_POW_EQ_1_IMP" > "HOLLight.hollight.REAL_POW_EQ_1_IMP"
-  "REAL_POW_EQ_1" > "HOLLight.hollight.REAL_POW_EQ_1"
-  "REAL_POW_EQ_0" > "HOLLight.hollight.REAL_POW_EQ_0"
-  "REAL_POW_EQ" > "HOLLight.hollight.REAL_POW_EQ"
-  "REAL_POW_DIV" > "HOLLight.hollight.REAL_POW_DIV"
-  "REAL_POW_ADD" > "HOLLight.hollight.REAL_POW_ADD"
-  "REAL_POW_2" > "HOLLight.hollight.REAL_POW_2"
-  "REAL_POW_1_LT" > "HOLLight.hollight.REAL_POW_1_LT"
-  "REAL_POW_1_LE" > "HOLLight.hollight.REAL_POW_1_LE"
-  "REAL_POW_1" > "HOLLight.hollight.REAL_POW_1"
-  "REAL_POW2_ABS" > "HOLLight.hollight.REAL_POW2_ABS"
-  "REAL_POS_NZ" > "HOLLight.hollight.REAL_POS_NZ"
-  "REAL_POS" > "HOLLight.hollight.REAL_POS"
-  "REAL_POLY_NEG_CLAUSES" > "HOLLight.hollight.REAL_POLY_NEG_CLAUSES"
-  "REAL_POLY_CLAUSES" > "HOLLight.hollight.REAL_POLY_CLAUSES"
-  "REAL_OF_NUM_SUM_NUMSEG" > "HOLLight.hollight.REAL_OF_NUM_SUM_NUMSEG"
-  "REAL_OF_NUM_SUM" > "HOLLight.hollight.REAL_OF_NUM_SUM"
-  "REAL_OF_NUM_SUC" > "HOLLight.hollight.REAL_OF_NUM_SUC"
-  "REAL_OF_NUM_SUB" > "HOLLight.hollight.REAL_OF_NUM_SUB"
-  "REAL_OF_NUM_POW" > "HOLLight.hollight.REAL_OF_NUM_POW"
-  "REAL_OF_NUM_MIN" > "HOLLight.hollight.REAL_OF_NUM_MIN"
-  "REAL_OF_NUM_MAX" > "HOLLight.hollight.REAL_OF_NUM_MAX"
-  "REAL_OF_NUM_LT" > "HOLLight.hollight.REAL_OF_NUM_LT"
-  "REAL_OF_NUM_GT" > "HOLLight.hollight.REAL_OF_NUM_GT"
-  "REAL_OF_NUM_GE" > "HOLLight.hollight.REAL_OF_NUM_GE"
-  "REAL_NOT_LT" > "HOLLight.hollight.REAL_NOT_LT"
-  "REAL_NOT_LE" > "HOLLight.hollight.real_lt_def"
-  "REAL_NOT_EQ" > "HOLLight.hollight.REAL_NOT_EQ"
-  "REAL_NEG_SUB" > "HOLLight.hollight.REAL_NEG_SUB"
-  "REAL_NEG_RMUL" > "HOLLight.hollight.REAL_NEG_RMUL"
-  "REAL_NEG_NEG" > "HOLLight.hollight.REAL_NEG_NEG"
-  "REAL_NEG_MUL2" > "HOLLight.hollight.REAL_NEG_MUL2"
-  "REAL_NEG_MINUS1" > "HOLLight.hollight.REAL_NEG_MINUS1"
-  "REAL_NEG_LT0" > "HOLLight.hollight.REAL_NEG_LT0"
-  "REAL_NEG_LMUL" > "HOLLight.hollight.REAL_NEG_LMUL"
-  "REAL_NEG_LE0" > "HOLLight.hollight.REAL_NEG_LE0"
-  "REAL_NEG_GT0" > "HOLLight.hollight.REAL_NEG_GT0"
-  "REAL_NEG_GE0" > "HOLLight.hollight.REAL_NEG_GE0"
-  "REAL_NEG_EQ_0" > "HOLLight.hollight.REAL_NEG_EQ_0"
-  "REAL_NEG_EQ" > "HOLLight.hollight.REAL_NEG_EQ"
-  "REAL_NEG_ADD" > "HOLLight.hollight.REAL_NEG_ADD"
-  "REAL_NEG_0" > "HOLLight.hollight.REAL_NEG_0"
-  "REAL_NEGNEG" > "HOLLight.hollight.REAL_NEGNEG"
-  "REAL_MUL_RZERO" > "HOLLight.hollight.REAL_MUL_RZERO"
-  "REAL_MUL_RNEG" > "HOLLight.hollight.REAL_MUL_RNEG"
-  "REAL_MUL_RINV_UNIQ" > "HOLLight.hollight.REAL_MUL_RINV_UNIQ"
-  "REAL_MUL_RINV" > "HOLLight.hollight.REAL_MUL_RINV"
-  "REAL_MUL_RID" > "HOLLight.hollight.REAL_MUL_RID"
-  "REAL_MUL_POS_LT" > "HOLLight.hollight.REAL_MUL_POS_LT"
-  "REAL_MUL_POS_LE" > "HOLLight.hollight.REAL_MUL_POS_LE"
-  "REAL_MUL_LZERO" > "HOLLight.hollight.REAL_MUL_LZERO"
-  "REAL_MUL_LNEG" > "HOLLight.hollight.REAL_MUL_LNEG"
-  "REAL_MUL_LINV_UNIQ" > "HOLLight.hollight.REAL_MUL_LINV_UNIQ"
-  "REAL_MUL_AC" > "HOLLight.hollight.REAL_MUL_AC"
-  "REAL_MUL_2" > "HOLLight.hollight.REAL_MUL_2"
-  "REAL_MIN_SYM" > "HOLLight.hollight.REAL_MIN_SYM"
-  "REAL_MIN_MIN" > "HOLLight.hollight.REAL_MIN_MIN"
-  "REAL_MIN_MAX" > "HOLLight.hollight.REAL_MIN_MAX"
-  "REAL_MIN_LT" > "HOLLight.hollight.REAL_MIN_LT"
-  "REAL_MIN_LE" > "HOLLight.hollight.REAL_MIN_LE"
-  "REAL_MIN_ASSOC" > "HOLLight.hollight.REAL_MIN_ASSOC"
-  "REAL_MIN_ACI" > "HOLLight.hollight.REAL_MIN_ACI"
-  "REAL_MAX_SYM" > "HOLLight.hollight.REAL_MAX_SYM"
-  "REAL_MAX_MIN" > "HOLLight.hollight.REAL_MAX_MIN"
-  "REAL_MAX_MAX" > "HOLLight.hollight.REAL_MAX_MAX"
-  "REAL_MAX_LT" > "HOLLight.hollight.REAL_MAX_LT"
-  "REAL_MAX_LE" > "HOLLight.hollight.REAL_MAX_LE"
-  "REAL_MAX_ASSOC" > "HOLLight.hollight.REAL_MAX_ASSOC"
-  "REAL_MAX_ACI" > "HOLLight.hollight.REAL_MAX_ACI"
-  "REAL_LT_TRANS" > "HOLLight.hollight.REAL_LT_TRANS"
-  "REAL_LT_TOTAL" > "HOLLight.hollight.REAL_LT_TOTAL"
-  "REAL_LT_SUB_RADD" > "HOLLight.hollight.REAL_LT_SUB_RADD"
-  "REAL_LT_SUB_LADD" > "HOLLight.hollight.REAL_LT_SUB_LADD"
-  "REAL_LT_SQUARE_ABS" > "HOLLight.hollight.REAL_LT_SQUARE_ABS"
-  "REAL_LT_SQUARE" > "HOLLight.hollight.REAL_LT_SQUARE"
-  "REAL_LT_RNEG" > "HOLLight.hollight.REAL_LT_RNEG"
-  "REAL_LT_RMUL_EQ" > "HOLLight.hollight.REAL_LT_RMUL_EQ"
-  "REAL_LT_RMUL" > "HOLLight.hollight.REAL_LT_RMUL"
-  "REAL_LT_RINV" > "HOLLight.hollight.REAL_LT_RINV"
-  "REAL_LT_REFL" > "HOLLight.hollight.REAL_LT_REFL"
-  "REAL_LT_RDIV_EQ" > "HOLLight.hollight.REAL_LT_RDIV_EQ"
-  "REAL_LT_RCANCEL_IMP" > "HOLLight.hollight.REAL_LT_RCANCEL_IMP"
-  "REAL_LT_RADD" > "HOLLight.hollight.REAL_LT_RADD"
-  "REAL_LT_POW2" > "HOLLight.hollight.REAL_LT_POW2"
-  "REAL_LT_NEGTOTAL" > "HOLLight.hollight.REAL_LT_NEGTOTAL"
-  "REAL_LT_NEG2" > "HOLLight.hollight.REAL_LT_NEG2"
-  "REAL_LT_NEG" > "HOLLight.hollight.REAL_LT_NEG"
-  "REAL_LT_MUL_EQ" > "HOLLight.hollight.REAL_LT_MUL_EQ"
-  "REAL_LT_MUL2" > "HOLLight.hollight.REAL_LT_MUL2"
-  "REAL_LT_MUL" > "HOLLight.hollight.REAL_LT_MUL"
-  "REAL_LT_MIN" > "HOLLight.hollight.REAL_LT_MIN"
-  "REAL_LT_MAX" > "HOLLight.hollight.REAL_LT_MAX"
-  "REAL_LT_LNEG" > "HOLLight.hollight.REAL_LT_LNEG"
-  "REAL_LT_LMUL_EQ" > "HOLLight.hollight.REAL_LT_LMUL_EQ"
-  "REAL_LT_LMUL" > "HOLLight.hollight.REAL_LT_LMUL"
-  "REAL_LT_LINV" > "HOLLight.hollight.REAL_LT_LINV"
-  "REAL_LT_LE" > "HOLLight.hollight.REAL_LT_LE"
-  "REAL_LT_LDIV_EQ" > "HOLLight.hollight.REAL_LT_LDIV_EQ"
-  "REAL_LT_LCANCEL_IMP" > "HOLLight.hollight.REAL_LT_LCANCEL_IMP"
-  "REAL_LT_LADD_IMP" > "HOLLight.hollight.REAL_LT_LADD_IMP"
-  "REAL_LT_LADD" > "HOLLight.hollight.REAL_LT_LADD"
-  "REAL_LT_INV_EQ" > "HOLLight.hollight.REAL_LT_INV_EQ"
-  "REAL_LT_INV2" > "HOLLight.hollight.REAL_LT_INV2"
-  "REAL_LT_INV" > "HOLLight.hollight.REAL_LT_INV"
-  "REAL_LT_IMP_NZ" > "HOLLight.hollight.REAL_LT_IMP_NZ"
-  "REAL_LT_IMP_NE" > "HOLLight.hollight.REAL_LT_IMP_NE"
-  "REAL_LT_IMP_LE" > "HOLLight.hollight.REAL_LT_IMP_LE"
-  "REAL_LT_GT" > "HOLLight.hollight.REAL_LT_GT"
-  "REAL_LT_DIV2_EQ" > "HOLLight.hollight.REAL_LT_DIV2_EQ"
-  "REAL_LT_DIV" > "HOLLight.hollight.REAL_LT_DIV"
-  "REAL_LT_ANTISYM" > "HOLLight.hollight.REAL_LT_ANTISYM"
-  "REAL_LT_ADD_SUB" > "HOLLight.hollight.REAL_LT_ADD_SUB"
-  "REAL_LT_ADDR" > "HOLLight.hollight.REAL_LT_ADDR"
-  "REAL_LT_ADDNEG2" > "HOLLight.hollight.REAL_LT_ADDNEG2"
-  "REAL_LT_ADDNEG" > "HOLLight.hollight.REAL_LT_ADDNEG"
-  "REAL_LT_ADDL" > "HOLLight.hollight.REAL_LT_ADDL"
-  "REAL_LT_ADD2" > "HOLLight.hollight.REAL_LT_ADD2"
-  "REAL_LT_ADD1" > "HOLLight.hollight.REAL_LT_ADD1"
-  "REAL_LT_ADD" > "HOLLight.hollight.REAL_LT_ADD"
-  "REAL_LT_01" > "HOLLight.hollight.REAL_LT_01"
-  "REAL_LTE_TRANS" > "HOLLight.hollight.REAL_LTE_TRANS"
-  "REAL_LTE_TOTAL" > "HOLLight.hollight.REAL_LTE_TOTAL"
-  "REAL_LTE_ANTISYM" > "HOLLight.hollight.REAL_LTE_ANTISYM"
-  "REAL_LTE_ADD2" > "HOLLight.hollight.REAL_LTE_ADD2"
-  "REAL_LTE_ADD" > "HOLLight.hollight.REAL_LTE_ADD"
-  "REAL_LNEG_UNIQ" > "HOLLight.hollight.REAL_LNEG_UNIQ"
-  "REAL_LE_SUB_RADD" > "HOLLight.hollight.REAL_LE_SUB_RADD"
-  "REAL_LE_SUB_LADD" > "HOLLight.hollight.REAL_LE_SUB_LADD"
-  "REAL_LE_SQUARE_ABS" > "HOLLight.hollight.REAL_LE_SQUARE_ABS"
-  "REAL_LE_SQUARE" > "HOLLight.hollight.REAL_LE_SQUARE"
-  "REAL_LE_RNEG" > "HOLLight.hollight.REAL_LE_RNEG"
-  "REAL_LE_RMUL_EQ" > "HOLLight.hollight.REAL_LE_RMUL_EQ"
-  "REAL_LE_RMUL" > "HOLLight.hollight.REAL_LE_RMUL"
-  "REAL_LE_RINV" > "HOLLight.hollight.REAL_LE_RINV"
-  "REAL_LE_RDIV_EQ" > "HOLLight.hollight.REAL_LE_RDIV_EQ"
-  "REAL_LE_RCANCEL_IMP" > "HOLLight.hollight.REAL_LE_RCANCEL_IMP"
-  "REAL_LE_RADD" > "HOLLight.hollight.REAL_LE_RADD"
-  "REAL_LE_POW_2" > "HOLLight.hollight.REAL_LE_POW_2"
-  "REAL_LE_POW2" > "HOLLight.hollight.REAL_LE_POW2"
-  "REAL_LE_NEGTOTAL" > "HOLLight.hollight.REAL_LE_NEGTOTAL"
-  "REAL_LE_NEGR" > "HOLLight.hollight.REAL_LE_NEGR"
-  "REAL_LE_NEGL" > "HOLLight.hollight.REAL_LE_NEGL"
-  "REAL_LE_NEG2" > "HOLLight.hollight.REAL_LE_NEG2"
-  "REAL_LE_NEG" > "HOLLight.hollight.REAL_LE_NEG"
-  "REAL_LE_MUL_EQ" > "HOLLight.hollight.REAL_LE_MUL_EQ"
-  "REAL_LE_MUL2" > "HOLLight.hollight.REAL_LE_MUL2"
-  "REAL_LE_MIN" > "HOLLight.hollight.REAL_LE_MIN"
-  "REAL_LE_MAX" > "HOLLight.hollight.REAL_LE_MAX"
-  "REAL_LE_LT" > "HOLLight.hollight.REAL_LE_LT"
-  "REAL_LE_LNEG" > "HOLLight.hollight.REAL_LE_LNEG"
-  "REAL_LE_LMUL_EQ" > "HOLLight.hollight.REAL_LE_LMUL_EQ"
-  "REAL_LE_LMUL" > "HOLLight.hollight.REAL_LE_LMUL"
-  "REAL_LE_LINV" > "HOLLight.hollight.REAL_LE_LINV"
-  "REAL_LE_LDIV_EQ" > "HOLLight.hollight.REAL_LE_LDIV_EQ"
-  "REAL_LE_LCANCEL_IMP" > "HOLLight.hollight.REAL_LE_LCANCEL_IMP"
-  "REAL_LE_LADD" > "HOLLight.hollight.REAL_LE_LADD"
-  "REAL_LE_INV_EQ" > "HOLLight.hollight.REAL_LE_INV_EQ"
-  "REAL_LE_INV2" > "HOLLight.hollight.REAL_LE_INV2"
-  "REAL_LE_INV" > "HOLLight.hollight.REAL_LE_INV"
-  "REAL_LE_DOUBLE" > "HOLLight.hollight.REAL_LE_DOUBLE"
-  "REAL_LE_DIV2_EQ" > "HOLLight.hollight.REAL_LE_DIV2_EQ"
-  "REAL_LE_DIV" > "HOLLight.hollight.REAL_LE_DIV"
-  "REAL_LE_ADDR" > "HOLLight.hollight.REAL_LE_ADDR"
-  "REAL_LE_ADDL" > "HOLLight.hollight.REAL_LE_ADDL"
-  "REAL_LE_ADD2" > "HOLLight.hollight.REAL_LE_ADD2"
-  "REAL_LE_ADD" > "HOLLight.hollight.REAL_LE_ADD"
-  "REAL_LE_01" > "HOLLight.hollight.REAL_LE_01"
-  "REAL_LET_TRANS" > "HOLLight.hollight.REAL_LET_TRANS"
-  "REAL_LET_TOTAL" > "HOLLight.hollight.REAL_LET_TOTAL"
-  "REAL_LET_ANTISYM" > "HOLLight.hollight.REAL_LET_ANTISYM"
-  "REAL_LET_ADD2" > "HOLLight.hollight.REAL_LET_ADD2"
-  "REAL_LET_ADD" > "HOLLight.hollight.REAL_LET_ADD"
-  "REAL_INV_POW" > "HOLLight.hollight.REAL_INV_POW"
-  "REAL_INV_NEG" > "HOLLight.hollight.REAL_INV_NEG"
-  "REAL_INV_MUL" > "HOLLight.hollight.REAL_INV_MUL"
-  "REAL_INV_LT_1" > "HOLLight.hollight.REAL_INV_LT_1"
-  "REAL_INV_LE_1" > "HOLLight.hollight.REAL_INV_LE_1"
-  "REAL_INV_INV" > "HOLLight.hollight.REAL_INV_INV"
-  "REAL_INV_EQ_1" > "HOLLight.hollight.REAL_INV_EQ_1"
-  "REAL_INV_EQ_0" > "HOLLight.hollight.REAL_INV_EQ_0"
-  "REAL_INV_DIV" > "HOLLight.hollight.REAL_INV_DIV"
-  "REAL_INV_1_LT" > "HOLLight.hollight.REAL_INV_1_LT"
-  "REAL_INV_1_LE" > "HOLLight.hollight.REAL_INV_1_LE"
-  "REAL_INV_1" > "HOLLight.hollight.REAL_INV_1"
-  "REAL_INTEGRAL" > "HOLLight.hollight.REAL_INTEGRAL"
-  "REAL_HREAL_LEMMA2" > "HOLLight.hollight.REAL_HREAL_LEMMA2"
-  "REAL_HREAL_LEMMA1" > "HOLLight.hollight.REAL_HREAL_LEMMA1"
-  "REAL_EQ_SUB_RADD" > "HOLLight.hollight.REAL_EQ_SUB_RADD"
-  "REAL_EQ_SUB_LADD" > "HOLLight.hollight.REAL_EQ_SUB_LADD"
-  "REAL_EQ_SQUARE_ABS" > "HOLLight.hollight.REAL_EQ_SQUARE_ABS"
-  "REAL_EQ_RDIV_EQ" > "HOLLight.hollight.REAL_EQ_RDIV_EQ"
-  "REAL_EQ_RCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_RCANCEL_IMP"
-  "REAL_EQ_NEG2" > "HOLLight.hollight.REAL_EQ_NEG2"
-  "REAL_EQ_MUL_RCANCEL" > "HOLLight.hollight.REAL_EQ_MUL_RCANCEL"
-  "REAL_EQ_MUL_LCANCEL" > "HOLLight.hollight.REAL_EQ_MUL_LCANCEL"
-  "REAL_EQ_LDIV_EQ" > "HOLLight.hollight.REAL_EQ_LDIV_EQ"
-  "REAL_EQ_LCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_LCANCEL_IMP"
-  "REAL_EQ_INV2" > "HOLLight.hollight.REAL_EQ_INV2"
-  "REAL_EQ_IMP_LE" > "HOLLight.hollight.REAL_EQ_IMP_LE"
-  "REAL_EQ_ADD_RCANCEL_0" > "HOLLight.hollight.REAL_EQ_ADD_RCANCEL_0"
-  "REAL_EQ_ADD_RCANCEL" > "HOLLight.hollight.REAL_EQ_ADD_RCANCEL"
-  "REAL_EQ_ADD_LCANCEL_0" > "HOLLight.hollight.REAL_EQ_ADD_LCANCEL_0"
-  "REAL_EQ_ADD_LCANCEL" > "HOLLight.hollight.REAL_EQ_ADD_LCANCEL"
-  "REAL_ENTIRE" > "HOLLight.hollight.REAL_ENTIRE"
-  "REAL_DOWN2" > "HOLLight.hollight.REAL_DOWN2"
-  "REAL_DOWN" > "HOLLight.hollight.REAL_DOWN"
-  "REAL_DIV_RMUL" > "HOLLight.hollight.REAL_DIV_RMUL"
-  "REAL_DIV_REFL" > "HOLLight.hollight.REAL_DIV_REFL"
-  "REAL_DIV_POW2_ALT" > "HOLLight.hollight.REAL_DIV_POW2_ALT"
-  "REAL_DIV_POW2" > "HOLLight.hollight.REAL_DIV_POW2"
-  "REAL_DIV_LMUL" > "HOLLight.hollight.REAL_DIV_LMUL"
-  "REAL_DIV_1" > "HOLLight.hollight.REAL_DIV_1"
-  "REAL_DIFFSQ" > "HOLLight.hollight.REAL_DIFFSQ"
-  "REAL_COMPLETE_SOMEPOS" > "HOLLight.hollight.REAL_COMPLETE_SOMEPOS"
-  "REAL_COMPLETE" > "HOLLight.hollight.REAL_COMPLETE"
-  "REAL_BOUNDS_LT" > "HOLLight.hollight.REAL_BOUNDS_LT"
-  "REAL_BOUNDS_LE" > "HOLLight.hollight.REAL_BOUNDS_LE"
-  "REAL_ADD_SUB2" > "HOLLight.hollight.REAL_ADD_SUB2"
-  "REAL_ADD_SUB" > "HOLLight.hollight.REAL_ADD_SUB"
-  "REAL_ADD_RINV" > "HOLLight.hollight.REAL_ADD_RINV"
-  "REAL_ADD_RID" > "HOLLight.hollight.REAL_ADD_RID"
-  "REAL_ADD_RDISTRIB" > "HOLLight.hollight.REAL_ADD_RDISTRIB"
-  "REAL_ADD_AC" > "HOLLight.hollight.REAL_ADD_AC"
-  "REAL_ADD2_SUB2" > "HOLLight.hollight.REAL_ADD2_SUB2"
-  "REAL_ABS_ZERO" > "HOLLight.hollight.REAL_ABS_ZERO"
-  "REAL_ABS_TRIANGLE_LT" > "HOLLight.hollight.REAL_ABS_TRIANGLE_LT"
-  "REAL_ABS_TRIANGLE_LE" > "HOLLight.hollight.REAL_ABS_TRIANGLE_LE"
-  "REAL_ABS_TRIANGLE" > "HOLLight.hollight.REAL_ABS_TRIANGLE"
-  "REAL_ABS_SUB_ABS" > "HOLLight.hollight.REAL_ABS_SUB_ABS"
-  "REAL_ABS_SUB" > "HOLLight.hollight.REAL_ABS_SUB"
-  "REAL_ABS_STILLNZ" > "HOLLight.hollight.REAL_ABS_STILLNZ"
-  "REAL_ABS_SIGN2" > "HOLLight.hollight.REAL_ABS_SIGN2"
-  "REAL_ABS_SIGN" > "HOLLight.hollight.REAL_ABS_SIGN"
-  "REAL_ABS_SGN" > "HOLLight.hollight.REAL_ABS_SGN"
-  "REAL_ABS_REFL" > "HOLLight.hollight.REAL_ABS_REFL"
-  "REAL_ABS_POW" > "HOLLight.hollight.REAL_ABS_POW"
-  "REAL_ABS_POS" > "HOLLight.hollight.REAL_ABS_POS"
-  "REAL_ABS_NZ" > "HOLLight.hollight.REAL_ABS_NZ"
-  "REAL_ABS_NUM" > "HOLLight.hollight.REAL_ABS_NUM"
-  "REAL_ABS_NEG" > "HOLLight.hollight.REAL_ABS_NEG"
-  "REAL_ABS_MUL" > "HOLLight.hollight.REAL_ABS_MUL"
-  "REAL_ABS_LE" > "HOLLight.hollight.REAL_ABS_LE"
-  "REAL_ABS_INV" > "HOLLight.hollight.REAL_ABS_INV"
-  "REAL_ABS_DIV" > "HOLLight.hollight.REAL_ABS_DIV"
-  "REAL_ABS_CIRCLE" > "HOLLight.hollight.REAL_ABS_CIRCLE"
-  "REAL_ABS_CASES" > "HOLLight.hollight.REAL_ABS_CASES"
-  "REAL_ABS_BOUNDS" > "HOLLight.hollight.REAL_ABS_BOUNDS"
-  "REAL_ABS_BOUND" > "HOLLight.hollight.REAL_ABS_BOUND"
-  "REAL_ABS_BETWEEN2" > "HOLLight.hollight.REAL_ABS_BETWEEN2"
-  "REAL_ABS_BETWEEN1" > "HOLLight.hollight.REAL_ABS_BETWEEN1"
-  "REAL_ABS_BETWEEN" > "HOLLight.hollight.REAL_ABS_BETWEEN"
-  "REAL_ABS_ABS" > "HOLLight.hollight.REAL_ABS_ABS"
-  "REAL_ABS_1" > "HOLLight.hollight.REAL_ABS_1"
-  "REAL_ABS_0" > "HOLLight.hollight.REAL_ABS_0"
-  "RAT_LEMMA5" > "HOLLight.hollight.RAT_LEMMA5"
-  "RAT_LEMMA4" > "HOLLight.hollight.RAT_LEMMA4"
-  "RAT_LEMMA3" > "HOLLight.hollight.RAT_LEMMA3"
-  "RAT_LEMMA2" > "HOLLight.hollight.RAT_LEMMA2"
-  "RAT_LEMMA1" > "HOLLight.hollight.RAT_LEMMA1"
-  "PSUBSET_UNIV" > "HOLLight.hollight.PSUBSET_UNIV"
-  "PSUBSET_TRANS" > "Orderings.order_less_trans"
-  "PSUBSET_SUBSET_TRANS" > "Orderings.order_less_le_trans"
-  "PSUBSET_MEMBER" > "HOLLight.hollight.PSUBSET_MEMBER"
-  "PSUBSET_IRREFL" > "Orderings.order_less_irrefl"
-  "PSUBSET_INSERT_SUBSET" > "HOLLight.hollight.PSUBSET_INSERT_SUBSET"
-  "PSUBSET_ALT" > "HOLLight.hollight.PSUBSET_ALT"
-  "PRE_ELIM_THM" > "HOLLight.hollight.PRE_ELIM_THM"
-  "POWERSET_CLAUSES" > "HOLLight.hollight.POWERSET_CLAUSES"
-  "PASSOC_def" > "HOLLight.hollight.PASSOC_def"
-  "PAIR_SURJECTIVE" > "Product_Type.prod.nchotomy"
-  "PAIR_EXISTS_THM" > "HOLLight.hollight.PAIR_EXISTS_THM"
-  "PAIR_EQ" > "Product_Type.Pair_eq"
-  "PAIRWISE_def" > "HOLLight.hollight.PAIRWISE_def"
-  "PAIRWISE_SING" > "HOLLight.hollight.PAIRWISE_SING"
-  "PAIRWISE_MONO" > "HOLLight.hollight.PAIRWISE_MONO"
-  "PAIRWISE_EMPTY" > "HOLLight.hollight.PAIRWISE_EMPTY"
-  "PAIR" > "HOLLightCompat.PAIR"
-  "OR_EXISTS_THM" > "HOL.ex_disj_distrib"
-  "OR_CLAUSES" > "HOLLight.hollight.OR_CLAUSES"
-  "ONE" > "Nat.One_nat_def"
-  "ODD_SUB" > "HOLLight.hollight.ODD_SUB"
-  "ODD_MULT" > "HOLLight.hollight.ODD_MULT"
-  "ODD_MOD" > "HOLLight.hollight.ODD_MOD"
-  "ODD_EXP" > "HOLLight.hollight.ODD_EXP"
-  "ODD_EXISTS" > "Parity.odd_Suc_mult_two_ex"
-  "ODD_DOUBLE" > "HOLLight.hollight.ODD_DOUBLE"
-  "ODD_ADD" > "Parity.odd_add"
-  "NUM_REP_def" > "HOLLight.hollight.NUM_REP_def"
-  "NUM_OF_INT_OF_NUM" > "HOLLight.hollight.NUM_OF_INT_OF_NUM"
-  "NUM_OF_INT" > "HOLLight.hollight.NUM_OF_INT"
-  "NUM_INTEGRAL_LEMMA" > "HOLLight.hollight.NUM_INTEGRAL_LEMMA"
-  "NUM_INTEGRAL" > "HOLLight.hollight.NUM_INTEGRAL"
-  "NUM_GCD" > "HOLLight.hollight.NUM_GCD"
-  "NUMSUM_def" > "HOLLight.hollight.NUMSUM_def"
-  "NUMSUM_INJ" > "HOLLight.hollight.NUMSUM_INJ"
-  "NUMSND_def" > "HOLLight.hollight.NUMSND_def"
-  "NUMSEG_SING" > "SetInterval.order_class.atLeastAtMost_singleton"
-  "NUMSEG_RREC" > "HOLLight.hollight.NUMSEG_RREC"
-  "NUMSEG_REC" > "SetInterval.atLeastAtMostSuc_conv"
-  "NUMSEG_OFFSET_IMAGE" > "SetInterval.image_add_atLeastAtMost"
-  "NUMSEG_LT" > "HOLLight.hollight.NUMSEG_LT"
-  "NUMSEG_LREC" > "HOLLight.hollight.NUMSEG_LREC"
-  "NUMSEG_LE" > "HOLLight.hollight.NUMSEG_LE"
-  "NUMSEG_EMPTY" > "HOLLight.hollight.NUMSEG_EMPTY"
-  "NUMSEG_COMBINE_R" > "HOLLight.hollight.NUMSEG_COMBINE_R"
-  "NUMSEG_COMBINE_L" > "HOLLight.hollight.NUMSEG_COMBINE_L"
-  "NUMSEG_CLAUSES" > "HOLLight.hollight.NUMSEG_CLAUSES"
-  "NUMSEG_ADD_SPLIT" > "HOLLight.hollight.NUMSEG_ADD_SPLIT"
-  "NUMRIGHT_def" > "HOLLight.hollight.NUMRIGHT_def"
-  "NUMPAIR_def" > "HOLLight.hollight.NUMPAIR_def"
-  "NUMPAIR_INJ_LEMMA" > "HOLLight.hollight.NUMPAIR_INJ_LEMMA"
-  "NUMPAIR_INJ" > "HOLLight.hollight.NUMPAIR_INJ"
-  "NUMLEFT_def" > "HOLLight.hollight.NUMLEFT_def"
-  "NUMFST_def" > "HOLLight.hollight.NUMFST_def"
-  "NSUM_UNION_RZERO" > "HOLLight.hollight.NSUM_UNION_RZERO"
-  "NSUM_UNION_NONZERO" > "HOLLight.hollight.NSUM_UNION_NONZERO"
-  "NSUM_UNION_LZERO" > "HOLLight.hollight.NSUM_UNION_LZERO"
-  "NSUM_UNION_EQ" > "HOLLight.hollight.NSUM_UNION_EQ"
-  "NSUM_UNIONS_NONZERO" > "HOLLight.hollight.NSUM_UNIONS_NONZERO"
-  "NSUM_UNION" > "HOLLight.hollight.NSUM_UNION"
-  "NSUM_TRIV_NUMSEG" > "HOLLight.hollight.NSUM_TRIV_NUMSEG"
-  "NSUM_SWAP_NUMSEG" > "HOLLight.hollight.NSUM_SWAP_NUMSEG"
-  "NSUM_SWAP" > "HOLLight.hollight.NSUM_SWAP"
-  "NSUM_SUPPORT" > "HOLLight.hollight.NSUM_SUPPORT"
-  "NSUM_SUPERSET" > "HOLLight.hollight.NSUM_SUPERSET"
-  "NSUM_SUBSET_SIMPLE" > "HOLLight.hollight.NSUM_SUBSET_SIMPLE"
-  "NSUM_SUBSET" > "HOLLight.hollight.NSUM_SUBSET"
-  "NSUM_SING_NUMSEG" > "HOLLight.hollight.NSUM_SING_NUMSEG"
-  "NSUM_SING" > "HOLLight.hollight.NSUM_SING"
-  "NSUM_RMUL" > "HOLLight.hollight.NSUM_RMUL"
-  "NSUM_RESTRICT_SET" > "HOLLight.hollight.NSUM_RESTRICT_SET"
-  "NSUM_RESTRICT" > "HOLLight.hollight.NSUM_RESTRICT"
-  "NSUM_POS_BOUND" > "HOLLight.hollight.NSUM_POS_BOUND"
-  "NSUM_PAIR" > "HOLLight.hollight.NSUM_PAIR"
-  "NSUM_OFFSET_0" > "HOLLight.hollight.NSUM_OFFSET_0"
-  "NSUM_OFFSET" > "HOLLight.hollight.NSUM_OFFSET"
-  "NSUM_NSUM_RESTRICT" > "HOLLight.hollight.NSUM_NSUM_RESTRICT"
-  "NSUM_NSUM_PRODUCT" > "HOLLight.hollight.NSUM_NSUM_PRODUCT"
-  "NSUM_MULTICOUNT_GEN" > "HOLLight.hollight.NSUM_MULTICOUNT_GEN"
-  "NSUM_MULTICOUNT" > "HOLLight.hollight.NSUM_MULTICOUNT"
-  "NSUM_LT_ALL" > "HOLLight.hollight.NSUM_LT_ALL"
-  "NSUM_LT" > "HOLLight.hollight.NSUM_LT"
-  "NSUM_LMUL" > "HOLLight.hollight.NSUM_LMUL"
-  "NSUM_LE_NUMSEG" > "HOLLight.hollight.NSUM_LE_NUMSEG"
-  "NSUM_LE" > "HOLLight.hollight.NSUM_LE"
-  "NSUM_INJECTION" > "HOLLight.hollight.NSUM_INJECTION"
-  "NSUM_INCL_EXCL" > "HOLLight.hollight.NSUM_INCL_EXCL"
-  "NSUM_IMAGE_NONZERO" > "HOLLight.hollight.NSUM_IMAGE_NONZERO"
-  "NSUM_IMAGE_GEN" > "HOLLight.hollight.NSUM_IMAGE_GEN"
-  "NSUM_IMAGE" > "HOLLight.hollight.NSUM_IMAGE"
-  "NSUM_GROUP" > "HOLLight.hollight.NSUM_GROUP"
-  "NSUM_EQ_SUPERSET" > "HOLLight.hollight.NSUM_EQ_SUPERSET"
-  "NSUM_EQ_NUMSEG" > "HOLLight.hollight.NSUM_EQ_NUMSEG"
-  "NSUM_EQ_GENERAL_INVERSES" > "HOLLight.hollight.NSUM_EQ_GENERAL_INVERSES"
-  "NSUM_EQ_GENERAL" > "HOLLight.hollight.NSUM_EQ_GENERAL"
-  "NSUM_EQ_0_NUMSEG" > "HOLLight.hollight.NSUM_EQ_0_NUMSEG"
-  "NSUM_EQ_0_IFF_NUMSEG" > "HOLLight.hollight.NSUM_EQ_0_IFF_NUMSEG"
-  "NSUM_EQ_0_IFF" > "HOLLight.hollight.NSUM_EQ_0_IFF"
-  "NSUM_EQ_0" > "HOLLight.hollight.NSUM_EQ_0"
-  "NSUM_EQ" > "HOLLight.hollight.NSUM_EQ"
-  "NSUM_DIFF" > "HOLLight.hollight.NSUM_DIFF"
-  "NSUM_DELTA" > "HOLLight.hollight.NSUM_DELTA"
-  "NSUM_DELETE" > "HOLLight.hollight.NSUM_DELETE"
-  "NSUM_CONST_NUMSEG" > "HOLLight.hollight.NSUM_CONST_NUMSEG"
-  "NSUM_CONST" > "HOLLight.hollight.NSUM_CONST"
-  "NSUM_CLOSED" > "HOLLight.hollight.NSUM_CLOSED"
-  "NSUM_CLAUSES_RIGHT" > "HOLLight.hollight.NSUM_CLAUSES_RIGHT"
-  "NSUM_CLAUSES_NUMSEG" > "HOLLight.hollight.NSUM_CLAUSES_NUMSEG"
-  "NSUM_CLAUSES_LEFT" > "HOLLight.hollight.NSUM_CLAUSES_LEFT"
-  "NSUM_CLAUSES" > "HOLLight.hollight.NSUM_CLAUSES"
-  "NSUM_CASES" > "HOLLight.hollight.NSUM_CASES"
-  "NSUM_BOUND_LT_GEN" > "HOLLight.hollight.NSUM_BOUND_LT_GEN"
-  "NSUM_BOUND_LT_ALL" > "HOLLight.hollight.NSUM_BOUND_LT_ALL"
-  "NSUM_BOUND_LT" > "HOLLight.hollight.NSUM_BOUND_LT"
-  "NSUM_BOUND_GEN" > "HOLLight.hollight.NSUM_BOUND_GEN"
-  "NSUM_BOUND" > "HOLLight.hollight.NSUM_BOUND"
-  "NSUM_BIJECTION" > "HOLLight.hollight.NSUM_BIJECTION"
-  "NSUM_ADD_SPLIT" > "HOLLight.hollight.NSUM_ADD_SPLIT"
-  "NSUM_ADD_NUMSEG" > "HOLLight.hollight.NSUM_ADD_NUMSEG"
-  "NSUM_ADD_GEN" > "HOLLight.hollight.NSUM_ADD_GEN"
-  "NSUM_ADD" > "HOLLight.hollight.NSUM_ADD"
-  "NSUM_0" > "HOLLight.hollight.NSUM_0"
-  "NOT_UNIV_PSUBSET" > "Orderings.top_class.not_top_less"
-  "NOT_SUC" > "Nat.Suc_not_Zero"
-  "NOT_PSUBSET_EMPTY" > "Orderings.bot_class.not_less_bot"
-  "NOT_ODD" > "HOLLight.hollight.NOT_ODD"
-  "NOT_LT" > "Orderings.linorder_class.not_less"
-  "NOT_LE" > "Orderings.linorder_class.not_le"
-  "NOT_IN_EMPTY" > "HOLLight.hollight.NOT_IN_EMPTY"
-  "NOT_INSERT_EMPTY" > "Set.insert_not_empty"
-  "NOT_FORALL_THM" > "HOL.not_all"
-  "NOT_EXISTS_THM" > "HOL.not_ex"
-  "NOT_EX" > "HOLLightList.NOT_EX"
-  "NOT_EVEN" > "HOLLight.hollight.NOT_EVEN"
-  "NOT_EQUAL_SETS" > "HOLLight.hollight.NOT_EQUAL_SETS"
-  "NOT_EMPTY_INSERT" > "Set.empty_not_insert"
-  "NOT_CONS_NIL" > "List.list.distinct_2"
-  "NOT_CLAUSES_WEAK" > "HOLLight.hollight.NOT_CLAUSES_WEAK"
-  "NOT_ALL" > "HOLLightList.NOT_ALL"
-  "NEUTRAL_REAL_MUL" > "HOLLight.hollight.NEUTRAL_REAL_MUL"
-  "NEUTRAL_REAL_ADD" > "HOLLight.hollight.NEUTRAL_REAL_ADD"
-  "NEUTRAL_MUL" > "HOLLight.hollight.NEUTRAL_MUL"
-  "NEUTRAL_ADD" > "HOLLight.hollight.NEUTRAL_ADD"
-  "NADD_UBOUND" > "HOLLight.hollight.NADD_UBOUND"
-  "NADD_SUC" > "HOLLight.hollight.NADD_SUC"
-  "NADD_RDISTRIB" > "HOLLight.hollight.NADD_RDISTRIB"
-  "NADD_OF_NUM_WELLDEF" > "HOLLight.hollight.NADD_OF_NUM_WELLDEF"
-  "NADD_OF_NUM_MUL" > "HOLLight.hollight.NADD_OF_NUM_MUL"
-  "NADD_OF_NUM_LE" > "HOLLight.hollight.NADD_OF_NUM_LE"
-  "NADD_OF_NUM_EQ" > "HOLLight.hollight.NADD_OF_NUM_EQ"
-  "NADD_OF_NUM_ADD" > "HOLLight.hollight.NADD_OF_NUM_ADD"
-  "NADD_OF_NUM" > "HOLLight.hollight.NADD_OF_NUM"
-  "NADD_NONZERO" > "HOLLight.hollight.NADD_NONZERO"
-  "NADD_MUL_WELLDEF_LEMMA" > "HOLLight.hollight.NADD_MUL_WELLDEF_LEMMA"
-  "NADD_MUL_WELLDEF" > "HOLLight.hollight.NADD_MUL_WELLDEF"
-  "NADD_MUL_SYM" > "HOLLight.hollight.NADD_MUL_SYM"
-  "NADD_MUL_LINV_LEMMA8" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA8"
-  "NADD_MUL_LINV_LEMMA7a" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA7a"
-  "NADD_MUL_LINV_LEMMA7" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA7"
-  "NADD_MUL_LINV_LEMMA6" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA6"
-  "NADD_MUL_LINV_LEMMA5" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA5"
-  "NADD_MUL_LINV_LEMMA4" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA4"
-  "NADD_MUL_LINV_LEMMA3" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA3"
-  "NADD_MUL_LINV_LEMMA2" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA2"
-  "NADD_MUL_LINV_LEMMA1" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA1"
-  "NADD_MUL_LINV_LEMMA0" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA0"
-  "NADD_MUL_LINV" > "HOLLight.hollight.NADD_MUL_LINV"
-  "NADD_MUL_LID" > "HOLLight.hollight.NADD_MUL_LID"
-  "NADD_MUL_ASSOC" > "HOLLight.hollight.NADD_MUL_ASSOC"
-  "NADD_MULTIPLICATIVE" > "HOLLight.hollight.NADD_MULTIPLICATIVE"
-  "NADD_MUL" > "HOLLight.hollight.NADD_MUL"
-  "NADD_LE_WELLDEF_LEMMA" > "HOLLight.hollight.NADD_LE_WELLDEF_LEMMA"
-  "NADD_LE_WELLDEF" > "HOLLight.hollight.NADD_LE_WELLDEF"
-  "NADD_LE_TRANS" > "HOLLight.hollight.NADD_LE_TRANS"
-  "NADD_LE_TOTAL_LEMMA" > "HOLLight.hollight.NADD_LE_TOTAL_LEMMA"
-  "NADD_LE_TOTAL" > "HOLLight.hollight.NADD_LE_TOTAL"
-  "NADD_LE_RMUL" > "HOLLight.hollight.NADD_LE_RMUL"
-  "NADD_LE_REFL" > "HOLLight.hollight.NADD_LE_REFL"
-  "NADD_LE_RADD" > "HOLLight.hollight.NADD_LE_RADD"
-  "NADD_LE_LMUL" > "HOLLight.hollight.NADD_LE_LMUL"
-  "NADD_LE_LADD" > "HOLLight.hollight.NADD_LE_LADD"
-  "NADD_LE_EXISTS" > "HOLLight.hollight.NADD_LE_EXISTS"
-  "NADD_LE_ANTISYM" > "HOLLight.hollight.NADD_LE_ANTISYM"
-  "NADD_LE_ADD" > "HOLLight.hollight.NADD_LE_ADD"
-  "NADD_LE_0" > "HOLLight.hollight.NADD_LE_0"
-  "NADD_LDISTRIB" > "HOLLight.hollight.NADD_LDISTRIB"
-  "NADD_LBOUND" > "HOLLight.hollight.NADD_LBOUND"
-  "NADD_INV_WELLDEF" > "HOLLight.hollight.NADD_INV_WELLDEF"
-  "NADD_INV_0" > "HOLLight.hollight.NADD_INV_0"
-  "NADD_INV" > "HOLLight.hollight.NADD_INV"
-  "NADD_EQ_TRANS" > "HOLLight.hollight.NADD_EQ_TRANS"
-  "NADD_EQ_SYM" > "HOLLight.hollight.NADD_EQ_SYM"
-  "NADD_EQ_REFL" > "HOLLight.hollight.NADD_EQ_REFL"
-  "NADD_EQ_IMP_LE" > "HOLLight.hollight.NADD_EQ_IMP_LE"
-  "NADD_DIST_LEMMA" > "HOLLight.hollight.NADD_DIST_LEMMA"
-  "NADD_DIST" > "HOLLight.hollight.NADD_DIST"
-  "NADD_COMPLETE" > "HOLLight.hollight.NADD_COMPLETE"
-  "NADD_CAUCHY" > "HOLLight.hollight.NADD_CAUCHY"
-  "NADD_BOUND" > "HOLLight.hollight.NADD_BOUND"
-  "NADD_ARCH_ZERO" > "HOLLight.hollight.NADD_ARCH_ZERO"
-  "NADD_ARCH_MULT" > "HOLLight.hollight.NADD_ARCH_MULT"
-  "NADD_ARCH_LEMMA" > "HOLLight.hollight.NADD_ARCH_LEMMA"
-  "NADD_ARCH" > "HOLLight.hollight.NADD_ARCH"
-  "NADD_ALTMUL" > "HOLLight.hollight.NADD_ALTMUL"
-  "NADD_ADD_WELLDEF" > "HOLLight.hollight.NADD_ADD_WELLDEF"
-  "NADD_ADD_SYM" > "HOLLight.hollight.NADD_ADD_SYM"
-  "NADD_ADD_LID" > "HOLLight.hollight.NADD_ADD_LID"
-  "NADD_ADD_LCANCEL" > "HOLLight.hollight.NADD_ADD_LCANCEL"
-  "NADD_ADD_ASSOC" > "HOLLight.hollight.NADD_ADD_ASSOC"
-  "NADD_ADDITIVE" > "HOLLight.hollight.NADD_ADDITIVE"
-  "NADD_ADD" > "HOLLight.hollight.NADD_ADD"
-  "MULT_SYM" > "Fields.linordered_field_class.sign_simps_40"
-  "MULT_SUC" > "Nat.mult_Suc_right"
-  "MULT_EXP" > "Power.comm_monoid_mult_class.power_mult_distrib"
-  "MULT_EQ_1" > "Nat.nat_mult_eq_1_iff"
-  "MULT_EQ_0" > "Nat.mult_is_0"
-  "MULT_DIV_LE" > "HOLLight.hollight.MULT_DIV_LE"
-  "MULT_CLAUSES" > "HOLLight.hollight.MULT_CLAUSES"
-  "MULT_ASSOC" > "Fields.linordered_field_class.sign_simps_41"
-  "MULT_AC" > "HOLLight.hollight.MULT_AC"
-  "MULT_2" > "Int.semiring_mult_2"
-  "MULT_0" > "Divides.arithmetic_simps_41"
-  "MONO_FORALL" > "Inductive.basic_monos_6"
-  "MONO_EXISTS" > "Inductive.basic_monos_5"
-  "MONO_COND" > "HOLLight.hollight.MONO_COND"
-  "MONO_ALL2" > "List.list_all2_mono"
-  "MONO_ALL" > "HOLLightList.MONO_ALL"
-  "MONOIDAL_REAL_MUL" > "HOLLight.hollight.MONOIDAL_REAL_MUL"
-  "MONOIDAL_REAL_ADD" > "HOLLight.hollight.MONOIDAL_REAL_ADD"
-  "MONOIDAL_MUL" > "HOLLight.hollight.MONOIDAL_MUL"
-  "MONOIDAL_ADD" > "HOLLight.hollight.MONOIDAL_ADD"
-  "MONOIDAL_AC" > "HOLLight.hollight.MONOIDAL_AC"
-  "MOD_UNIQ" > "HOLLight.hollight.MOD_UNIQ"
-  "MOD_MULT_RMOD" > "HOLLight.hollight.MOD_MULT_RMOD"
-  "MOD_MULT_MOD2" > "HOLLight.hollight.MOD_MULT_MOD2"
-  "MOD_MULT_LMOD" > "HOLLight.hollight.MOD_MULT_LMOD"
-  "MOD_MULT_ADD" > "Divides.mod_mult_self3"
-  "MOD_MULT2" > "HOLLight.hollight.MOD_MULT2"
-  "MOD_MOD_REFL" > "HOLLight.hollight.MOD_MOD_REFL"
-  "MOD_MOD" > "HOLLight.hollight.MOD_MOD"
-  "MOD_LT" > "Divides.mod_less"
-  "MOD_LE" > "HOLLight.hollight.MOD_LE"
-  "MOD_EXP_MOD" > "HOLLight.hollight.MOD_EXP_MOD"
-  "MOD_EXISTS" > "HOLLight.hollight.MOD_EXISTS"
-  "MOD_EQ_0" > "HOLLight.hollight.MOD_EQ_0"
-  "MOD_EQ" > "HOLLight.hollight.MOD_EQ"
-  "MOD_ADD_MOD" > "HOLLight.hollight.MOD_ADD_MOD"
-  "MK_REC_INJ" > "HOLLight.hollight.MK_REC_INJ"
-  "MINIMAL" > "HOLLight.hollight.MINIMAL"
-  "MEM_MAP" > "HOLLightList.MEM_MAP"
-  "MEM_FILTER" > "HOLLightList.MEM_FILTER"
-  "MEM_EXISTS_EL" > "HOLLightList.MEM_EXISTS_EL"
-  "MEM_EL" > "List.nth_mem"
-  "MEM_APPEND" > "HOLLightList.MEM_APPEND"
-  "MEMBER_NOT_EMPTY" > "Set.ex_in_conv"
-  "MEASURE_LE" > "HOLLight.hollight.MEASURE_LE"
-  "MATCH_SEQPATTERN" > "HOLLight.hollight.MATCH_SEQPATTERN"
-  "MAP_o" > "List.map.compositionality"
-  "MAP_SND_ZIP" > "List.map_snd_zip"
-  "MAP_ID" > "List.map_ident"
-  "MAP_I" > "List.map.id"
-  "MAP_FST_ZIP" > "List.map_fst_zip"
-  "MAP_EQ_NIL" > "List.map_is_Nil_conv"
-  "MAP_EQ_DEGEN" > "HOLLightList.MAP_EQ_DEGEN"
-  "MAP_EQ_ALL2" > "HOLLightList.MAP_EQ_ALL2"
-  "MAP_EQ" > "HOLLightList.MAP_EQ"
-  "MAP_APPEND" > "List.map_append"
-  "MAP2" > "HOLLightList.MAP2"
-  "LT_TRANS" > "Orderings.order_less_trans"
-  "LT_SUC_LE" > "Nat.le_simps_2"
-  "LT_SUC" > "Nat.Suc_less_eq"
-  "LT_REFL" > "Nat.less_not_refl"
-  "LT_NZ" > "Nat.neq0_conv"
-  "LT_MULT_RCANCEL" > "HOLLight.hollight.LT_MULT_RCANCEL"
-  "LT_MULT_LCANCEL" > "HOLLight.hollight.LT_MULT_LCANCEL"
-  "LT_MULT2" > "HOLLight.hollight.LT_MULT2"
-  "LT_MULT" > "Nat.nat_0_less_mult_iff"
-  "LT_LMULT" > "HOLLight.hollight.LT_LMULT"
-  "LT_LE" > "Nat.nat_less_le"
-  "LT_IMP_LE" > "FunDef.termination_basic_simps_5"
-  "LT_EXP" > "HOLLight.hollight.LT_EXP"
-  "LT_EXISTS" > "HOLLight.hollight.LT_EXISTS"
-  "LT_CASES" > "HOLLight.hollight.LT_CASES"
-  "LT_ANTISYM" > "HOLLight.hollight.LT_ANTISYM"
-  "LT_ADD_RCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_right"
-  "LT_ADD_LCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_left"
-  "LT_ADDR" > "HOLLight.hollight.LT_ADDR"
-  "LT_ADD2" > "Groups.add_mono_thms_linordered_field_5"
-  "LT_ADD" > "HOLLight.hollight.LT_ADD"
-  "LT_0" > "Nat.zero_less_Suc"
-  "LTE_TRANS" > "Orderings.order_less_le_trans"
-  "LTE_CASES" > "HOLLight.hollight.LTE_CASES"
-  "LTE_ANTISYM" > "HOLLight.hollight.LTE_ANTISYM"
-  "LTE_ADD2" > "Groups.add_mono_thms_linordered_field_3"
-  "LE_TRANS" > "Nat.le_trans"
-  "LE_SUC_LT" > "Nat.Suc_le_eq"
-  "LE_SUC" > "Nat.Suc_le_mono"
-  "LE_SQUARE_REFL" > "Nat.le_square"
-  "LE_REFL" > "Nat.le_refl"
-  "LE_RDIV_EQ" > "HOLLight.hollight.LE_RDIV_EQ"
-  "LE_MULT_RCANCEL" > "HOLLight.hollight.LE_MULT_RCANCEL"
-  "LE_MULT_LCANCEL" > "HOLLight.hollight.LE_MULT_LCANCEL"
-  "LE_MULT2" > "Nat.mult_le_mono"
-  "LE_LT" > "Nat.le_eq_less_or_eq"
-  "LE_LDIV_EQ" > "HOLLight.hollight.LE_LDIV_EQ"
-  "LE_LDIV" > "HOLLight.hollight.LE_LDIV"
-  "LE_EXP" > "HOLLight.hollight.LE_EXP"
-  "LE_EXISTS" > "Nat.le_iff_add"
-  "LE_CASES" > "Nat.nat_le_linear"
-  "LE_C" > "HOLLight.hollight.LE_C"
-  "LE_ANTISYM" > "Orderings.order_class.eq_iff"
-  "LE_ADD_RCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_right"
-  "LE_ADD_LCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left"
-  "LE_ADDR" > "Nat.le_add2"
-  "LE_ADD2" > "Groups.add_mono_thms_linordered_semiring_1"
-  "LE_ADD" > "Nat.le_add1"
-  "LE_1" > "HOLLight.hollight.LE_1"
-  "LE_0" > "Nat.le0"
-  "LET_TRANS" > "Orderings.order_le_less_trans"
-  "LET_END_def" > "HOLLight.hollight.LET_END_def"
-  "LET_CASES" > "Orderings.linorder_class.le_less_linear"
-  "LET_ANTISYM" > "HOLLight.hollight.LET_ANTISYM"
-  "LET_ADD2" > "Groups.add_mono_thms_linordered_field_4"
-  "LENGTH_TL" > "HOLLightList.LENGTH_TL"
-  "LENGTH_REPLICATE" > "List.length_replicate"
-  "LENGTH_MAP2" > "HOLLightList.LENGTH_MAP2"
-  "LENGTH_MAP" > "List.length_map"
-  "LENGTH_EQ_NIL" > "List.length_0_conv"
-  "LENGTH_EQ_CONS" > "List.length_Suc_conv"
-  "LENGTH_APPEND" > "List.length_append"
-  "LEFT_SUB_DISTRIB" > "Nat.diff_mult_distrib2"
-  "LEFT_OR_FORALL_THM" > "HOL.all_simps_3"
-  "LEFT_OR_EXISTS_THM" > "HOL.ex_simps_3"
-  "LEFT_OR_DISTRIB" > "Groebner_Basis.dnf_1"
-  "LEFT_IMP_FORALL_THM" > "HOL.ex_simps_5"
-  "LEFT_IMP_EXISTS_THM" > "HOL.all_simps_5"
-  "LEFT_FORALL_OR_THM" > "HOL.all_simps_3"
-  "LEFT_FORALL_IMP_THM" > "HOL.all_simps_5"
-  "LEFT_EXISTS_IMP_THM" > "HOL.ex_simps_5"
-  "LEFT_EXISTS_AND_THM" > "HOL.ex_simps_1"
-  "LEFT_AND_FORALL_THM" > "HOL.all_simps_1"
-  "LEFT_AND_EXISTS_THM" > "HOL.ex_simps_1"
-  "LEFT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_25"
-  "LAST_EL" > "List.last_conv_nth"
-  "LAST_CLAUSES" > "HOLLightList.LAST_CLAUSES"
-  "LAST_APPEND" > "List.last_append"
-  "LAMBDA_UNIQUE" > "HOLLight.hollight.LAMBDA_UNIQUE"
-  "LAMBDA_PAIR_THM" > "HOLLight.hollight.LAMBDA_PAIR_THM"
-  "LAMBDA_ETA" > "HOLLight.hollight.LAMBDA_ETA"
-  "LAMBDA_BETA" > "HOLLight.hollight.LAMBDA_BETA"
-  "I_THM" > "HOL.refl"
-  "I_O_ID" > "HOLLight.hollight.I_O_ID"
-  "ITSET_def" > "HOLLight.hollight.ITSET_def"
-  "ITSET_EQ" > "HOLLight.hollight.ITSET_EQ"
-  "ITLIST_EXTRA" > "HOLLightList.ITLIST_EXTRA"
-  "ITLIST_APPEND" > "List.foldr_append"
-  "ITLIST2" > "HOLLightList.ITLIST2"
-  "ITERATE_UNION_NONZERO" > "HOLLight.hollight.ITERATE_UNION_NONZERO"
-  "ITERATE_UNION_GEN" > "HOLLight.hollight.ITERATE_UNION_GEN"
-  "ITERATE_UNION" > "HOLLight.hollight.ITERATE_UNION"
-  "ITERATE_SUPPORT" > "HOLLight.hollight.ITERATE_SUPPORT"
-  "ITERATE_SUPERSET" > "HOLLight.hollight.ITERATE_SUPERSET"
-  "ITERATE_SING" > "HOLLight.hollight.ITERATE_SING"
-  "ITERATE_RELATED" > "HOLLight.hollight.ITERATE_RELATED"
-  "ITERATE_PAIR" > "HOLLight.hollight.ITERATE_PAIR"
-  "ITERATE_OP_GEN" > "HOLLight.hollight.ITERATE_OP_GEN"
-  "ITERATE_OP" > "HOLLight.hollight.ITERATE_OP"
-  "ITERATE_ITERATE_PRODUCT" > "HOLLight.hollight.ITERATE_ITERATE_PRODUCT"
-  "ITERATE_INJECTION" > "HOLLight.hollight.ITERATE_INJECTION"
-  "ITERATE_INCL_EXCL" > "HOLLight.hollight.ITERATE_INCL_EXCL"
-  "ITERATE_IMAGE_NONZERO" > "HOLLight.hollight.ITERATE_IMAGE_NONZERO"
-  "ITERATE_IMAGE" > "HOLLight.hollight.ITERATE_IMAGE"
-  "ITERATE_EXPAND_CASES" > "HOLLight.hollight.ITERATE_EXPAND_CASES"
-  "ITERATE_EQ_NEUTRAL" > "HOLLight.hollight.ITERATE_EQ_NEUTRAL"
-  "ITERATE_EQ_GENERAL_INVERSES" > "HOLLight.hollight.ITERATE_EQ_GENERAL_INVERSES"
-  "ITERATE_EQ_GENERAL" > "HOLLight.hollight.ITERATE_EQ_GENERAL"
-  "ITERATE_EQ" > "HOLLight.hollight.ITERATE_EQ"
-  "ITERATE_DIFF_GEN" > "HOLLight.hollight.ITERATE_DIFF_GEN"
-  "ITERATE_DIFF" > "HOLLight.hollight.ITERATE_DIFF"
-  "ITERATE_DELTA" > "HOLLight.hollight.ITERATE_DELTA"
-  "ITERATE_DELETE" > "HOLLight.hollight.ITERATE_DELETE"
-  "ITERATE_CLOSED" > "HOLLight.hollight.ITERATE_CLOSED"
-  "ITERATE_CLAUSES_NUMSEG" > "HOLLight.hollight.ITERATE_CLAUSES_NUMSEG"
-  "ITERATE_CLAUSES_GEN" > "HOLLight.hollight.ITERATE_CLAUSES_GEN"
-  "ITERATE_CLAUSES" > "HOLLight.hollight.ITERATE_CLAUSES"
-  "ITERATE_CASES" > "HOLLight.hollight.ITERATE_CASES"
-  "ITERATE_BIJECTION" > "HOLLight.hollight.ITERATE_BIJECTION"
-  "ISO_def" > "HOLLight.hollight.ISO_def"
-  "ISO_USAGE" > "HOLLight.hollight.ISO_USAGE"
-  "ISO_REFL" > "HOLLight.hollight.ISO_REFL"
-  "ISO_FUN" > "HOLLight.hollight.ISO_FUN"
-  "IN_UNIV" > "Set.UNIV_I"
-  "IN_UNIONS" > "HOLLight.hollight.IN_UNIONS"
-  "IN_UNION" > "Complete_Lattices.mem_simps_3"
-  "IN_SUPPORT" > "HOLLight.hollight.IN_SUPPORT"
-  "IN_SING" > "Set.singleton_iff"
-  "IN_SET_OF_LIST" > "HOLLightList.IN_SET_OF_LIST"
-  "IN_REST" > "HOLLight.hollight.IN_REST"
-  "IN_NUMSEG_0" > "HOLLight.hollight.IN_NUMSEG_0"
-  "IN_NUMSEG" > "SetInterval.ord_class.atLeastAtMost_iff"
-  "IN_INTERS" > "HOLLight.hollight.IN_INTERS"
-  "IN_INTER" > "Complete_Lattices.mem_simps_4"
-  "IN_INSERT" > "Complete_Lattices.mem_simps_1"
-  "IN_IMAGE" > "HOLLight.hollight.IN_IMAGE"
-  "IN_ELIM_THM" > "HOLLight.hollight.IN_ELIM_THM"
-  "IN_ELIM_PAIR_THM" > "HOLLight.hollight.IN_ELIM_PAIR_THM"
-  "IN_DISJOINT" > "HOLLight.hollight.IN_DISJOINT"
-  "IN_DIFF" > "Complete_Lattices.mem_simps_6"
-  "IN_DELETE_EQ" > "HOLLight.hollight.IN_DELETE_EQ"
-  "IN_DELETE" > "HOLLight.hollight.IN_DELETE"
-  "IN_CROSS" > "HOLLight.hollight.IN_CROSS"
-  "INT_WOP" > "HOLLight.hollight.INT_WOP"
-  "INT_POW" > "HOLLight.hollight.INT_POW"
-  "INT_OF_NUM_OF_INT" > "HOLLight.hollight.INT_OF_NUM_OF_INT"
-  "INT_LT_DISCRETE" > "HOLLight.hollight.INT_LT_DISCRETE"
-  "INT_LT" > "HOLLight.hollight.INT_LT"
-  "INT_INTEGRAL" > "HOLLight.hollight.INT_INTEGRAL"
-  "INT_IMAGE" > "HOLLight.hollight.INT_IMAGE"
-  "INT_GT_DISCRETE" > "HOLLight.hollight.INT_GT_DISCRETE"
-  "INT_GT" > "HOLLight.hollight.INT_GT"
-  "INT_GE" > "HOLLight.hollight.INT_GE"
-  "INT_GCD_EXISTS_POS" > "HOLLight.hollight.INT_GCD_EXISTS_POS"
-  "INT_GCD_EXISTS" > "HOLLight.hollight.INT_GCD_EXISTS"
-  "INT_FORALL_POS" > "HOLLight.hollight.INT_FORALL_POS"
-  "INT_FORALL_ABS" > "HOLLight.hollight.INT_FORALL_ABS"
-  "INT_EXISTS_POS" > "HOLLight.hollight.INT_EXISTS_POS"
-  "INT_EXISTS_ABS" > "HOLLight.hollight.INT_EXISTS_ABS"
-  "INT_DIVMOD_UNIQ" > "HOLLight.hollight.INT_DIVMOD_UNIQ"
-  "INT_DIVMOD_EXIST_0" > "HOLLight.hollight.INT_DIVMOD_EXIST_0"
-  "INT_DIVISION" > "HOLLight.hollight.INT_DIVISION"
-  "INT_ARCH" > "HOLLight.hollight.INT_ARCH"
-  "INT_ABS_MUL_1" > "HOLLight.hollight.INT_ABS_MUL_1"
-  "INT_ABS" > "HOLLight.hollight.INT_ABS"
-  "INTER_UNIV" > "HOLLight.hollight.INTER_UNIV"
-  "INTER_UNIONS" > "HOLLight.hollight.INTER_UNIONS"
-  "INTER_SUBSET" > "HOLLight.hollight.INTER_SUBSET"
-  "INTER_OVER_UNION" > "Lattices.distrib_lattice_class.distrib_1"
-  "INTER_IDEMPOT" > "Big_Operators.lattice_class.Inf_fin.idem"
-  "INTER_EMPTY" > "HOLLight.hollight.INTER_EMPTY"
-  "INTER_COMM" > "Lattices.lattice_class.inf_sup_aci_1"
-  "INTER_ASSOC" > "Lattices.lattice_class.inf_sup_aci_2"
-  "INTER_ACI" > "HOLLight.hollight.INTER_ACI"
-  "INTERS_UNIONS" > "HOLLight.hollight.INTERS_UNIONS"
-  "INTERS_INSERT" > "Complete_Lattices.Inter_insert"
-  "INTERS_IMAGE" > "HOLLight.hollight.INTERS_IMAGE"
-  "INTERS_GSPEC" > "HOLLight.hollight.INTERS_GSPEC"
-  "INTERS_2" > "Complete_Lattices.Int_eq_Inter"
-  "INTERS_1" > "Complete_Lattices.complete_lattice_class.Inf_singleton"
-  "INTERS_0" > "Complete_Lattices.Inter_empty"
-  "INSERT_UNIV" > "HOLLight.hollight.INSERT_UNIV"
-  "INSERT_UNION_EQ" > "Set.Un_insert_left"
-  "INSERT_UNION" > "HOLLight.hollight.INSERT_UNION"
-  "INSERT_SUBSET" > "Set.insert_subset"
-  "INSERT_INTER" > "Set.Int_insert_left"
-  "INSERT_INSERT" > "Set.insert_absorb2"
-  "INSERT_DIFF" > "Set.insert_Diff_if"
-  "INSERT_DELETE" > "Set.insert_Diff"
-  "INSERT_COMM" > "Set.insert_commute"
-  "INSERT_AC" > "HOLLight.hollight.INSERT_AC"
-  "INSERT" > "HOLLight.hollight.INSERT"
-  "INJ_def" > "HOLLight.hollight.INJ_def"
-  "INJ_INVERSE2" > "HOLLight.hollight.INJ_INVERSE2"
-  "INJP_def" > "HOLLight.hollight.INJP_def"
-  "INJP_INJ" > "HOLLight.hollight.INJP_INJ"
-  "INJN_def" > "HOLLight.hollight.INJN_def"
-  "INJN_INJ" > "HOLLight.hollight.INJN_INJ"
-  "INJF_def" > "HOLLight.hollight.INJF_def"
-  "INJF_INJ" > "HOLLight.hollight.INJF_INJ"
-  "INJECTIVE_ON_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_ON_LEFT_INVERSE"
-  "INJECTIVE_ON_IMAGE" > "HOLLight.hollight.INJECTIVE_ON_IMAGE"
-  "INJECTIVE_MAP" > "HOLLightList.INJECTIVE_MAP"
-  "INJECTIVE_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_LEFT_INVERSE"
-  "INJECTIVE_IMAGE" > "HOLLight.hollight.INJECTIVE_IMAGE"
-  "INJA_def" > "HOLLight.hollight.INJA_def"
-  "INJA_INJ" > "HOLLight.hollight.INJA_INJ"
-  "INFINITE_NONEMPTY" > "Infinite_Set.infinite_imp_nonempty"
-  "INFINITE_IMAGE_INJ" > "HOLLight.hollight.INFINITE_IMAGE_INJ"
-  "INFINITE_DIFF_FINITE" > "Infinite_Set.Diff_infinite_finite"
-  "IND_SUC_def" > "HOLLight.hollight.IND_SUC_def"
-  "IND_SUC_0_EXISTS" > "HOLLight.hollight.IND_SUC_0_EXISTS"
-  "IND_0_def" > "HOLLight.hollight.IND_0_def"
-  "IMP_EQ_CLAUSE" > "HOLLight.hollight.IMP_EQ_CLAUSE"
-  "IMP_CONJ_ALT" > "HOLLight.hollight.IMP_CONJ_ALT"
-  "IMP_CONJ" > "HOL.imp_conjL"
-  "IMP_CLAUSES" > "HOLLight.hollight.IMP_CLAUSES"
-  "IMAGE_o" > "Fun.image_compose"
-  "IMAGE_UNIONS" > "HOLLight.hollight.IMAGE_UNIONS"
-  "IMAGE_UNION" > "Set.image_Un"
-  "IMAGE_SUBSET" > "Set.image_mono"
-  "IMAGE_INTER_INJ" > "HOLLight.hollight.IMAGE_INTER_INJ"
-  "IMAGE_INJECTIVE_IMAGE_OF_SUBSET" > "HOLLight.hollight.IMAGE_INJECTIVE_IMAGE_OF_SUBSET"
-  "IMAGE_IMP_INJECTIVE_GEN" > "HOLLight.hollight.IMAGE_IMP_INJECTIVE_GEN"
-  "IMAGE_IMP_INJECTIVE" > "HOLLight.hollight.IMAGE_IMP_INJECTIVE"
-  "IMAGE_ID" > "Set.image_ident"
-  "IMAGE_I" > "Fun.image_id"
-  "IMAGE_EQ_EMPTY" > "Set.image_is_empty"
-  "IMAGE_DIFF_INJ" > "HOLLight.hollight.IMAGE_DIFF_INJ"
-  "IMAGE_DELETE_INJ" > "HOLLight.hollight.IMAGE_DELETE_INJ"
-  "IMAGE_CONST" > "Set.image_constant_conv"
-  "IMAGE_CLAUSES" > "HOLLight.hollight.IMAGE_CLAUSES"
-  "HREAL_MUL_RZERO" > "HOLLight.hollight.HREAL_MUL_RZERO"
-  "HREAL_MUL_LZERO" > "HOLLight.hollight.HREAL_MUL_LZERO"
-  "HREAL_LE_MUL_RCANCEL_IMP" > "HOLLight.hollight.HREAL_LE_MUL_RCANCEL_IMP"
-  "HREAL_LE_EXISTS_DEF" > "HOLLight.hollight.HREAL_LE_EXISTS_DEF"
-  "HREAL_LE_ADD_RCANCEL" > "HOLLight.hollight.HREAL_LE_ADD_RCANCEL"
-  "HREAL_LE_ADD_LCANCEL" > "HOLLight.hollight.HREAL_LE_ADD_LCANCEL"
-  "HREAL_LE_ADD2" > "HOLLight.hollight.HREAL_LE_ADD2"
-  "HREAL_EQ_ADD_RCANCEL" > "HOLLight.hollight.HREAL_EQ_ADD_RCANCEL"
-  "HREAL_EQ_ADD_LCANCEL" > "HOLLight.hollight.HREAL_EQ_ADD_LCANCEL"
-  "HREAL_ADD_RID" > "HOLLight.hollight.HREAL_ADD_RID"
-  "HREAL_ADD_RDISTRIB" > "HOLLight.hollight.HREAL_ADD_RDISTRIB"
-  "HREAL_ADD_AC" > "HOLLight.hollight.HREAL_ADD_AC"
-  "HD_APPEND" > "List.hd_append"
-  "HAS_SIZE_def" > "HOLLight.hollight.HAS_SIZE_def"
-  "HAS_SIZE_UNIONS" > "HOLLight.hollight.HAS_SIZE_UNIONS"
-  "HAS_SIZE_UNION" > "HOLLight.hollight.HAS_SIZE_UNION"
-  "HAS_SIZE_SUC" > "HOLLight.hollight.HAS_SIZE_SUC"
-  "HAS_SIZE_PRODUCT_DEPENDENT" > "HOLLight.hollight.HAS_SIZE_PRODUCT_DEPENDENT"
-  "HAS_SIZE_PRODUCT" > "HOLLight.hollight.HAS_SIZE_PRODUCT"
-  "HAS_SIZE_POWERSET" > "HOLLight.hollight.HAS_SIZE_POWERSET"
-  "HAS_SIZE_NUMSEG_LT" > "HOLLight.hollight.HAS_SIZE_NUMSEG_LT"
-  "HAS_SIZE_NUMSEG_LE" > "HOLLight.hollight.HAS_SIZE_NUMSEG_LE"
-  "HAS_SIZE_NUMSEG_1" > "HOLLight.hollight.HAS_SIZE_NUMSEG_1"
-  "HAS_SIZE_NUMSEG" > "HOLLight.hollight.HAS_SIZE_NUMSEG"
-  "HAS_SIZE_INDEX" > "HOLLight.hollight.HAS_SIZE_INDEX"
-  "HAS_SIZE_IMAGE_INJ_EQ" > "HOLLight.hollight.HAS_SIZE_IMAGE_INJ_EQ"
-  "HAS_SIZE_IMAGE_INJ" > "HOLLight.hollight.HAS_SIZE_IMAGE_INJ"
-  "HAS_SIZE_FUNSPACE" > "HOLLight.hollight.HAS_SIZE_FUNSPACE"
-  "HAS_SIZE_FINITE_IMAGE" > "HOLLight.hollight.HAS_SIZE_FINITE_IMAGE"
-  "HAS_SIZE_DIFF" > "HOLLight.hollight.HAS_SIZE_DIFF"
-  "HAS_SIZE_CROSS" > "HOLLight.hollight.HAS_SIZE_CROSS"
-  "HAS_SIZE_CLAUSES" > "HOLLight.hollight.HAS_SIZE_CLAUSES"
-  "HAS_SIZE_CARD" > "HOLLight.hollight.HAS_SIZE_CARD"
-  "HAS_SIZE_1" > "HOLLight.hollight.HAS_SIZE_1"
-  "HAS_SIZE_0" > "HOLLight.hollight.HAS_SIZE_0"
-  "GE_C" > "HOLLight.hollight.GE_C"
-  "FUN_IN_IMAGE" > "Set.imageI"
-  "FUN_EQ_THM" > "HOL.fun_eq_iff"
-  "FUNCTION_FACTORS_RIGHT" > "HOLLight.hollight.FUNCTION_FACTORS_RIGHT"
-  "FUNCTION_FACTORS_LEFT" > "HOLLight.hollight.FUNCTION_FACTORS_LEFT"
-  "FST" > "Product_Type.fst_conv"
-  "FORALL_UNWIND_THM2" > "HOL.simp_thms_41"
-  "FORALL_UNWIND_THM1" > "HOL.simp_thms_42"
-  "FORALL_UNCURRY" > "HOLLight.hollight.FORALL_UNCURRY"
-  "FORALL_TRIPLED_THM" > "HOLLight.hollight.FORALL_TRIPLED_THM"
-  "FORALL_SUBSET_IMAGE" > "HOLLight.hollight.FORALL_SUBSET_IMAGE"
-  "FORALL_SIMP" > "HOL.simp_thms_35"
-  "FORALL_PAIR_THM" > "Product_Type.split_paired_All"
-  "FORALL_PAIRED_THM" > "HOLLight.hollight.FORALL_PAIRED_THM"
-  "FORALL_NOT_THM" > "HOL.not_ex"
-  "FORALL_IN_UNIONS" > "HOLLight.hollight.FORALL_IN_UNIONS"
-  "FORALL_IN_INSERT" > "HOLLight.hollight.FORALL_IN_INSERT"
-  "FORALL_IN_IMAGE" > "HOLLight.hollight.FORALL_IN_IMAGE"
-  "FORALL_IN_GSPEC" > "HOLLight.hollight.FORALL_IN_GSPEC"
-  "FORALL_IN_CLAUSES" > "HOLLight.hollight.FORALL_IN_CLAUSES"
-  "FORALL_FINITE_INDEX" > "HOLLight.hollight.FORALL_FINITE_INDEX"
-  "FORALL_BOOL_THM" > "Set.all_bool_eq"
-  "FORALL_AND_THM" > "HOL.all_conj_distrib"
-  "FORALL_ALL" > "HOLLightList.FORALL_ALL"
-  "FNIL_def" > "HOLLight.hollight.FNIL_def"
-  "FINREC_def" > "HOLLight.hollight.FINREC_def"
-  "FINREC_UNIQUE_LEMMA" > "HOLLight.hollight.FINREC_UNIQUE_LEMMA"
-  "FINREC_SUC_LEMMA" > "HOLLight.hollight.FINREC_SUC_LEMMA"
-  "FINREC_FUN_LEMMA" > "HOLLight.hollight.FINREC_FUN_LEMMA"
-  "FINREC_FUN" > "HOLLight.hollight.FINREC_FUN"
-  "FINREC_EXISTS_LEMMA" > "HOLLight.hollight.FINREC_EXISTS_LEMMA"
-  "FINREC_1_LEMMA" > "HOLLight.hollight.FINREC_1_LEMMA"
-  "FINITE_UNION_IMP" > "Finite_Set.finite_UnI"
-  "FINITE_UNIONS" > "HOLLight.hollight.FINITE_UNIONS"
-  "FINITE_UNION" > "Finite_Set.finite_Un"
-  "FINITE_SUPPORT_DELTA" > "HOLLight.hollight.FINITE_SUPPORT_DELTA"
-  "FINITE_SUPPORT" > "HOLLight.hollight.FINITE_SUPPORT"
-  "FINITE_SUM_IMAGE" > "HOLLight.hollight.FINITE_SUM_IMAGE"
-  "FINITE_SUBSET_IMAGE_IMP" > "HOLLight.hollight.FINITE_SUBSET_IMAGE_IMP"
-  "FINITE_SUBSET_IMAGE" > "HOLLight.hollight.FINITE_SUBSET_IMAGE"
-  "FINITE_SUBSET" > "Finite_Set.finite_subset"
-  "FINITE_SING" > "HOLLight.hollight.FINITE_SING"
-  "FINITE_RESTRICT" > "HOLLight.hollight.FINITE_RESTRICT"
-  "FINITE_RECURSION_DELETE" > "HOLLight.hollight.FINITE_RECURSION_DELETE"
-  "FINITE_RECURSION" > "HOLLight.hollight.FINITE_RECURSION"
-  "FINITE_REAL_INTERVAL" > "HOLLight.hollight.FINITE_REAL_INTERVAL"
-  "FINITE_PRODUCT_DEPENDENT" > "HOLLight.hollight.FINITE_PRODUCT_DEPENDENT"
-  "FINITE_PRODUCT" > "HOLLight.hollight.FINITE_PRODUCT"
-  "FINITE_POWERSET" > "HOLLight.hollight.FINITE_POWERSET"
-  "FINITE_NUMSEG_LT" > "HOLLight.hollight.FINITE_NUMSEG_LT"
-  "FINITE_NUMSEG_LE" > "HOLLight.hollight.FINITE_NUMSEG_LE"
-  "FINITE_NUMSEG" > "SetInterval.finite_atLeastAtMost"
-  "FINITE_INTSEG" > "HOLLight.hollight.FINITE_INTSEG"
-  "FINITE_INTER" > "Finite_Set.finite_Int"
-  "FINITE_INSERT" > "Finite_Set.finite_insert"
-  "FINITE_INDUCT_STRONG" > "Finite_Set.finite_induct"
-  "FINITE_INDUCT_DELETE" > "HOLLight.hollight.FINITE_INDUCT_DELETE"
-  "FINITE_INDEX_WORKS" > "HOLLight.hollight.FINITE_INDEX_WORKS"
-  "FINITE_INDEX_NUMSEG" > "HOLLight.hollight.FINITE_INDEX_NUMSEG"
-  "FINITE_INDEX_NUMBERS" > "HOLLight.hollight.FINITE_INDEX_NUMBERS"
-  "FINITE_INDEX_INRANGE" > "HOLLight.hollight.FINITE_INDEX_INRANGE"
-  "FINITE_INDEX_INJ" > "HOLLight.hollight.FINITE_INDEX_INJ"
-  "FINITE_IMAGE_INJ_GENERAL" > "HOLLight.hollight.FINITE_IMAGE_INJ_GENERAL"
-  "FINITE_IMAGE_INJ_EQ" > "HOLLight.hollight.FINITE_IMAGE_INJ_EQ"
-  "FINITE_IMAGE_INJ" > "HOLLight.hollight.FINITE_IMAGE_INJ"
-  "FINITE_IMAGE_IMAGE" > "HOLLight.hollight.FINITE_IMAGE_IMAGE"
-  "FINITE_IMAGE_EXPAND" > "HOLLight.hollight.FINITE_IMAGE_EXPAND"
-  "FINITE_IMAGE" > "Finite_Set.finite_imageI"
-  "FINITE_HAS_SIZE" > "HOLLight.hollight.FINITE_HAS_SIZE"
-  "FINITE_FUNSPACE" > "HOLLight.hollight.FINITE_FUNSPACE"
-  "FINITE_FINITE_UNIONS" > "HOLLight.hollight.FINITE_FINITE_UNIONS"
-  "FINITE_FINITE_PREIMAGE_GENERAL" > "HOLLight.hollight.FINITE_FINITE_PREIMAGE_GENERAL"
-  "FINITE_FINITE_PREIMAGE" > "HOLLight.hollight.FINITE_FINITE_PREIMAGE"
-  "FINITE_FINITE_IMAGE" > "HOLLight.hollight.FINITE_FINITE_IMAGE"
-  "FINITE_EMPTY" > "Finite_Set.finite.emptyI"
-  "FINITE_DIFF" > "Finite_Set.finite_Diff"
-  "FINITE_DELETE_IMP" > "HOLLight.hollight.FINITE_DELETE_IMP"
-  "FINITE_DELETE" > "HOLLight.hollight.FINITE_DELETE"
-  "FINITE_CROSS" > "HOLLight.hollight.FINITE_CROSS"
-  "FINITE_CART" > "HOLLight.hollight.FINITE_CART"
-  "FILTER_MAP" > "List.filter_map"
-  "FILTER_APPEND" > "List.filter_append"
-  "FCONS_def" > "HOLLight.hollight.FCONS_def"
-  "FCONS_UNDO" > "HOLLight.hollight.FCONS_UNDO"
-  "FACT_NZ" > "Fact.fact_nonzero_nat"
-  "FACT_MONO" > "Fact.fact_mono_nat"
-  "FACT_LT" > "Fact.fact_gt_zero_nat"
-  "FACT_LE" > "Fact.fact_ge_one_nat"
-  "EX_MEM" > "HOLLightList.EX_MEM"
-  "EX_IMP" > "HOLLightList.EX_IMP"
-  "EXTENSION" > "Set.set_eq_iff"
-  "EXP_ZERO" > "Power.power_0_left"
-  "EXP_ONE" > "Power.monoid_mult_class.power_one"
-  "EXP_MULT" > "Power.monoid_mult_class.power_mult"
-  "EXP_MONO_LT_IMP" > "HOLLight.hollight.EXP_MONO_LT_IMP"
-  "EXP_MONO_LT" > "HOLLight.hollight.EXP_MONO_LT"
-  "EXP_MONO_LE_IMP" > "HOLLight.hollight.EXP_MONO_LE_IMP"
-  "EXP_MONO_LE" > "HOLLight.hollight.EXP_MONO_LE"
-  "EXP_MONO_EQ" > "HOLLight.hollight.EXP_MONO_EQ"
-  "EXP_LT_0" > "HOLLight.hollight.EXP_LT_0"
-  "EXP_EQ_1" > "HOLLight.hollight.EXP_EQ_1"
-  "EXP_EQ_0" > "Power.power_eq_0_iff"
-  "EXP_ADD" > "Power.monoid_mult_class.power_add"
-  "EXP_2" > "Nat_Numeral.monoid_mult_class.power2_eq_square"
-  "EXP_1" > "Power.monoid_mult_class.power_one_right"
-  "EXISTS_UNIQUE_THM" > "HOLLightCompat.EXISTS_UNIQUE_THM"
-  "EXISTS_UNIQUE_REFL" > "HOL.ex1_eq_1"
-  "EXISTS_UNIQUE_ALT" > "HOLLight.hollight.EXISTS_UNIQUE_ALT"
-  "EXISTS_UNIQUE" > "HOL.Ex1_def"
-  "EXISTS_UNCURRY" > "HOLLight.hollight.EXISTS_UNCURRY"
-  "EXISTS_TRIPLED_THM" > "HOLLight.hollight.EXISTS_TRIPLED_THM"
-  "EXISTS_THM" > "Importer.EXISTS_DEF"
-  "EXISTS_SUBSET_IMAGE" > "HOLLight.hollight.EXISTS_SUBSET_IMAGE"
-  "EXISTS_SIMP" > "HOL.simp_thms_36"
-  "EXISTS_REFL" > "HOL.simp_thms_37"
-  "EXISTS_PAIR_THM" > "Product_Type.split_paired_Ex"
-  "EXISTS_PAIRED_THM" > "HOLLight.hollight.EXISTS_PAIRED_THM"
-  "EXISTS_OR_THM" > "HOL.ex_disj_distrib"
-  "EXISTS_ONE_REP" > "HOLLight.hollight.EXISTS_ONE_REP"
-  "EXISTS_NOT_THM" > "HOL.not_all"
-  "EXISTS_IN_UNIONS" > "HOLLight.hollight.EXISTS_IN_UNIONS"
-  "EXISTS_IN_INSERT" > "HOLLight.hollight.EXISTS_IN_INSERT"
-  "EXISTS_IN_IMAGE" > "HOLLight.hollight.EXISTS_IN_IMAGE"
-  "EXISTS_IN_GSPEC" > "HOLLight.hollight.EXISTS_IN_GSPEC"
-  "EXISTS_IN_CLAUSES" > "HOLLight.hollight.EXISTS_IN_CLAUSES"
-  "EXISTS_FINITE_SUBSET_IMAGE" > "HOLLight.hollight.EXISTS_FINITE_SUBSET_IMAGE"
-  "EXISTS_EX" > "HOLLightList.EXISTS_EX"
-  "EXISTS_BOOL_THM" > "Set.ex_bool_eq"
-  "EXCLUDED_MIDDLE" > "HOLLight.hollight.EXCLUDED_MIDDLE"
-  "EVEN_SUB" > "HOLLight.hollight.EVEN_SUB"
-  "EVEN_OR_ODD" > "HOLLight.hollight.EVEN_OR_ODD"
-  "EVEN_ODD_DECOMPOSITION" > "HOLLight.hollight.EVEN_ODD_DECOMPOSITION"
-  "EVEN_MULT" > "Parity.even_product_nat"
-  "EVEN_MOD" > "HOLLight.hollight.EVEN_MOD"
-  "EVEN_EXP" > "HOLLight.hollight.EVEN_EXP"
-  "EVEN_EXISTS_LEMMA" > "HOLLight.hollight.EVEN_EXISTS_LEMMA"
-  "EVEN_EXISTS" > "Parity.even_mult_two_ex"
-  "EVEN_DOUBLE" > "HOLLight.hollight.EVEN_DOUBLE"
-  "EVEN_AND_ODD" > "HOLLight.hollight.EVEN_AND_ODD"
-  "EVEN_ADD" > "Parity.even_add"
-  "EQ_UNIV" > "HOLLight.hollight.EQ_UNIV"
-  "EQ_TRANS" > "HOL.trans"
-  "EQ_SYM_EQ" > "HOL.eq_ac_1"
-  "EQ_SYM" > "HOL.eq_reflection"
-  "EQ_REFL" > "HOL.refl"
-  "EQ_MULT_RCANCEL" > "Nat.mult_cancel2"
-  "EQ_MULT_LCANCEL" > "Numeral_Simprocs.nat_mult_eq_cancel_disj"
-  "EQ_IMP_LE" > "Nat.eq_imp_le"
-  "EQ_EXT" > "HOL.eq_reflection"
-  "EQ_EXP" > "HOLLight.hollight.EQ_EXP"
-  "EQ_CLAUSES" > "HOLLight.hollight.EQ_CLAUSES"
-  "EQ_ADD_RCANCEL_0" > "HOLLight.hollight.EQ_ADD_RCANCEL_0"
-  "EQ_ADD_RCANCEL" > "Groups.cancel_semigroup_add_class.add_right_cancel"
-  "EQ_ADD_LCANCEL_0" > "HOLLight.hollight.EQ_ADD_LCANCEL_0"
-  "EQ_ADD_LCANCEL" > "Groups.cancel_semigroup_add_class.add_left_cancel"
-  "EMPTY_UNIONS" > "HOLLight.hollight.EMPTY_UNIONS"
-  "EMPTY_UNION" > "Lattices.bounded_lattice_bot_class.sup_eq_bot_iff"
-  "EMPTY_SUBSET" > "Orderings.bot_class.bot_least"
-  "EMPTY_NOT_UNIV" > "HOLLight.hollight.EMPTY_NOT_UNIV"
-  "EMPTY_GSPEC" > "HOLLight.hollight.EMPTY_GSPEC"
-  "EMPTY_DIFF" > "Set.empty_Diff"
-  "EMPTY_DELETE" > "HOLLight.hollight.EMPTY_DELETE"
-  "EL_CONS" > "List.nth_Cons'"
-  "EL_APPEND" > "List.nth_append"
-  "DIV_UNIQ" > "HOLLight.hollight.DIV_UNIQ"
-  "DIV_REFL" > "Divides.semiring_div_class.div_self"
-  "DIV_MUL_LE" > "HOLLight.hollight.DIV_MUL_LE"
-  "DIV_MULT2" > "HOLLight.hollight.DIV_MULT2"
-  "DIV_MONO_LT" > "HOLLight.hollight.DIV_MONO_LT"
-  "DIV_MONO2" > "HOLLight.hollight.DIV_MONO2"
-  "DIV_MONO" > "HOLLight.hollight.DIV_MONO"
-  "DIV_MOD" > "HOLLight.hollight.DIV_MOD"
-  "DIV_LT" > "Divides.div_less"
-  "DIV_LE_EXCLUSION" > "HOLLight.hollight.DIV_LE_EXCLUSION"
-  "DIV_LE" > "HOLLight.hollight.DIV_LE"
-  "DIV_EQ_EXCLUSION" > "HOLLight.hollight.DIV_EQ_EXCLUSION"
-  "DIV_EQ_0" > "HOLLight.hollight.DIV_EQ_0"
-  "DIV_DIV" > "HOLLight.hollight.DIV_DIV"
-  "DIV_ADD_MOD" > "HOLLight.hollight.DIV_ADD_MOD"
-  "DIVMOD_UNIQ_LEMMA" > "HOLLight.hollight.DIVMOD_UNIQ_LEMMA"
-  "DIVMOD_UNIQ" > "HOLLight.hollight.DIVMOD_UNIQ"
-  "DIVMOD_EXIST_0" > "HOLLight.hollight.DIVMOD_EXIST_0"
-  "DIVMOD_EXIST" > "HOLLight.hollight.DIVMOD_EXIST"
-  "DIVMOD_ELIM_THM" > "HOLLight.hollight.DIVMOD_ELIM_THM"
-  "DIVISION" > "HOLLight.hollight.DIVISION"
-  "DIST_TRIANGLE_LE" > "HOLLight.hollight.DIST_TRIANGLE_LE"
-  "DIST_TRIANGLES_LE" > "HOLLight.hollight.DIST_TRIANGLES_LE"
-  "DIST_SYM" > "HOLLight.hollight.DIST_SYM"
-  "DIST_RZERO" > "HOLLight.hollight.DIST_RZERO"
-  "DIST_RMUL" > "HOLLight.hollight.DIST_RMUL"
-  "DIST_REFL" > "HOLLight.hollight.DIST_REFL"
-  "DIST_RADD_0" > "HOLLight.hollight.DIST_RADD_0"
-  "DIST_RADD" > "HOLLight.hollight.DIST_RADD"
-  "DIST_LZERO" > "HOLLight.hollight.DIST_LZERO"
-  "DIST_LMUL" > "HOLLight.hollight.DIST_LMUL"
-  "DIST_LE_CASES" > "HOLLight.hollight.DIST_LE_CASES"
-  "DIST_LADD_0" > "HOLLight.hollight.DIST_LADD_0"
-  "DIST_LADD" > "HOLLight.hollight.DIST_LADD"
-  "DIST_EQ_0" > "HOLLight.hollight.DIST_EQ_0"
-  "DIST_ELIM_THM" > "HOLLight.hollight.DIST_ELIM_THM"
-  "DISJ_SYM" > "Groebner_Basis.dnf_4"
-  "DISJ_ASSOC" > "HOL.disj_ac_3"
-  "DISJ_ACI" > "HOLLight.hollight.DISJ_ACI"
-  "DISJOINT_UNION" > "HOLLight.hollight.DISJOINT_UNION"
-  "DISJOINT_SYM" > "HOLLight.hollight.DISJOINT_SYM"
-  "DISJOINT_NUMSEG" > "HOLLight.hollight.DISJOINT_NUMSEG"
-  "DISJOINT_INSERT" > "HOLLight.hollight.DISJOINT_INSERT"
-  "DISJOINT_EMPTY_REFL" > "HOLLight.hollight.DISJOINT_EMPTY_REFL"
-  "DISJOINT_EMPTY" > "HOLLight.hollight.DISJOINT_EMPTY"
-  "DISJOINT_DELETE_SYM" > "HOLLight.hollight.DISJOINT_DELETE_SYM"
-  "DIMINDEX_UNIV" > "HOLLight.hollight.DIMINDEX_UNIV"
-  "DIMINDEX_UNIQUE" > "HOLLight.hollight.DIMINDEX_UNIQUE"
-  "DIMINDEX_NONZERO" > "HOLLight.hollight.DIMINDEX_NONZERO"
-  "DIMINDEX_GE_1" > "HOLLight.hollight.DIMINDEX_GE_1"
-  "DIMINDEX_FINITE_IMAGE" > "HOLLight.hollight.DIMINDEX_FINITE_IMAGE"
-  "DIFF_UNIV" > "Set.Diff_UNIV"
-  "DIFF_INTERS" > "HOLLight.hollight.DIFF_INTERS"
-  "DIFF_INSERT" > "Set.Diff_insert2"
-  "DIFF_EQ_EMPTY" > "Set.Diff_cancel"
-  "DIFF_EMPTY" > "Set.Diff_empty"
-  "DIFF_DIFF" > "Set.Diff_idemp"
-  "DEST_REC_INJ" > "HOLLight.hollight.recspace._dest_rec_inject"
-  "DELETE_SUBSET" > "HOLLight.hollight.DELETE_SUBSET"
-  "DELETE_NON_ELEMENT" > "HOLLight.hollight.DELETE_NON_ELEMENT"
-  "DELETE_INTER" > "HOLLight.hollight.DELETE_INTER"
-  "DELETE_INSERT" > "HOLLight.hollight.DELETE_INSERT"
-  "DELETE_DELETE" > "HOLLight.hollight.DELETE_DELETE"
-  "DELETE_COMM" > "HOLLight.hollight.DELETE_COMM"
-  "DEF_~" > "Groebner_Basis.bool_simps_19"
-  "DEF_vector" > "HOLLight.hollight.DEF_vector"
-  "DEF_treal_of_num" > "HOLLight.hollight.DEF_treal_of_num"
-  "DEF_treal_neg" > "HOLLight.hollight.DEF_treal_neg"
-  "DEF_treal_mul" > "HOLLight.hollight.DEF_treal_mul"
-  "DEF_treal_le" > "HOLLight.hollight.DEF_treal_le"
-  "DEF_treal_inv" > "HOLLight.hollight.DEF_treal_inv"
-  "DEF_treal_eq" > "HOLLight.hollight.DEF_treal_eq"
-  "DEF_treal_add" > "HOLLight.hollight.DEF_treal_add"
-  "DEF_tailadmissible" > "HOLLight.hollight.DEF_tailadmissible"
-  "DEF_support" > "HOLLight.hollight.DEF_support"
-  "DEF_superadmissible" > "HOLLight.hollight.DEF_superadmissible"
-  "DEF_sum" > "HOLLight.hollight.DEF_sum"
-  "DEF_sndcart" > "HOLLight.hollight.DEF_sndcart"
-  "DEF_set_of_list" > "HOLLightList.DEF_set_of_list"
-  "DEF_rem" > "HOLLight.hollight.DEF_rem"
-  "DEF_real_sub" > "HOLLight.hollight.DEF_real_sub"
-  "DEF_real_sgn" > "HOLLight.hollight.DEF_real_sgn"
-  "DEF_real_pow" > "HOLLight.hollight.DEF_real_pow"
-  "DEF_real_of_num" > "HOLLight.hollight.DEF_real_of_num"
-  "DEF_real_neg" > "HOLLight.hollight.DEF_real_neg"
-  "DEF_real_mul" > "HOLLight.hollight.DEF_real_mul"
-  "DEF_real_mod" > "HOLLight.hollight.DEF_real_mod"
-  "DEF_real_min" > "HOLLight.hollight.DEF_real_min"
-  "DEF_real_max" > "HOLLight.hollight.DEF_real_max"
-  "DEF_real_lt" > "HOLLight.hollight.DEF_real_lt"
-  "DEF_real_le" > "HOLLight.hollight.DEF_real_le"
-  "DEF_real_inv" > "HOLLight.hollight.DEF_real_inv"
-  "DEF_real_gt" > "HOLLight.hollight.DEF_real_gt"
-  "DEF_real_ge" > "HOLLight.hollight.DEF_real_ge"
-  "DEF_real_div" > "HOLLight.hollight.DEF_real_div"
-  "DEF_real_add" > "HOLLight.hollight.DEF_real_add"
-  "DEF_real_abs" > "HOLLight.hollight.DEF_real_abs"
-  "DEF_pastecart" > "HOLLight.hollight.DEF_pastecart"
-  "DEF_pairwise" > "HOLLight.hollight.DEF_pairwise"
-  "DEF_o" > "Fun.comp_def"
-  "DEF_num_of_int" > "HOLLight.hollight.DEF_num_of_int"
-  "DEF_num_mod" > "HOLLight.hollight.DEF_num_mod"
-  "DEF_num_gcd" > "HOLLight.hollight.DEF_num_gcd"
-  "DEF_num_divides" > "HOLLight.hollight.DEF_num_divides"
-  "DEF_num_coprime" > "HOLLight.hollight.DEF_num_coprime"
-  "DEF_nsum" > "HOLLight.hollight.DEF_nsum"
-  "DEF_neutral" > "HOLLight.hollight.DEF_neutral"
-  "DEF_nadd_rinv" > "HOLLight.hollight.DEF_nadd_rinv"
-  "DEF_nadd_of_num" > "HOLLight.hollight.DEF_nadd_of_num"
-  "DEF_nadd_mul" > "HOLLight.hollight.DEF_nadd_mul"
-  "DEF_nadd_le" > "HOLLight.hollight.DEF_nadd_le"
-  "DEF_nadd_inv" > "HOLLight.hollight.DEF_nadd_inv"
-  "DEF_nadd_eq" > "HOLLight.hollight.DEF_nadd_eq"
-  "DEF_nadd_add" > "HOLLight.hollight.DEF_nadd_add"
-  "DEF_monoidal" > "HOLLight.hollight.DEF_monoidal"
-  "DEF_minimal" > "HOLLight.hollight.DEF_minimal"
-  "DEF_lambda" > "HOLLight.hollight.DEF_lambda"
-  "DEF_iterate" > "HOLLight.hollight.DEF_iterate"
-  "DEF_is_nadd" > "HOLLight.hollight.DEF_is_nadd"
-  "DEF_integer" > "HOLLight.hollight.DEF_integer"
-  "DEF_int_sub" > "HOLLight.hollight.DEF_int_sub"
-  "DEF_int_sgn" > "HOLLight.hollight.DEF_int_sgn"
-  "DEF_int_pow" > "HOLLight.hollight.DEF_int_pow"
-  "DEF_int_of_num" > "HOLLight.hollight.DEF_int_of_num"
-  "DEF_int_neg" > "HOLLight.hollight.DEF_int_neg"
-  "DEF_int_mul" > "HOLLight.hollight.DEF_int_mul"
-  "DEF_int_mod" > "HOLLight.hollight.DEF_int_mod"
-  "DEF_int_min" > "HOLLight.hollight.DEF_int_min"
-  "DEF_int_max" > "HOLLight.hollight.DEF_int_max"
-  "DEF_int_lt" > "HOLLight.hollight.DEF_int_lt"
-  "DEF_int_le" > "HOLLight.hollight.DEF_int_le"
-  "DEF_int_gt" > "HOLLight.hollight.DEF_int_gt"
-  "DEF_int_ge" > "HOLLight.hollight.DEF_int_ge"
-  "DEF_int_gcd" > "HOLLight.hollight.DEF_int_gcd"
-  "DEF_int_divides" > "HOLLight.hollight.DEF_int_divides"
-  "DEF_int_coprime" > "HOLLight.hollight.DEF_int_coprime"
-  "DEF_int_add" > "HOLLight.hollight.DEF_int_add"
-  "DEF_int_abs" > "HOLLight.hollight.DEF_int_abs"
-  "DEF_hreal_of_num" > "HOLLight.hollight.DEF_hreal_of_num"
-  "DEF_hreal_mul" > "HOLLight.hollight.DEF_hreal_mul"
-  "DEF_hreal_le" > "HOLLight.hollight.DEF_hreal_le"
-  "DEF_hreal_inv" > "HOLLight.hollight.DEF_hreal_inv"
-  "DEF_hreal_add" > "HOLLight.hollight.DEF_hreal_add"
-  "DEF_fstcart" > "HOLLight.hollight.DEF_fstcart"
-  "DEF_div" > "HOLLight.hollight.DEF_div"
-  "DEF_dist" > "HOLLight.hollight.DEF_dist"
-  "DEF_dimindex" > "HOLLight.hollight.DEF_dimindex"
-  "DEF_admissible" > "HOLLight.hollight.DEF_admissible"
-  "DEF__star_" > "HOLLightCompat.DEF__star_"
-  "DEF__slash__backslash_" > "HOLLightCompat.DEF__slash__backslash_"
-  "DEF__questionmark__exclamationmark_" > "HOLLightCompat.EXISTS_UNIQUE_THM"
-  "DEF__questionmark_" > "HOL.Ex_def"
-  "DEF__lessthan__equal__c" > "HOLLight.hollight.DEF__lessthan__equal__c"
-  "DEF__lessthan__equal_" > "HOLLightCompat.DEF__lessthan__equal_"
-  "DEF__lessthan__c" > "HOLLight.hollight.DEF__lessthan__c"
-  "DEF__lessthan_" > "HOLLightCompat.DEF__lessthan_"
-  "DEF__greaterthan__equal__c" > "HOLLight.hollight.DEF__greaterthan__equal__c"
-  "DEF__greaterthan__equal_" > "HOLLightCompat.DEF__greaterthan__equal_"
-  "DEF__greaterthan__c" > "HOLLight.hollight.DEF__greaterthan__c"
-  "DEF__greaterthan_" > "HOLLightCompat.DEF__greaterthan_"
-  "DEF__exclamationmark_" > "HOL.All_def"
-  "DEF__equal__equal__greaterthan_" > "HOLLightCompat.DEF__equal__equal__greaterthan_"
-  "DEF__equal__equal_" > "HOLLight.hollight.DEF__equal__equal_"
-  "DEF__equal__c" > "HOLLight.hollight.DEF__equal__c"
-  "DEF__dot__dot_" > "HOLLightCompat.dotdot_def"
-  "DEF__backslash__slash_" > "HOL.or_def"
-  "DEF__UNGUARDED_PATTERN" > "HOLLight.hollight.DEF__UNGUARDED_PATTERN"
-  "DEF__SEQPATTERN" > "HOLLight.hollight.DEF__SEQPATTERN"
-  "DEF__MATCH" > "HOLLight.hollight.DEF__MATCH"
-  "DEF__GUARDED_PATTERN" > "HOLLight.hollight.DEF__GUARDED_PATTERN"
-  "DEF__FUNCTION" > "HOLLight.hollight.DEF__FUNCTION"
-  "DEF__FALSITY_" > "HOLLight.hollight.DEF__FALSITY_"
-  "DEF__11937" > "HOLLight.hollight.DEF__11937"
-  "DEF_ZRECSPACE" > "HOLLight.hollight.DEF_ZRECSPACE"
-  "DEF_ZCONSTR" > "HOLLight.hollight.DEF_ZCONSTR"
-  "DEF_ZBOT" > "HOLLight.hollight.DEF_ZBOT"
-  "DEF_WF" > "HOLLightCompat.DEF_WF"
-  "DEF_UNIV" > "HOLLightCompat.DEF_UNIV"
-  "DEF_UNIONS" > "HOLLightCompat.DEF_UNIONS"
-  "DEF_UNION" > "HOLLightCompat.DEF_UNION"
-  "DEF_UNCURRY" > "HOLLight.hollight.DEF_UNCURRY"
-  "DEF_T" > "HOL.True_def"
-  "DEF_SURJ" > "HOLLight.hollight.DEF_SURJ"
-  "DEF_SUBSET" > "HOLLightCompat.DEF_SUBSET"
-  "DEF_SND" > "HOLLightCompat.DEF_SND"
-  "DEF_SING" > "HOLLight.hollight.DEF_SING"
-  "DEF_SETSPEC" > "HOLLightCompat.SETSPEC_def"
-  "DEF_REVERSE" > "HOLLightList.DEF_REVERSE"
-  "DEF_REST" > "HOLLight.hollight.DEF_REST"
-  "DEF_REPLICATE" > "HOLLightList.DEF_REPLICATE"
-  "DEF_PSUBSET" > "HOLLightCompat.DEF_PSUBSET"
-  "DEF_PRE" > "HOLLightCompat.DEF_PRE"
-  "DEF_PASSOC" > "HOLLight.hollight.DEF_PASSOC"
-  "DEF_PAIRWISE" > "HOLLight.hollight.DEF_PAIRWISE"
-  "DEF_ONTO" > "Fun.surj_def"
-  "DEF_ONE_ONE" > "HOLLightCompat.DEF_ONE_ONE"
-  "DEF_ODD" > "HOLLightCompat.DEF_ODD"
-  "DEF_NUM_REP" > "HOLLight.hollight.DEF_NUM_REP"
-  "DEF_NUMSUM" > "HOLLight.hollight.DEF_NUMSUM"
-  "DEF_NUMSND" > "HOLLight.hollight.DEF_NUMSND"
-  "DEF_NUMRIGHT" > "HOLLight.hollight.DEF_NUMRIGHT"
-  "DEF_NUMPAIR" > "HOLLight.hollight.DEF_NUMPAIR"
-  "DEF_NUMLEFT" > "HOLLight.hollight.DEF_NUMLEFT"
-  "DEF_NUMFST" > "HOLLight.hollight.DEF_NUMFST"
-  "DEF_NUMERAL" > "HOLLightCompat.NUMERAL_def"
-  "DEF_NULL" > "HOLLightList.DEF_NULL"
-  "DEF_MOD" > "HOLLightCompat.DEF_MOD"
-  "DEF_MIN" > "Orderings.ord_class.min_def"
-  "DEF_MEM" > "HOLLightList.DEF_MEM"
-  "DEF_MEASURE" > "HOLLightCompat.MEASURE_def"
-  "DEF_MAX" > "Orderings.ord_class.max_def"
-  "DEF_MAP" > "HOLLightList.DEF_MAP"
-  "DEF_LET_END" > "HOLLight.hollight.DEF_LET_END"
-  "DEF_LET" > "HOLLightCompat.LET_def"
-  "DEF_LENGTH" > "HOLLightList.DEF_LENGTH"
-  "DEF_LAST" > "HOLLightList.DEF_LAST"
-  "DEF_ITSET" > "HOLLight.hollight.DEF_ITSET"
-  "DEF_ITLIST" > "HOLLightList.DEF_ITLIST"
-  "DEF_ISO" > "HOLLight.hollight.DEF_ISO"
-  "DEF_INTERS" > "HOLLightCompat.DEF_INTERS"
-  "DEF_INTER" > "HOLLightCompat.DEF_INTER"
-  "DEF_INSERT" > "HOLLightCompat.DEF_INSERT"
-  "DEF_INJP" > "HOLLight.hollight.DEF_INJP"
-  "DEF_INJN" > "HOLLight.hollight.DEF_INJN"
-  "DEF_INJF" > "HOLLight.hollight.DEF_INJF"
-  "DEF_INJA" > "HOLLight.hollight.DEF_INJA"
-  "DEF_INJ" > "HOLLight.hollight.DEF_INJ"
-  "DEF_INFINITE" > "HOLLightCompat.DEF_INFINITE"
-  "DEF_IND_SUC" > "HOLLight.hollight.DEF_IND_SUC"
-  "DEF_IND_0" > "HOLLight.hollight.DEF_IND_0"
-  "DEF_IN" > "Set.mem_def"
-  "DEF_IMAGE" > "HOLLightCompat.DEF_IMAGE"
-  "DEF_I" > "Fun.id_apply"
-  "DEF_HAS_SIZE" > "HOLLight.hollight.DEF_HAS_SIZE"
-  "DEF_GSPEC" > "Set.Collect_def"
-  "DEF_GEQ" > "HOLLightCompat.DEF_GEQ"
-  "DEF_GABS" > "HOLLightCompat.DEF_GABS"
-  "DEF_FST" > "HOLLightCompat.DEF_FST"
-  "DEF_FNIL" > "HOLLight.hollight.DEF_FNIL"
-  "DEF_FINREC" > "HOLLight.hollight.DEF_FINREC"
-  "DEF_FINITE" > "HOLLightCompat.DEF_FINITE"
-  "DEF_FILTER" > "HOLLightList.DEF_FILTER"
-  "DEF_FCONS" > "HOLLight.hollight.DEF_FCONS"
-  "DEF_FACT" > "HOLLightCompat.DEF_FACT"
-  "DEF_F" > "HOL.False_def"
-  "DEF_EXP" > "HOLLightCompat.DEF_EXP"
-  "DEF_EX" > "HOLLightList.DEF_EX"
-  "DEF_EVEN" > "HOLLightCompat.DEF_EVEN"
-  "DEF_EMPTY" > "HOLLightCompat.DEF_EMPTY"
-  "DEF_EL" > "HOLLightList.DEF_EL"
-  "DEF_DIV" > "HOLLightCompat.DEF_DIV"
-  "DEF_DISJOINT" > "HOLLightCompat.DEF_DISJOINT"
-  "DEF_DIFF" > "HOLLightCompat.DEF_DIFF"
-  "DEF_DELETE" > "HOLLightCompat.DEF_DELETE"
-  "DEF_DECIMAL" > "HOLLight.hollight.DEF_DECIMAL"
-  "DEF_CURRY" > "Product_Type.curry_conv"
-  "DEF_CROSS" > "HOLLight.hollight.DEF_CROSS"
-  "DEF_COUNTABLE" > "HOLLight.hollight.DEF_COUNTABLE"
-  "DEF_CONSTR" > "HOLLight.hollight.DEF_CONSTR"
-  "DEF_COND" > "HOLLightCompat.COND_DEF"
-  "DEF_CHOICE" > "HOLLightCompat.DEF_CHOICE"
-  "DEF_CASEWISE" > "HOLLight.hollight.DEF_CASEWISE"
-  "DEF_CARD" > "HOLLight.hollight.DEF_CARD"
-  "DEF_BUTLAST" > "HOLLightList.DEF_BUTLAST"
-  "DEF_BOTTOM" > "HOLLight.hollight.DEF_BOTTOM"
-  "DEF_BIT1" > "HOLLightCompat.BIT1_DEF"
-  "DEF_BIT0" > "HOLLightCompat.BIT0_DEF"
-  "DEF_BIJ" > "HOLLight.hollight.DEF_BIJ"
-  "DEF_ASCII" > "HOLLight.hollight.DEF_ASCII"
-  "DEF_APPEND" > "HOLLightList.DEF_APPEND"
-  "DEF_ALL2" > "HOLLightList.DEF_ALL2"
-  "DEF_ALL" > "HOLLightList.DEF_ALL"
-  "DEF_-" > "HOLLightCompat.DEF_MINUS"
-  "DEF_+" > "HOLLightCompat.DEF_PLUS"
-  "DEF_$" > "HOLLight.hollight.DEF_$"
-  "DECOMPOSITION" > "HOLLight.hollight.DECOMPOSITION"
-  "DECIMAL_def" > "HOLLight.hollight.DECIMAL_def"
-  "CROSS_def" > "HOLLight.hollight.CROSS_def"
-  "CROSS_EQ_EMPTY" > "HOLLight.hollight.CROSS_EQ_EMPTY"
-  "COUNTABLE_def" > "HOLLight.hollight.COUNTABLE_def"
-  "CONS_HD_TL" > "List.hd_Cons_tl"
-  "CONS_11" > "List.list.inject"
-  "CONSTR_def" > "HOLLight.hollight.CONSTR_def"
-  "CONSTR_REC" > "HOLLight.hollight.CONSTR_REC"
-  "CONSTR_INJ" > "HOLLight.hollight.CONSTR_INJ"
-  "CONSTR_IND" > "HOLLight.hollight.CONSTR_IND"
-  "CONSTR_BOT" > "HOLLight.hollight.CONSTR_BOT"
-  "CONJ_SYM" > "Groebner_Basis.dnf_3"
-  "CONJ_ASSOC" > "HOL.conj_ac_3"
-  "CONJ_ACI" > "HOLLight.hollight.CONJ_ACI"
-  "COND_RATOR" > "HOLLight.hollight.COND_RATOR"
-  "COND_RAND" > "HOL.if_distrib"
-  "COND_ID" > "HOL.if_cancel"
-  "COND_EXPAND" > "HOLLight.hollight.COND_EXPAND"
-  "COND_EQ_CLAUSE" > "HOLLight.hollight.COND_EQ_CLAUSE"
-  "COND_ELIM_THM" > "HOL.if_splits_1"
-  "COND_CLAUSES" > "HOLLight.hollight.COND_CLAUSES"
-  "COND_ABS" > "HOLLight.hollight.COND_ABS"
-  "COMPONENT" > "Set.insertI1"
-  "CHOOSE_SUBSET_STRONG" > "HOLLight.hollight.CHOOSE_SUBSET_STRONG"
-  "CHOOSE_SUBSET" > "HOLLight.hollight.CHOOSE_SUBSET"
-  "CHOICE_DEF" > "HOLLight.hollight.CHOICE_DEF"
-  "CASEWISE_def" > "HOLLight.hollight.CASEWISE_def"
-  "CART_EQ_FULL" > "HOLLight.hollight.CART_EQ_FULL"
-  "CART_EQ" > "HOLLight.hollight.CART_EQ"
-  "CARD_def" > "HOLLight.hollight.CARD_def"
-  "CARD_UNION_OVERLAP_EQ" > "HOLLight.hollight.CARD_UNION_OVERLAP_EQ"
-  "CARD_UNION_OVERLAP" > "HOLLight.hollight.CARD_UNION_OVERLAP"
-  "CARD_UNION_LE" > "HOLLight.hollight.CARD_UNION_LE"
-  "CARD_UNION_GEN" > "HOLLight.hollight.CARD_UNION_GEN"
-  "CARD_UNION_EQ" > "HOLLight.hollight.CARD_UNION_EQ"
-  "CARD_UNIONS_LE" > "HOLLight.hollight.CARD_UNIONS_LE"
-  "CARD_UNIONS" > "HOLLight.hollight.CARD_UNIONS"
-  "CARD_UNION" > "HOLLight.hollight.CARD_UNION"
-  "CARD_SUBSET_LE" > "HOLLight.hollight.CARD_SUBSET_LE"
-  "CARD_SUBSET_IMAGE" > "HOLLight.hollight.CARD_SUBSET_IMAGE"
-  "CARD_SUBSET_EQ" > "HOLLight.hollight.CARD_SUBSET_EQ"
-  "CARD_SUBSET" > "HOLLight.hollight.CARD_SUBSET"
-  "CARD_PSUBSET" > "HOLLight.hollight.CARD_PSUBSET"
-  "CARD_PRODUCT" > "HOLLight.hollight.CARD_PRODUCT"
-  "CARD_POWERSET" > "HOLLight.hollight.CARD_POWERSET"
-  "CARD_NUMSEG_LT" > "HOLLight.hollight.CARD_NUMSEG_LT"
-  "CARD_NUMSEG_LEMMA" > "HOLLight.hollight.CARD_NUMSEG_LEMMA"
-  "CARD_NUMSEG_LE" > "HOLLight.hollight.CARD_NUMSEG_LE"
-  "CARD_NUMSEG_1" > "HOLLight.hollight.CARD_NUMSEG_1"
-  "CARD_NUMSEG" > "HOLLight.hollight.CARD_NUMSEG"
-  "CARD_LE_INJ" > "HOLLight.hollight.CARD_LE_INJ"
-  "CARD_IMAGE_LE" > "HOLLight.hollight.CARD_IMAGE_LE"
-  "CARD_IMAGE_INJ_EQ" > "HOLLight.hollight.CARD_IMAGE_INJ_EQ"
-  "CARD_IMAGE_INJ" > "HOLLight.hollight.CARD_IMAGE_INJ"
-  "CARD_FUNSPACE" > "HOLLight.hollight.CARD_FUNSPACE"
-  "CARD_FINITE_IMAGE" > "HOLLight.hollight.CARD_FINITE_IMAGE"
-  "CARD_EQ_SUM" > "HOLLight.hollight.CARD_EQ_SUM"
-  "CARD_EQ_NSUM" > "HOLLight.hollight.CARD_EQ_NSUM"
-  "CARD_EQ_BIJECTIONS" > "HOLLight.hollight.CARD_EQ_BIJECTIONS"
-  "CARD_EQ_BIJECTION" > "HOLLight.hollight.CARD_EQ_BIJECTION"
-  "CARD_EQ_0" > "HOLLight.hollight.CARD_EQ_0"
-  "CARD_DIFF" > "HOLLight.hollight.CARD_DIFF"
-  "CARD_DELETE" > "HOLLight.hollight.CARD_DELETE"
-  "CARD_CROSS" > "HOLLight.hollight.CARD_CROSS"
-  "CARD_CLAUSES" > "HOLLight.hollight.CARD_CLAUSES"
-  "BOUNDS_NOTZERO" > "HOLLight.hollight.BOUNDS_NOTZERO"
-  "BOUNDS_LINEAR_0" > "HOLLight.hollight.BOUNDS_LINEAR_0"
-  "BOUNDS_LINEAR" > "HOLLight.hollight.BOUNDS_LINEAR"
-  "BOUNDS_IGNORE" > "HOLLight.hollight.BOUNDS_IGNORE"
-  "BOUNDS_DIVIDED" > "HOLLight.hollight.BOUNDS_DIVIDED"
-  "BOTTOM_def" > "HOLLight.hollight.BOTTOM_def"
-  "BOOL_CASES_AX" > "HOL.True_or_False"
-  "BIT1_THM" > "HOLLight.hollight.BIT1_THM"
-  "BIT1" > "HOLLight.hollight.BIT1"
-  "BIT0_THM" > "Int.semiring_mult_2"
-  "BIT0" > "Int.semiring_mult_2"
-  "BIJ_def" > "HOLLight.hollight.BIJ_def"
-  "BIJECTIVE_ON_LEFT_RIGHT_INVERSE" > "HOLLight.hollight.BIJECTIVE_ON_LEFT_RIGHT_INVERSE"
-  "BIJECTIVE_LEFT_RIGHT_INVERSE" > "HOLLight.hollight.BIJECTIVE_LEFT_RIGHT_INVERSE"
-  "BIJECTIONS_HAS_SIZE_EQ" > "HOLLight.hollight.BIJECTIONS_HAS_SIZE_EQ"
-  "BIJECTIONS_HAS_SIZE" > "HOLLight.hollight.BIJECTIONS_HAS_SIZE"
-  "BIJECTIONS_CARD_EQ" > "HOLLight.hollight.BIJECTIONS_CARD_EQ"
-  "BETA_THM" > "HOL.eta_contract_eq"
-  "ASCII_def" > "HOLLight.hollight.ASCII_def"
-  "ARITH_ZERO" > "HOLLight.hollight.ARITH_ZERO"
-  "ARITH_SUC" > "HOLLight.hollight.ARITH_SUC"
-  "ARITH_SUB" > "HOLLight.hollight.ARITH_SUB"
-  "ARITH_PRE" > "HOLLight.hollight.ARITH_PRE"
-  "ARITH_ODD" > "HOLLight.hollight.ARITH_ODD"
-  "ARITH_MULT" > "HOLLight.hollight.ARITH_MULT"
-  "ARITH_LT" > "HOLLight.hollight.ARITH_LT"
-  "ARITH_LE" > "HOLLight.hollight.ARITH_LE"
-  "ARITH_EXP" > "HOLLight.hollight.ARITH_EXP"
-  "ARITH_EVEN" > "HOLLight.hollight.ARITH_EVEN"
-  "ARITH_EQ" > "HOLLight.hollight.ARITH_EQ"
-  "ARITH_ADD" > "HOLLight.hollight.ARITH_ADD"
-  "APPEND_NIL" > "List.append_Nil2"
-  "APPEND_EQ_NIL" > "List.append_is_Nil_conv"
-  "APPEND_BUTLAST_LAST" > "List.append_butlast_last_id"
-  "APPEND_ASSOC" > "List.append_assoc"
-  "AND_FORALL_THM" > "HOL.all_conj_distrib"
-  "AND_CLAUSES" > "HOLLight.hollight.AND_CLAUSES"
-  "AND_ALL2" > "HOLLightList.AND_ALL2"
-  "AND_ALL" > "HOLLightList.AND_ALL"
-  "ALL_T" > "HOLLightList.ALL_T"
-  "ALL_MP" > "HOLLightList.ALL_MP"
-  "ALL_MEM" > "HOLLightList.ALL_MEM"
-  "ALL_IMP" > "HOLLightList.ALL_IMP"
-  "ALL_EL" > "List.list_all_length"
-  "ALL_APPEND" > "List.list_all_append"
-  "ALL2_MAP2" > "HOLLightList.ALL2_MAP2"
-  "ALL2_MAP" > "HOLLightList.ALL2_MAP"
-  "ALL2_AND_RIGHT" > "HOLLightList.ALL2_AND_RIGHT"
-  "ALL2_ALL" > "HOLLightList.ALL2_ALL"
-  "ALL2" > "HOLLightList.ALL2"
-  "ADMISSIBLE_UNGUARDED_PATTERN" > "HOLLight.hollight.ADMISSIBLE_UNGUARDED_PATTERN"
-  "ADMISSIBLE_SUM" > "HOLLight.hollight.ADMISSIBLE_SUM"
-  "ADMISSIBLE_SEQPATTERN" > "HOLLight.hollight.ADMISSIBLE_SEQPATTERN"
-  "ADMISSIBLE_RAND" > "HOLLight.hollight.ADMISSIBLE_RAND"
-  "ADMISSIBLE_NSUM" > "HOLLight.hollight.ADMISSIBLE_NSUM"
-  "ADMISSIBLE_NEST" > "HOLLight.hollight.ADMISSIBLE_NEST"
-  "ADMISSIBLE_MATCH_SEQPATTERN" > "HOLLight.hollight.ADMISSIBLE_MATCH_SEQPATTERN"
-  "ADMISSIBLE_MATCH" > "HOLLight.hollight.ADMISSIBLE_MATCH"
-  "ADMISSIBLE_MAP" > "HOLLight.hollight.ADMISSIBLE_MAP"
-  "ADMISSIBLE_LAMBDA" > "HOLLight.hollight.ADMISSIBLE_LAMBDA"
-  "ADMISSIBLE_IMP_SUPERADMISSIBLE" > "HOLLight.hollight.ADMISSIBLE_IMP_SUPERADMISSIBLE"
-  "ADMISSIBLE_GUARDED_PATTERN" > "HOLLight.hollight.ADMISSIBLE_GUARDED_PATTERN"
-  "ADMISSIBLE_CONST" > "HOLLight.hollight.ADMISSIBLE_CONST"
-  "ADMISSIBLE_COND" > "HOLLight.hollight.ADMISSIBLE_COND"
-  "ADMISSIBLE_COMB" > "HOLLight.hollight.ADMISSIBLE_COMB"
-  "ADMISSIBLE_BASE" > "HOLLight.hollight.ADMISSIBLE_BASE"
-  "ADD_SYM" > "Fields.linordered_field_class.sign_simps_43"
-  "ADD_SUC" > "Nat.add_Suc_right"
-  "ADD_SUBR2" > "Nat.diff_add_0"
-  "ADD_SUBR" > "HOLLight.hollight.ADD_SUBR"
-  "ADD_SUB2" > "Nat.diff_add_inverse"
-  "ADD_SUB" > "Nat.diff_add_inverse2"
-  "ADD_EQ_0" > "Nat.add_is_0"
-  "ADD_CLAUSES" > "HOLLight.hollight.ADD_CLAUSES"
-  "ADD_ASSOC" > "Fields.linordered_field_class.sign_simps_44"
-  "ADD_AC" > "HOLLight.hollight.ADD_AC"
-  "ADD_0" > "Divides.arithmetic_simps_39"
-  "ADD1" > "Nat_Numeral.Suc_eq_plus1"
-  "ABS_SIMP" > "HOL.refl"
-  "ABSORPTION" > "HOLLight.hollight.ABSORPTION"
-  ">_c_def" > "HOLLight.hollight.>_c_def"
-  ">=_c_def" > "HOLLight.hollight.>=_c_def"
-  "=_c_def" > "HOLLight.hollight.=_c_def"
-  "<_c_def" > "HOLLight.hollight.<_c_def"
-  "<=_c_def" > "HOLLight.hollight.<=_c_def"
-  "$_def" > "HOLLight.hollight.$_def"
-
-ignore_thms
-  "WF_REC_CASES"
-  "TYDEF_sum"
-  "TYDEF_prod"
-  "TYDEF_option"
-  "TYDEF_num"
-  "TYDEF_list"
-  "TYDEF_1"
-  "SNDCART_PASTECART"
-  "SET_OF_LIST_OF_SET"
-  "REP_ABS_PAIR"
-  "RECURSION_CASEWISE_PAIRWISE"
-  "RECURSION_CASEWISE"
-  "PASTECART_FST_SND"
-  "PASTECART_EQ"
-  "MEM_LIST_OF_SET"
-  "MEM_ASSOC"
-  "LIST_OF_SET_PROPERTIES"
-  "LENGTH_LIST_OF_SET"
-  "HAS_SIZE_SET_OF_LIST"
-  "FSTCART_PASTECART"
-  "FORALL_PASTECART"
-  "FINITE_SET_OF_LIST"
-  "EX_MAP"
-  "EXISTS_PASTECART"
-  "EL_TL"
-  "DIMINDEX_HAS_SIZE_FINITE_SUM"
-  "DIMINDEX_FINITE_SUM"
-  "DEF_one"
-  "DEF_mk_pair"
-  "DEF_list_of_set"
-  "DEF__0"
-  "DEF_ZIP"
-  "DEF_TL"
-  "DEF_SUC"
-  "DEF_SOME"
-  "DEF_OUTR"
-  "DEF_OUTL"
-  "DEF_NONE"
-  "DEF_NIL"
-  "DEF_MAP2"
-  "DEF_ITLIST2"
-  "DEF_INR"
-  "DEF_INL"
-  "DEF_HD"
-  "DEF_CONS"
-  "DEF_ASSOC"
-  "DEF_,"
-  "CASEWISE_WORKS"
-  "CASEWISE_CASES"
-  "CASEWISE"
-  "CARD_SET_OF_LIST_LE"
-  "ALL_MAP"
-
-end
--- a/src/HOL/Import/HOL_Light/HOLLightInt.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,208 +0,0 @@
-(*  Title:      HOL/Import/HOL_Light/HOLLightInt.thy
-    Author:     Cezary Kaliszyk
-*)
-
-header {* Compatibility theorems for HOL Light integers *}
-
-theory HOLLightInt imports Main Real GCD begin
-
-fun int_coprime where "int_coprime ((a :: int), (b :: int)) = coprime a b"
-
-lemma DEF_int_coprime:
-  "int_coprime = (\<lambda>u. \<exists>x y. ((fst u) * x) + ((snd u) * y) = int 1)"
-  apply (auto simp add: fun_eq_iff)
-  apply (metis bezout_int mult_commute)
-  by (metis coprime_divisors_nat dvd_triv_left gcd_1_int gcd_add2_int)
-
-lemma INT_FORALL_POS:
-  "(\<forall>n. P (int n)) = (\<forall>i\<ge>(int 0). P i)"
-  by (auto, drule_tac x="nat i" in spec) simp
-
-lemma INT_LT_DISCRETE:
-  "(x < y) = (x + int 1 \<le> y)"
-  by auto
-
-lemma INT_ABS_MUL_1:
-  "(abs (x * y) = int 1) = (abs x = int 1 \<and> abs y = int 1)"
-  by simp (metis dvd_mult_right zdvd1_eq abs_zmult_eq_1 abs_mult mult_1_right)
-
-lemma dest_int_rep:
-  "\<exists>(n :: nat). real (i :: int) = real n \<or> real i = - real n"
-  by (metis (full_types) of_int_of_nat real_eq_of_int real_of_nat_def)
-
-lemma DEF_int_add:
-  "op + = (\<lambda>u ua. floor (real u + real ua))"
-  by simp
-
-lemma DEF_int_sub:
-  "op - = (\<lambda>u ua. floor (real u - real ua))"
-  by simp
-
-lemma DEF_int_mul:
-  "op * = (\<lambda>u ua. floor (real u * real ua))"
-  by (metis floor_real_of_int real_of_int_mult)
-
-lemma DEF_int_abs:
-  "abs = (\<lambda>u. floor (abs (real u)))"
-  by (metis floor_real_of_int real_of_int_abs)
-
-lemma DEF_int_sgn:
-  "sgn = (\<lambda>u. floor (sgn (real u)))"
-  by (simp add: sgn_if fun_eq_iff)
-
-lemma int_sgn_th:
-  "real (sgn (x :: int)) = sgn (real x)"
-  by (simp add: sgn_if)
-
-lemma DEF_int_max:
-  "max = (\<lambda>u ua. floor (max (real u) (real ua)))"
-  by (metis floor_real_of_int real_of_int_le_iff sup_absorb1 sup_commute sup_max linorder_linear)
-
-lemma int_max_th:
-  "real (max (x :: int) y) = max (real x) (real y)"
-  by (metis min_max.le_iff_sup min_max.sup_absorb1 real_of_int_le_iff linorder_linear)
-
-lemma DEF_int_min:
-  "min = (\<lambda>u ua. floor (min (real u) (real ua)))"
-  by (metis floor_real_of_int inf_absorb1 inf_absorb2 inf_int_def inf_real_def real_of_int_le_iff linorder_linear)
-
-lemma int_min_th:
-  "real (min (x :: int) y) = min (real x) (real y)"
-  by (metis inf_absorb1 inf_absorb2 inf_int_def inf_real_def real_of_int_le_iff linorder_linear)
-
-lemma INT_IMAGE:
-  "(\<exists>n. x = int n) \<or> (\<exists>n. x = - int n)"
-  by (metis of_int_eq_id id_def of_int_of_nat)
-
-lemma DEF_int_pow:
-  "op ^ = (\<lambda>u ua. floor (real u ^ ua))"
-  by (simp add: floor_power)
-
-lemma DEF_int_divides:
-  "op dvd = (\<lambda>(u :: int) ua. \<exists>x. ua = u * x)"
-  by (metis dvdE dvdI)
-
-lemma DEF_int_divides':
-  "(a :: int) dvd b = (\<exists>x. b = a * x)"
-  by (metis dvdE dvdI)
-
-definition "int_mod (u :: int) ua ub = (u dvd (ua - ub))"
-
-lemma int_mod_def':
-  "int_mod = (\<lambda>u ua ub. (u dvd (ua - ub)))"
-  by (simp add: int_mod_def [abs_def])
-
-lemma int_congruent:
-  "\<forall>x xa xb. int_mod xb x xa = (\<exists>d. x - xa = xb * d)"
-  unfolding int_mod_def'
-  by (auto simp add: DEF_int_divides')
-
-lemma int_congruent':
-  "\<forall>(x :: int) y n. (n dvd x - y) = (\<exists>d. x - y = n * d)"
-  using int_congruent[unfolded int_mod_def] .
-
-fun int_gcd where
-  "int_gcd ((a :: int), (b :: int)) = gcd a b"
-
-definition "hl_mod (k\<Colon>int) (l\<Colon>int) = (if 0 \<le> l then k mod l else k mod - l)"
-
-lemma hl_mod_nonneg:
-  "b \<noteq> 0 \<Longrightarrow> hl_mod a b \<ge> 0"
-  by (simp add: hl_mod_def)
-
-lemma hl_mod_lt_abs:
-  "b \<noteq> 0 \<Longrightarrow> hl_mod a b < abs b"
-  by (simp add: hl_mod_def)
-
-definition "hl_div k l = (if 0 \<le> l then k div l else -(k div (-l)))"
-
-lemma hl_mod_div:
-  "n \<noteq> (0\<Colon>int) \<Longrightarrow> m = hl_div m n * n + hl_mod m n"
-  unfolding hl_div_def hl_mod_def
-  by auto (metis zmod_zdiv_equality mult_commute mult_minus_left)
-
-lemma sth:
-  "(\<forall>(x :: int) y z. x + (y + z) = x + y + z) \<and>
-   (\<forall>(x :: int) y. x + y = y + x) \<and>
-   (\<forall>(x :: int). int 0 + x = x) \<and>
-   (\<forall>(x :: int) y z. x * (y * z) = x * y * z) \<and>
-   (\<forall>(x :: int) y. x * y = y * x) \<and>
-   (\<forall>(x :: int). int 1 * x = x) \<and>
-   (\<forall>(x :: int). int 0 * x = int 0) \<and>
-   (\<forall>(x :: int) y z. x * (y + z) = x * y + x * z) \<and>
-   (\<forall>(x :: int). x ^ 0 = int 1) \<and> (\<forall>(x :: int) n. x ^ Suc n = x * x ^ n)"
-  by (simp_all add: right_distrib)
-
-lemma INT_DIVISION:
-  "n ~= int 0 \<Longrightarrow> m = hl_div m n * n + hl_mod m n \<and> int 0 \<le> hl_mod m n \<and> hl_mod m n < abs n"
-  by (auto simp add: hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
-
-lemma INT_DIVMOD_EXIST_0:
-  "\<exists>q r. if n = int 0 then q = int 0 \<and> r = m
-         else int 0 \<le> r \<and> r < abs n \<and> m = q * n + r"
-  apply (rule_tac x="hl_div m n" in exI)
-  apply (rule_tac x="hl_mod m n" in exI)
-  apply (auto simp add: hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
-  unfolding hl_div_def hl_mod_def
-  by auto
-
-lemma DEF_div:
-  "hl_div = (SOME q. \<exists>r. \<forall>m n. if n = int 0 then q m n = int 0 \<and> r m n = m
-     else (int 0) \<le> (r m n) \<and> (r m n) < (abs n) \<and> m = ((q m n) * n) + (r m n))"
-  apply (rule some_equality[symmetric])
-  apply (rule_tac x="hl_mod" in exI)
-  apply (auto simp add: fun_eq_iff hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
-  apply (simp add: hl_div_def)
-  apply (simp add: hl_mod_def)
-  apply (drule_tac x="x" in spec)
-  apply (drule_tac x="xa" in spec)
-  apply (case_tac "0 = xa")
-  apply (simp add: hl_mod_def hl_div_def)
-  apply (case_tac "xa > 0")
-  apply (simp add: hl_mod_def hl_div_def)
-  apply (metis comm_semiring_1_class.normalizing_semiring_rules(24) div_mult_self2 not_less_iff_gr_or_eq order_less_le add_0 zdiv_eq_0_iff mult_commute)
-  apply (simp add: hl_mod_def hl_div_def)
-  by (metis add.comm_neutral add_pos_nonneg div_mult_self1 less_minus_iff minus_add minus_add_cancel minus_minus mult_zero_right not_square_less_zero zdiv_eq_0_iff div_minus_right)
-
-lemma DEF_rem:
-  "hl_mod = (SOME r. \<forall>m n. if n = int 0 then
-     (if 0 \<le> n then m div n else - (m div - n)) = int 0 \<and> r m n = m
-     else int 0 \<le> r m n \<and> r m n < abs n \<and>
-            m = (if 0 \<le> n then m div n else - (m div - n)) * n + r m n)"
-  apply (rule some_equality[symmetric])
-  apply (fold hl_div_def)
-  apply (auto simp add: fun_eq_iff hl_mod_nonneg hl_mod_lt_abs hl_mod_div)
-  apply (simp add: hl_div_def)
-  apply (simp add: hl_mod_def)
-  apply (drule_tac x="x" in spec)
-  apply (drule_tac x="xa" in spec)
-  apply (case_tac "0 = xa")
-  apply (simp add: hl_mod_def hl_div_def)
-  apply (case_tac "xa > 0")
-  apply (simp add: hl_mod_def hl_div_def)
-  apply (metis add_left_cancel mod_div_equality)
-  apply (simp add: hl_mod_def hl_div_def)
-  by (metis minus_mult_right mod_mult_self2 mod_pos_pos_trivial add_commute zminus_zmod mod_minus_right mult_commute)
-
-lemma DEF_int_gcd:
-  "int_gcd = (SOME d. \<forall>a b. (int 0) \<le> (d (a, b)) \<and> (d (a, b)) dvd a \<and>
-       (d (a, b)) dvd b \<and> (\<exists>x y. d (a, b) = (a * x) + (b * y)))"
-  apply (rule some_equality[symmetric])
-  apply auto
-  apply (metis bezout_int mult_commute)
-  apply (auto simp add: fun_eq_iff)
-  apply (drule_tac x="a" in spec)
-  apply (drule_tac x="b" in spec)
-  using gcd_greatest_int zdvd_antisym_nonneg
-  by auto
-
-definition "eqeq x y (r :: 'a \<Rightarrow> 'b \<Rightarrow> bool) = r x y"
-
-lemma INT_INTEGRAL:
-  "(\<forall>x. int 0 * x = int 0) \<and>
-   (\<forall>(x :: int) y z. (x + y = x + z) = (y = z)) \<and>
-   (\<forall>(w :: int) x y z. (w * y + x * z = w * z + x * y) = (w = x \<or> y = z))"
-  by (auto simp add: crossproduct_eq)
-
-end
-
--- a/src/HOL/Import/HOL_Light/HOLLightList.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,340 +0,0 @@
-(*  Title:      HOL/Import/HOL_Light/HOLLightList.thy
-    Author:     Cezary Kaliszyk
-*)
-
-header {* Compatibility theorems for HOL Light lists *}
-
-theory HOLLightList
-imports List
-begin
-
-lemma FINITE_SET_OF_LIST:
-  "finite (set l)"
-  by simp
-
-lemma AND_ALL2:
-  "(list_all2 P l m \<and> list_all2 Q l m) = list_all2 (\<lambda>x y. P x y \<and> Q x y) l m"
-  by (induct l m rule: list_induct2') auto
-
-lemma MEM_EXISTS_EL:
-  "(x \<in> set l) = (\<exists>i<length l. x = l ! i)"
-  by (auto simp add: in_set_conv_nth)
-
-lemma INJECTIVE_MAP:
-  "(\<forall>l m. map f l = map f m --> l = m) = (\<forall>x y. f x = f y --> x = y)"
-proof (intro iffI allI impI)
-  fix x y
-  assume "\<forall>l m. map f l = map f m \<longrightarrow> l = m" "f x = f y"
-  then show "x = y"
-    by (drule_tac x="[x]" in spec) (drule_tac x="[y]" in spec, simp)
-next
-  fix l m
-  assume a: "\<forall>x y. f x = f y \<longrightarrow> x = y"
-  assume "map f l = map f m"
-  then show "l = m"
-    by (induct l m rule: list_induct2') (simp_all add: a)
-qed
-
-lemma SURJECTIVE_MAP:
-  "(\<forall>m. EX l. map f l = m) = (\<forall>y. EX x. f x = y)"
-  apply (intro iffI allI)
-  apply (drule_tac x="[y]" in spec)
-  apply (elim exE)
-  apply (case_tac l)
-  apply simp
-  apply (rule_tac x="a" in exI)
-  apply simp
-  apply (induct_tac m)
-  apply simp
-  apply (drule_tac x="a" in spec)
-  apply (elim exE)
-  apply (rule_tac x="x # l" in exI)
-  apply simp
-  done
-
-lemma LENGTH_TL:
-  "l \<noteq> [] \<longrightarrow> length (tl l) = length l - 1"
-  by simp
-
-lemma DEF_APPEND:
-  "op @ = (SOME APPEND. (\<forall>l. APPEND [] l = l) \<and> (\<forall>h t l. APPEND (h # t) l = h # APPEND t l))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_REVERSE:
-  "rev = (SOME REVERSE. REVERSE [] = [] \<and> (\<forall>l x. REVERSE (x # l) = (REVERSE l) @ [x]))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_LENGTH:
-  "length = (SOME LENGTH. LENGTH [] = 0 \<and> (\<forall>h t. LENGTH (h # t) = Suc (LENGTH t)))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_MAP:
-  "map = (SOME MAP. (\<forall>f. MAP f [] = []) \<and> (\<forall>f h t. MAP f (h # t) = f h # MAP f t))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac xa)
-  apply simp_all
-  done
-
-lemma DEF_REPLICATE:
-  "replicate =
-    (SOME REPLICATE. (\<forall>x. REPLICATE 0 x = []) \<and> (\<forall>n x. REPLICATE (Suc n) x = x # REPLICATE n x))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma DEF_ITLIST:
-  "foldr = (SOME ITLIST. (\<forall>f b. ITLIST f [] b = b) \<and> (\<forall>h f t b. ITLIST f (h # t) b = f h (ITLIST f t b)))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac xa)
-  apply simp_all
-  done
-
-lemma DEF_ALL2: "list_all2 =
-  (SOME ALL2.
-    (\<forall>P l2. ALL2 P [] l2 = (l2 = [])) \<and>
-    (\<forall>h1 P t1 l2.
-      ALL2 P (h1 # t1) l2 = (if l2 = [] then False else P h1 (hd l2) \<and> ALL2 P t1 (tl l2))))"
-  apply (rule some_equality[symmetric])
-  apply (auto)
-  apply (case_tac l2, simp_all)
-  apply (case_tac l2, simp_all)
-  apply (case_tac l2, simp_all)
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa xb rule: list_induct2')
-  apply simp_all
-  done
-
-lemma ALL2:
- "list_all2 P [] [] = True \<and>
-  list_all2 P (h1 # t1) [] = False \<and>
-  list_all2 P [] (h2 # t2) = False \<and>
-  list_all2 P (h1 # t1) (h2 # t2) = (P h1 h2 \<and> list_all2 P t1 t2)"
-  by simp
-
-lemma DEF_FILTER:
-  "filter = (SOME FILTER. (\<forall>P. FILTER P [] = []) \<and>
-     (\<forall>h P t. FILTER P (h # t) = (if P h then h # FILTER P t else FILTER P t)))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff)
-  apply (induct_tac xa)
-  apply simp_all
-  done
-
-fun map2 where
-  "map2 f [] [] = []"
-| "map2 f (h1 # t1) (h2 # t2) = (f h1 h2) # (map2 f t1 t2)"
-
-lemma MAP2:
-  "map2 f [] [] = [] \<and> map2 f (h1 # t1) (h2 # t2) = f h1 h2 # map2 f t1 t2"
-  by simp
-
-fun fold2 where
-  "fold2 f [] [] b = b"
-| "fold2 f (h1 # t1) (h2 # t2) b = f h1 h2 (fold2 f t1 t2 b)"
-
-lemma ITLIST2:
-  "fold2 f [] [] b = b \<and> fold2 f (h1 # t1) (h2 # t2) b = f h1 h2 (fold2 f t1 t2 b)"
-  by simp
-
-definition [simp]: "list_el x xs = nth xs x"
-
-lemma ZIP:
-  "zip [] [] = [] \<and> zip (h1 # t1) (h2 # t2) = (h1, h2) # (zip t1 t2)"
-  by simp
-
-lemma LAST_CLAUSES:
-  "last [h] = h \<and> last (h # k # t) = last (k # t)"
-  by simp
-
-lemma DEF_NULL:
-  "List.null = (SOME NULL. NULL [] = True \<and> (\<forall>h t. NULL (h # t) = False))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: fun_eq_iff null_def)
-  apply (case_tac x)
-  apply simp_all
-  done
-
-lemma DEF_ALL:
-  "list_all = (SOME u. (\<forall>P. u P [] = True) \<and> (\<forall>h P t. u P (h # t) = (P h \<and> u P t)))"
-  apply (rule some_equality[symmetric])
-  apply auto[1]
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply simp_all
-  done
-
-lemma MAP_EQ:
-  "list_all (\<lambda>x. f x = g x) l \<longrightarrow> map f l = map g l"
-  by (induct l) auto
-
-definition [simp]: "list_mem x xs = List.member xs x"
-
-lemma DEF_MEM:
-  "list_mem = (SOME MEM. (\<forall>x. MEM x [] = False) \<and> (\<forall>h x t. MEM x (h # t) = (x = h \<or> MEM x t)))"
-  apply (rule some_equality[symmetric])
-  apply (auto simp add: member_def)[1]
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply (simp_all add: member_def)
-  done
-
-lemma DEF_EX:
-  "list_ex = (SOME u. (\<forall>P. u P [] = False) \<and> (\<forall>h P t. u P (h # t) = (P h \<or> u P t)))"
-  apply (rule some_equality[symmetric])
-  apply (auto)
-  apply (simp add: fun_eq_iff)
-  apply (intro allI)
-  apply (induct_tac xa)
-  apply (simp_all)
-  done
-
-lemma ALL_IMP:
-  "(\<forall>x. x \<in> set l \<and> P x \<longrightarrow> Q x) \<and> list_all P l \<longrightarrow> list_all Q l"
-  by (simp add: list_all_iff)
-
-lemma NOT_EX: "(\<not> list_ex P l) = list_all (\<lambda>x. \<not> P x) l"
-  by (simp add: list_all_iff list_ex_iff)
-
-lemma NOT_ALL: "(\<not> list_all P l) = list_ex (\<lambda>x. \<not> P x) l"
-  by (simp add: list_all_iff list_ex_iff)
-
-lemma ALL_MAP: "list_all P (map f l) = list_all (P o f) l"
-  by (simp add: list_all_iff)
-
-lemma ALL_T: "list_all (\<lambda>x. True) l"
-  by (simp add: list_all_iff)
-
-lemma MAP_EQ_ALL2: "list_all2 (\<lambda>x y. f x = f y) l m \<longrightarrow> map f l = map f m"
-  by (induct l m rule: list_induct2') simp_all
-
-lemma ALL2_MAP: "list_all2 P (map f l) l = list_all (\<lambda>a. P (f a) a) l"
-  by (induct l) simp_all
-
-lemma MAP_EQ_DEGEN: "list_all (\<lambda>x. f x = x) l --> map f l = l"
-  by (induct l) simp_all
-
-lemma ALL2_AND_RIGHT:
-   "list_all2 (\<lambda>x y. P x \<and> Q x y) l m = (list_all P l \<and> list_all2 Q l m)"
-  by (induct l m rule: list_induct2') auto
-
-lemma ITLIST_EXTRA:
-  "foldr f (l @ [a]) b = foldr f l (f a b)"
-  by simp
-
-lemma ALL_MP:
-  "list_all (\<lambda>x. P x \<longrightarrow> Q x) l \<and> list_all P l \<longrightarrow> list_all Q l"
-  by (simp add: list_all_iff)
-
-lemma AND_ALL:
-  "(list_all P l \<and> list_all Q l) = list_all (\<lambda>x. P x \<and> Q x) l"
-  by (auto simp add: list_all_iff)
-
-lemma EX_IMP:
-  "(\<forall>x. x\<in>set l \<and> P x \<longrightarrow> Q x) \<and> list_ex P l \<longrightarrow> list_ex Q l"
-  by (auto simp add: list_ex_iff)
-
-lemma ALL_MEM:
-  "(\<forall>x. x\<in>set l \<longrightarrow> P x) = list_all P l"
-  by (auto simp add: list_all_iff)
-
-lemma EX_MAP:
-  "ALL P f l. list_ex P (map f l) = list_ex (P o f) l"
-  by (simp add: list_ex_iff)
-
-lemma EXISTS_EX:
-  "\<forall>P l. (EX x. list_ex (P x) l) = list_ex (\<lambda>s. EX x. P x s) l"
-  by (auto simp add: list_ex_iff)
-
-lemma FORALL_ALL:
-  "\<forall>P l. (\<forall>x. list_all (P x) l) = list_all (\<lambda>s. \<forall>x. P x s) l"
-  by (auto simp add: list_all_iff)
-
-lemma MEM_APPEND: "\<forall>x l1 l2. (x\<in>set (l1 @ l2)) = (x\<in>set l1 \<or> x\<in>set l2)"
-  by simp
-
-lemma MEM_MAP: "\<forall>f y l. (y\<in>set (map f l)) = (EX x. x\<in>set l \<and> y = f x)"
-  by auto
-
-lemma MEM_FILTER: "\<forall>P l x. (x\<in>set (filter P l)) = (P x \<and> x\<in>set l)"
-  by auto
-
-lemma EX_MEM: "(EX x. P x \<and> x\<in>set l) = list_ex P l"
-  by (auto simp add: list_ex_iff)
-
-lemma ALL2_MAP2:
-  "list_all2 P (map f l) (map g m) = list_all2 (\<lambda>x y. P (f x) (g y)) l m"
-  by (simp add: list_all2_map1 list_all2_map2)
-
-lemma ALL2_ALL:
-  "list_all2 P l l = list_all (\<lambda>x. P x x) l"
-  by (induct l) simp_all
-
-lemma LENGTH_MAP2:
-  "length l = length m \<longrightarrow> length (map2 f l m) = length m"
-  by (induct l m rule: list_induct2') simp_all
-
-lemma DEF_set_of_list:
-  "set = (SOME sol. sol [] = {} \<and> (\<forall>h t. sol (h # t) = insert h (sol t)))"
-  apply (rule some_equality[symmetric])
-  apply (simp_all)
-  apply (rule ext)
-  apply (induct_tac x)
-  apply simp_all
-  done
-
-lemma IN_SET_OF_LIST:
-  "(x : set l) = (x : set l)"
-  by simp
-
-lemma DEF_BUTLAST:
-  "butlast = (SOME B. B [] = [] \<and> (\<forall>h t. B (h # t) = (if t = [] then [] else h # B t)))"
-  apply (rule some_equality[symmetric])
-  apply auto
-  apply (rule ext)
-  apply (induct_tac x)
-  apply auto
-  done
-
-lemma MONO_ALL:
-  "(ALL x. P x \<longrightarrow> Q x) \<longrightarrow> list_all P l \<longrightarrow> list_all Q l"
-  by (simp add: list_all_iff)
-
-lemma EL_TL: "l \<noteq> [] \<Longrightarrow> tl l ! x = l ! (x + 1)"
-  by (induct l) (simp_all)
-
-(* Assume the same behaviour outside of the usual domain.
-   For HD, LAST, EL it follows from "undefined = SOME _. False".
-
-   The definitions of TL and ZIP are different for empty lists.
- *)
-axiomatization where
-  DEF_HD: "hd = (SOME HD. \<forall>t h. HD (h # t) = h)"
-
-axiomatization where
-  DEF_LAST: "last =
-    (SOME LAST. \<forall>h t. LAST (h # t) = (if t = [] then h else LAST t))"
-
-axiomatization where
-  DEF_EL: "list_el =
-    (SOME EL. (\<forall>l. EL 0 l = hd l) \<and> (\<forall>n l. EL (Suc n) l = EL n (tl l)))"
-
-end
--- a/src/HOL/Import/HOL_Light/HOLLightReal.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,329 +0,0 @@
-(*  Title:      HOL/Import/HOL_Light/HOLLightReal.thy
-    Author:     Cezary Kaliszyk
-*)
-
-header {* Compatibility theorems for HOL Light reals *}
-
-theory HOLLightReal imports Real begin
-
-lemma REAL_OF_NUM_MAX:
-  "max (real (m :: nat)) (real n) = real (max m n)"
-  by simp
-
-lemma REAL_OF_NUM_MIN:
-  "min (real (m :: nat)) (real n) = real (min m n)"
-  by simp
-
-lemma REAL_POLY_NEG_CLAUSES:
-  "(\<forall>(x :: real). - x = - 1 * x) \<and> (\<forall>(x :: real) y. x - y = x + - 1 * y)"
-  by simp
-
-lemma REAL_MUL_AC:
-  "(m :: real) * n = n * m \<and> m * n * p = m * (n * p) \<and> m * (n * p) = n * (m * p)"
-  by simp
-
-lemma REAL_EQ_ADD_LCANCEL_0:
-  "((x :: real) + y = x) = (y = 0)"
-  by simp
-
-lemma REAL_EQ_ADD_RCANCEL_0:
-  "((x :: real) + y = y) = (x = 0)"
-  by simp
-
-lemma REAL_LT_ANTISYM:
-  "\<not> ((x :: real) < y \<and> y < x)"
-  by simp
-
-lemma REAL_LET_ANTISYM:
-  "\<not> ((x :: real) \<le> y \<and> y < x)"
-  by simp
-
-lemma REAL_LT_NEGTOTAL:
-  "(x :: real) = 0 \<or> 0 < x \<or> 0 < - x"
-  by auto
-
-lemma REAL_LT_ADDNEG:
-  "((y :: real) < x + - z) = (y + z < x)"
-  by auto
-
-lemma REAL_LT_ADDNEG2:
-  "((x :: real) + - y < z) = (x < z + y)"
-  by auto
-
-lemma REAL_LT_ADD1:
-  "(x :: real) \<le> y \<longrightarrow> x < y + 1"
-  by simp
-
-lemma REAL_SUB_ADD2:
-  "(y :: real) + (x - y) = x"
-  by simp
-
-lemma REAL_ADD_SUB:
-  "(x :: real) + y - x = y"
-  by simp
-
-lemma REAL_NEG_EQ:
-  "(- (x :: real) = y) = (x = - y)"
-  by auto
-
-lemma REAL_LE_ADDR:
-  "((x :: real) \<le> x + y) = (0 \<le> y)"
-  by simp
-
-lemma REAL_LE_ADDL:
-  "((y :: real) \<le> x + y) = (0 \<le> x)"
-  by simp
-
-lemma REAL_LT_ADDR:
-  "((x :: real) < x + y) = (0 < y)"
-  by simp
-
-lemma REAL_LT_ADDL:
-  "((y :: real) < x + y) = (0 < x)"
-  by simp
-
-lemma REAL_SUB_SUB:
-  "(x :: real) - y - x = - y"
-  by simp
-
-lemma REAL_SUB_LNEG:
-  "- (x :: real) - y = - (x + y)"
-  by simp
-
-lemma REAL_SUB_NEG2:
-  "- (x :: real) - - y = y - x"
-  by simp
-
-lemma REAL_SUB_TRIANGLE:
-  "(a :: real) - b + (b - c) = a - c"
-  by simp
-
-lemma REAL_SUB_SUB2:
-  "(x :: real) - (x - y) = y"
-  by simp
-
-lemma REAL_ADD_SUB2:
-  "(x :: real) - (x + y) = - y"
-  by simp
-
-lemma REAL_POS_NZ:
-  "0 < (x :: real) \<longrightarrow> x \<noteq> 0"
-  by simp
-
-lemma REAL_DIFFSQ:
-  "((x :: real) + y) * (x - y) = x * x - y * y"
-  by (simp add: comm_semiring_1_class.normalizing_semiring_rules(7) right_distrib mult_diff_mult)
-
-lemma REAL_ABS_TRIANGLE_LE:
-  "abs (x :: real) + abs (y - x) \<le> z \<longrightarrow> abs y \<le> z"
-  by auto
-
-lemma REAL_ABS_TRIANGLE_LT:
-  "abs (x :: real) + abs (y - x) < z \<longrightarrow> abs y < z"
-  by auto
-
-lemma REAL_ABS_REFL:
-  "(abs (x :: real) = x) = (0 \<le> x)"
-  by auto
-
-lemma REAL_ABS_BETWEEN:
-  "(0 < (d :: real) \<and> x - d < y \<and> y < x + d) = (abs (y - x) < d)"
-  by auto
-
-lemma REAL_ABS_BOUND:
-  "abs ((x :: real) - y) < d \<longrightarrow> y < x + d"
-  by auto
-
-lemma REAL_ABS_STILLNZ:
-  "abs ((x :: real) - y) < abs y \<longrightarrow> x \<noteq> 0"
-  by auto
-
-lemma REAL_ABS_CASES:
-  "(x :: real) = 0 \<or> 0 < abs x"
-  by simp
-
-lemma REAL_ABS_BETWEEN1:
-  "(x :: real) < z \<and> abs (y - x) < z - x \<longrightarrow> y < z"
-  by auto
-
-lemma REAL_ABS_SIGN:
-  "abs ((x :: real) - y) < y \<longrightarrow> 0 < x"
-  by auto
-
-lemma REAL_ABS_SIGN2:
-  "abs ((x :: real) - y) < - y \<longrightarrow> x < 0"
-  by auto
-
-lemma REAL_ABS_CIRCLE:
-  "abs (h :: real) < abs y - abs x \<longrightarrow> abs (x + h) < abs y"
-  by auto
-
-lemma REAL_BOUNDS_LT:
-  "(- (k :: real) < x \<and> x < k) = (abs x < k)"
-  by auto
-
-lemma REAL_MIN_MAX:
-  "min (x :: real) y = - max (- x) (- y)"
-  by auto
-
-lemma REAL_MAX_MIN:
-  "max (x :: real) y = - min (- x) (- y)"
-  by auto
-
-lemma REAL_MAX_MAX:
-  "(x :: real) \<le> max x y \<and> y \<le> max x y"
-  by simp
-
-lemma REAL_MIN_MIN:
-  "min (x :: real) y \<le> x \<and> min x y \<le> y"
-  by simp
-
-lemma REAL_MAX_ACI:
-  "max (x :: real) y = max y x \<and>
-   max (max x y) z = max x (max y z) \<and>
-   max x (max y z) = max y (max x z) \<and> max x x = x \<and> max x (max x y) = max x y"
-  by auto
-
-
-lemma REAL_MIN_ACI:
-  "min (x :: real) y = min y x \<and>
-   min (min x y) z = min x (min y z) \<and>
-   min x (min y z) = min y (min x z) \<and> min x x = x \<and> min x (min x y) = min x y"
-  by auto
-
-lemma REAL_EQ_MUL_RCANCEL:
-  "((x :: real) * z = y * z) = (x = y \<or> z = 0)"
-  by auto
-
-lemma REAL_MUL_LINV_UNIQ:
-  "(x :: real) * y = 1 \<longrightarrow> inverse y = x"
-  by (metis inverse_inverse_eq inverse_unique)
-
-lemma REAL_DIV_RMUL:
-  "(y :: real) \<noteq> 0 \<longrightarrow> x / y * y = x"
-  by simp
-
-lemma REAL_DIV_LMUL:
-  "(y :: real) \<noteq> 0 \<longrightarrow> y * (x / y) = x"
-  by simp
-
-lemma REAL_LT_IMP_NZ:
-  "0 < (x :: real) \<longrightarrow> x \<noteq> 0"
-  by simp
-
-lemma REAL_LT_LCANCEL_IMP:
-  "0 < (x :: real) \<and> x * y < x * z \<longrightarrow> y < z"
-  by (auto simp add: mult_less_cancel_left_disj not_less_iff_gr_or_eq)
-
-lemma REAL_LT_RCANCEL_IMP:
-  "0 < (z :: real) \<and> x * z < y * z \<longrightarrow> x < y"
-  by (auto simp add: mult_less_cancel_left_disj not_less_iff_gr_or_eq)
-
-lemma REAL_MUL_POS_LE:
-  "(0 \<le> (x :: real) * y) = (x = 0 \<or> y = 0 \<or> 0 < x \<and> 0 < y \<or> x < 0 \<and> y < 0)"
-  by (metis less_eq_real_def mult_eq_0_iff zero_le_mult_iff)
-
-lemma REAL_EQ_RDIV_EQ:
-  "0 < (z :: real) \<longrightarrow> (x = y / z) = (x * z = y)"
-  by auto
-
-lemma REAL_EQ_LDIV_EQ:
-  "0 < (z :: real) \<longrightarrow> (x / z = y) = (x = y * z)"
-  by auto
-
-lemma REAL_SUB_INV:
-  "(x :: real) \<noteq> 0 \<and> y \<noteq> 0 \<longrightarrow> inverse x - inverse y = (y - x) / (x * y)"
-  by (simp add: division_ring_inverse_diff divide_real_def)
-
-lemma REAL_DOWN:
-  "0 < (d :: real) \<longrightarrow> (\<exists>e>0. e < d)"
-  by (intro impI exI[of _ "d / 2"]) simp
-
-lemma REAL_POW_MONO_LT:
-  "1 < (x :: real) \<and> m < n \<longrightarrow> x ^ m < x ^ n"
-  by simp
-
-lemma REAL_POW_MONO:
-  "1 \<le> (x :: real) \<and> m \<le> n \<longrightarrow> x ^ m \<le> x ^ n"
-  by (cases "m < n", cases "x = 1") auto
-
-lemma REAL_EQ_LCANCEL_IMP:
-  "(z :: real) \<noteq> 0 \<and> z * x = z * y \<longrightarrow> x = y"
-  by auto
-
-lemma REAL_LE_DIV:
-  "0 \<le> (x :: real) \<and> 0 \<le> y \<longrightarrow> 0 \<le> x / y"
-  by (simp add: zero_le_divide_iff)
-
-lemma REAL_10: "(1::real) \<noteq> 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 \<noteq> (0::real) \<Longrightarrow> inverse x * x = 1"
-  by simp
-
-lemma REAL_LT_TOTAL: "((x::real) = y) \<or> x < y \<or> y < x"
-  by auto
-
-lemma real_lte: "((x::real) \<le> y) = (\<not>(y < x))"
-  by auto
-
-lemma real_of_num: "((0::real) = 0) \<and> (!n. real (Suc n) = real n + 1)"
-  by (simp add: real_of_nat_Suc)
-
-lemma abs: "abs (x::real) = (if 0 \<le> x then x else -x)"
-  by (simp add: abs_if)
-
-lemma pow: "(!x::real. x ^ 0 = 1) \<and> (!x::real. \<forall>n. x ^ (Suc n) = x * x ^ n)"
-  by simp
-
-lemma REAL_POLY_CLAUSES:
-  "(\<forall>(x :: real) y z. x + (y + z) = x + y + z) \<and>
-   (\<forall>(x :: real) y. x + y = y + x) \<and>
-   (\<forall>(x :: real). 0 + x = x) \<and>
-   (\<forall>(x :: real) y z. x * (y * z) = x * y * z) \<and>
-   (\<forall>(x :: real) y. x * y = y * x) \<and>
-   (\<forall>(x :: real). 1 * x = x) \<and>
-   (\<forall>(x :: real). 0 * x = 0) \<and>
-   (\<forall>(x :: real) y z. x * (y + z) = x * y + x * z) \<and>
-   (\<forall>(x :: real). x ^ 0 = 1) \<and> (\<forall>(x :: real) n. x ^ Suc n = x * x ^ n)"
-  by (auto simp add: right_distrib)
-
-lemma REAL_COMPLETE:
-  "(\<exists>(x :: real). P x) \<and> (\<exists>(M :: real). \<forall>x. P x \<longrightarrow> x \<le> M) \<longrightarrow>
-   (\<exists>M. (\<forall>x. P x \<longrightarrow> x \<le> M) \<and>
-          (\<forall>M'. (\<forall>x. P x \<longrightarrow> x \<le> M') \<longrightarrow> M \<le> M'))"
-  using complete_real[unfolded Ball_def, of "Collect P"] by auto
-
-lemma REAL_COMPLETE_SOMEPOS:
-  "(\<exists>(x :: real). P x \<and> 0 \<le> x) \<and> (\<exists>M. \<forall>x. P x \<longrightarrow> x \<le> M) \<longrightarrow>
-   (\<exists>M. (\<forall>x. P x \<longrightarrow> x \<le> M) \<and>
-          (\<forall>M'. (\<forall>x. P x \<longrightarrow> x \<le> M') \<longrightarrow> M \<le> M'))"
-  using REAL_COMPLETE by auto
-
-lemma REAL_ADD_AC:
-  "(m :: real) + n = n + m \<and> m + n + p = m + (n + p) \<and> m + (n + p) = n + (m + p)"
-  by simp
-
-lemma REAL_LE_RNEG:
-  "((x :: real) \<le> - y) = (x + y \<le> 0)"
-  by auto
-
-lemma REAL_LE_NEGTOTAL:
-  "0 \<le> (x :: real) \<or> 0 \<le> - x"
-  by auto
-
-lemma DEF_real_sgn:
-  "sgn = (\<lambda>u. if (0 :: real) < u then 1 else if u < 0 then - 1 else 0)"
-  by (simp add: ext)
-
-end
-
--- a/src/HOL/Import/HOL_Light/Imported.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-theory Imported
-imports "Generated/HOLLight"
-begin
-
-end
-
--- a/src/HOL/Import/HOL_Light/ROOT.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2 +0,0 @@
-
-use_thy "Compatibility";
--- a/src/HOL/Import/HOL_Light/Template/GenHOLLight.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,220 +0,0 @@
-(*  Title:      HOL/Import/HOL_Light/Template/GenHOLLight.thy
-    Author:     Steven Obua, TU Muenchen
-    Author:     Cezary Kaliszyk
-*)
-
-theory GenHOLLight
-imports "../../Importer" "../Compatibility"
-begin
-
-import_segment "hollight"
-
-setup_dump "../Generated" "HOLLight"
-
-append_dump {*theory HOL4Base
-imports "../../Importer" "../Compatibility"
-begin
-*}
-
-import_theory "~~/src/HOL/Import/HOL_Light/Generated" hollight
-
-;ignore_thms
-  (* Unit type *)
-  TYDEF_1 DEF_one
-  (* Product type *)
-  TYDEF_prod "DEF_," DEF_mk_pair REP_ABS_PAIR
-  (* Option type *)
-  TYDEF_option DEF_NONE DEF_SOME
-  (* Sum type *)
-  TYDEF_sum DEF_INL DEF_INR DEF_OUTL DEF_OUTR
-  (* Naturals *)
-  TYDEF_num DEF__0 DEF_SUC
-  (* Lists *)
-  TYDEF_list DEF_NIL DEF_CONS DEF_HD DEF_TL DEF_MAP2 DEF_ITLIST2 ALL_MAP EX_MAP
-  DEF_ASSOC MEM_ASSOC DEF_ZIP EL_TL
-
-  (* list_of_set uses Isabelle lists with HOLLight CARD *)
-  DEF_list_of_set LIST_OF_SET_PROPERTIES SET_OF_LIST_OF_SET LENGTH_LIST_OF_SET
-  MEM_LIST_OF_SET HAS_SIZE_SET_OF_LIST FINITE_SET_OF_LIST
-  (* UNIV *)
-  DIMINDEX_FINITE_SUM  DIMINDEX_HAS_SIZE_FINITE_SUM FSTCART_PASTECART
-  SNDCART_PASTECART PASTECART_FST_SND PASTECART_EQ FORALL_PASTECART EXISTS_PASTECART
-  (* Reals *)
-  (* TYDEF_real DEF_real_of_num DEF_real_neg DEF_real_add DEF_real_mul DEF_real_le
-  DEF_real_inv REAL_HREAL_LEMMA1 REAL_HREAL_LEMMA2 *)
-  (* Integers *)
-  (* TYDEF_int DEF_int_divides DEF_int_coprime*)
-  (* HOLLight CARD and CASEWISE with Isabelle lists *)
-  CARD_SET_OF_LIST_LE CASEWISE CASEWISE_CASES RECURSION_CASEWISE CASEWISE_WORKS
-  WF_REC_CASES RECURSION_CASEWISE_PAIRWISE
-
-;type_maps
-  bool > HOL.bool
-  "fun" > "fun"
-  N_1 >  Product_Type.unit
-  prod > Product_Type.prod
-  ind > Nat.ind
-  num > Nat.nat
-  sum > Sum_Type.sum
-  option > Datatype.option
-  list > List.list
-(*real > RealDef.real *)
-(*int > Int.int *)
-
-;const_renames
-  "==" > "eqeq"
-  "ALL" > list_ALL
-  "EX" > list_EX
-
-;const_maps
-  T > HOL.True
-  F > HOL.False
-  "=" > HOL.eq
-  "==>" > HOL.implies
-  "/\\" > HOL.conj
-  "\\/" > HOL.disj
-  "!" > HOL.All
-  "?" > HOL.Ex
-  "?!" > HOL.Ex1
-  "~" > HOL.Not
-  COND > HOL.If
-  ONE_ONE > Fun.inj
-  ONTO > Fun.surj
-  o > Fun.comp
-  "@" > Hilbert_Choice.Eps
-  CHOICE > Hilbert_Choice.Eps
-  I > Fun.id
-  one > Product_Type.Unity
-  LET > HOLLightCompat.LET
-  mk_pair > Product_Type.Pair_Rep
-  "," > Product_Type.Pair
-  FST > Product_Type.fst
-  SND > Product_Type.snd
-  CURRY > Product_Type.curry
-  "_0" > Groups.zero_class.zero :: nat
-  SUC > Nat.Suc
-  PRE > HOLLightCompat.Pred
-  NUMERAL > HOLLightCompat.NUMERAL
-  mk_num > Fun.id
-  "+" > Groups.plus_class.plus :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  "*" > Groups.times_class.times :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  "-" > Groups.minus_class.minus :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  "<" > Orderings.ord_class.less :: "nat \<Rightarrow> nat \<Rightarrow> bool"
-  "<=" > Orderings.ord_class.less_eq :: "nat \<Rightarrow> nat \<Rightarrow> bool"
-  ">" > Orderings.ord_class.greater :: "nat \<Rightarrow> nat \<Rightarrow> bool"
-  ">=" > Orderings.ord_class.greater_eq :: "nat \<Rightarrow> nat \<Rightarrow> bool"
-  EXP > Power.power_class.power :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  MAX > Orderings.ord_class.max :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  MIN > Orderings.ord_class.min :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  DIV > Divides.div_class.div :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  MOD > Divides.div_class.mod :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-  BIT0 > HOLLightCompat.NUMERAL_BIT0
-  BIT1 > HOLLightCompat.NUMERAL_BIT1
-  INL > Sum_Type.Inl
-  INR > Sum_Type.Inr
-  OUTL > HOLLightCompat.OUTL
-  OUTR > HOLLightCompat.OUTR
-  NONE > Datatype.None
-  SOME > Datatype.Some
-  EVEN > Parity.even_odd_class.even :: "nat \<Rightarrow> bool"
-  ODD > HOLLightCompat.ODD
-  FACT > Fact.fact_class.fact :: "nat \<Rightarrow> nat"
-  WF > Wellfounded.wfP
-  NIL > List.list.Nil
-  CONS > List.list.Cons
-  APPEND > List.append
-  REVERSE > List.rev
-  LENGTH > List.length
-  MAP > List.map
-  LAST > List.last
-  BUTLAST > List.butlast
-  REPLICATE > List.replicate
-  ITLIST > List.foldr
-  list_ALL > List.list_all
-  ALL2 > List.list_all2
-  list_EX > List.list_ex
-  FILTER > List.filter
-  NULL > List.null
-  HD > List.hd
-  TL > List.tl
-  EL > HOLLightList.list_el
-  ZIP > List.zip
-  MAP2 > HOLLightList.map2
-  ITLIST2 > HOLLightList.fold2
-  MEM > HOLLightList.list_mem
-  set_of_list > List.set
-  IN > Set.member
-  INSERT > Set.insert
-  EMPTY > Orderings.bot_class.bot :: "'a \<Rightarrow> bool"
-  GABS > Hilbert_Choice.Eps
-  GEQ > HOL.eq
-  GSPEC > Set.Collect
-  SETSPEC > HOLLightCompat.SETSPEC
-  UNION > Lattices.sup_class.sup :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
-  UNIONS > Complete_Lattices.Sup_class.Sup :: "(('a \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
-  INTER > Lattices.inf_class.inf :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
-  INTERS > Complete_Lattices.Inf_class.Inf :: "(('a \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
-  DIFF > Groups.minus_class.minus :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
-  SUBSET > Orderings.ord_class.less_eq :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
-  PSUBSET > Orderings.ord_class.less :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
-  DELETE > HOLLightCompat.DELETE
-  DISJOINT > HOLLightCompat.DISJOINT
-  IMAGE > Set.image
-  FINITE > Finite_Set.finite
-  INFINITE > HOLLightCompat.INFINITE
-  ".." > HOLLightCompat.dotdot
-  UNIV > Orderings.top_class.top :: "'a \<Rightarrow> bool"
-  MEASURE > HOLLightCompat.MEASURE
-(*real_of_num > RealDef.real :: "nat => real"
-  real_neg > Groups.uminus_class.uminus :: "real => real"
-  real_inv > Fields.inverse_class.inverse :: "real => real"
-  real_add > Groups.plus_class.plus :: "real => real => real"
-  real_sub > Groups.minus_class.minus :: "real => real => real"
-  real_mul > Groups.times_class.times :: "real => real => real"
-  real_div > Fields.inverse_class.divide :: "real => real => real"
-  real_lt > Orderings.ord_class.less :: "real \<Rightarrow> real \<Rightarrow> bool"
-  real_le > Orderings.ord_class.less_eq :: "real \<Rightarrow> real \<Rightarrow> bool"
-  real_gt > Orderings.ord_class.greater :: "real \<Rightarrow> real \<Rightarrow> bool"
-  real_ge > Orderings.ord_class.greater_eq :: "real \<Rightarrow> real \<Rightarrow> bool"
-  real_pow > Power.power_class.power :: "real \<Rightarrow> nat \<Rightarrow> real"
-  real_abs > Groups.abs_class.abs :: "real \<Rightarrow> real"
-  real_max > Orderings.ord_class.max :: "real \<Rightarrow> real \<Rightarrow> real"
-  real_min > Orderings.ord_class.min :: "real \<Rightarrow> real \<Rightarrow> real"
-  real_sgn > Groups.sgn_class.sgn :: "real \<Rightarrow> real"*)
-(*real_of_int > RealDef.real :: "int => real"
-  int_of_real > Archimedean_Field.floor :: "real \<Rightarrow> int"
-  dest_int > RealDef.real :: "int => real"
-  mk_int > Archimedean_Field.floor :: "real \<Rightarrow> int"
-  int_lt > Orderings.ord_class.less :: "int \<Rightarrow> int \<Rightarrow> bool"
-  int_le > Orderings.ord_class.less_eq :: "int \<Rightarrow> int \<Rightarrow> bool"
-  int_gt > Orderings.ord_class.greater :: "int \<Rightarrow> int \<Rightarrow> bool"
-  int_ge > Orderings.ord_class.greater_eq :: "int \<Rightarrow> int \<Rightarrow> bool"
-  int_of_num > Nat.semiring_1_class.of_nat :: "nat \<Rightarrow> int"
-  int_neg > Groups.uminus_class.uminus :: "int \<Rightarrow> int"
-  int_add > Groups.plus_class.plus :: "int => int => int"
-  int_sub > Groups.minus_class.minus :: "int => int => int"
-  int_mul > Groups.times_class.times :: "int => int => int"
-  int_abs > Groups.abs_class.abs :: "int \<Rightarrow> int"
-  int_max > Orderings.ord_class.max :: "int \<Rightarrow> int \<Rightarrow> int"
-  int_min > Orderings.ord_class.min :: "int \<Rightarrow> int \<Rightarrow> int"
-  int_sgn > Groups.sgn_class.sgn :: "int \<Rightarrow> int"
-  int_pow > Power.power_class.power :: "int \<Rightarrow> nat \<Rightarrow> int"
-  int_div > HOLLightInt.hl_div :: "int \<Rightarrow> int \<Rightarrow> int"
-  div > HOLLightInt.hl_div :: "int \<Rightarrow> int \<Rightarrow> int"
-  mod_int > HOLLightInt.hl_mod :: "int \<Rightarrow> int \<Rightarrow> int"
-  rem > HOLLightInt.hl_mod :: "int \<Rightarrow> int \<Rightarrow> int"
-  int_divides > Rings.dvd_class.dvd :: "int \<Rightarrow> int \<Rightarrow> bool"
-  int_mod > HOLLightInt.int_mod :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"
-  int_gcd > HOLLightInt.int_gcd :: "int \<times> int \<Rightarrow> int"
-  int_coprime > HOLLightInt.int_coprime :: "int \<times> int \<Rightarrow> bool"
-  eqeq > HOLLightInt.eqeq*)
-
-;end_import
-
-;append_dump "end"
-
-;flush_dump
-
-;import_segment ""
-
-end
--- a/src/HOL/Import/HOL_Light/generate.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2 +0,0 @@
-
-use_thy "Generate";
--- a/src/HOL/Import/HOL_Light/imported.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2 +0,0 @@
-
-Unsynchronized.setmp quick_and_dirty true use_thy "Imported";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL_Light_Import.thy	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,15 @@
+(*  Title:      HOL/Import/HOL_Light_Import.thy
+    Author:     Cezary Kaliszyk, University of Innsbruck
+    Author:     Alexander Krauss, QAware GmbH
+*)
+
+header {* Main HOL Light importer *}
+
+theory HOL_Light_Import
+imports HOL_Light_Maps
+begin
+
+import_file "$HOL_LIGHT_DUMP"
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL_Light_Maps.thy	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,324 @@
+(*  Title:      HOL/Import/HOL_Light_Maps.thy
+    Author:     Cezary Kaliszyk, University of Innsbruck
+    Author:     Alexander Krauss, QAware GmbH
+
+Based on earlier code by Steven Obua and Sebastian Skalberg
+*)
+
+header {* Type and constant mappings of HOL Light importer *}
+
+theory HOL_Light_Maps
+imports Import_Setup
+begin
+
+lemma [import_const T]:
+  "True = ((\<lambda>p :: bool. p) = (\<lambda>p. p))"
+  by simp
+
+lemma [import_const "/\\"]:
+  "(op \<and>) = (\<lambda>p q. (\<lambda>f. f p q \<Colon> bool) = (\<lambda>f. f True True))"
+  by metis
+
+lemma [import_const "==>"]:
+  "(op \<longrightarrow>) = (\<lambda>(p\<Colon>bool) q\<Colon>bool. (p \<and> q) = p)"
+  by auto
+
+lemma [import_const "!"]:
+  "All = (\<lambda>P\<Colon>'A \<Rightarrow> bool. P = (\<lambda>x\<Colon>'A. True))"
+  by auto
+
+lemma [import_const "?"]:
+  "Ex = (\<lambda>P\<Colon>'A \<Rightarrow> bool. All (\<lambda>q\<Colon>bool. (All (\<lambda>x\<Colon>'A\<Colon>type. (P x) \<longrightarrow> q)) \<longrightarrow> q))"
+  by auto
+
+lemma [import_const "\\/"]:
+  "(op \<or>) = (\<lambda>p q. \<forall>r. (p \<longrightarrow> r) \<longrightarrow> (q \<longrightarrow> r) \<longrightarrow> r)"
+  by auto
+
+lemma [import_const F]:
+  "False = (\<forall>p. p)"
+  by auto
+
+lemma [import_const "~"]:
+  "Not = (\<lambda>p. p \<longrightarrow> False)"
+  by simp
+
+lemma [import_const "?!"]:
+  "Ex1 = (\<lambda>P\<Colon>'A \<Rightarrow> bool. Ex P \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y))"
+  by auto
+
+lemma [import_const "_FALSITY_"]: "False = False"
+  by simp
+
+lemma hl_ax1: "\<forall>t\<Colon>'A \<Rightarrow> 'B. (\<lambda>x. t x) = t"
+  by metis
+
+lemma hl_ax2: "\<forall>P x\<Colon>'A. P x \<longrightarrow> P (Eps P)"
+  by (auto intro: someI)
+
+lemma [import_const COND]:
+  "If = (\<lambda>t (t1 :: 'A) t2. SOME x. (t = True \<longrightarrow> x = t1) \<and> (t = False \<longrightarrow> x = t2))"
+  unfolding fun_eq_iff by auto
+
+lemma [import_const o]:
+  "(op \<circ>) = (\<lambda>(f\<Colon>'B \<Rightarrow> 'C) g x\<Colon>'A. f (g x))"
+  unfolding fun_eq_iff by simp
+
+lemma [import_const I]: "id = (\<lambda>x\<Colon>'A. x)"
+  by auto
+
+lemma [import_type 1 one_ABS one_REP]:
+  "type_definition Rep_unit Abs_unit (Collect (\<lambda>b. b))"
+  by (metis (full_types) Collect_cong singleton_conv2 type_definition_unit)
+
+lemma [import_const one]: "() = (SOME x\<Colon>unit. True)"
+  by auto
+
+lemma [import_const mk_pair]:
+  "Pair_Rep = (\<lambda>(x\<Colon>'A) (y\<Colon>'B) (a\<Colon>'A) b\<Colon>'B. a = x \<and> b = y)"
+  by (simp add: Pair_Rep_def fun_eq_iff)
+
+lemma [import_type prod ABS_prod REP_prod]:
+  "type_definition Rep_prod Abs_prod (Collect (\<lambda>x\<Colon>'A \<Rightarrow> 'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b))"
+  using type_definition_prod[unfolded Product_Type.prod_def] by simp
+
+lemma [import_const ","]: "Pair = (\<lambda>(x\<Colon>'A) y\<Colon>'B. Abs_prod (Pair_Rep x y))"
+  by (metis Pair_def)
+
+lemma [import_const FST]:
+  "fst = (\<lambda>p\<Colon>'A \<times> 'B. SOME x\<Colon>'A. \<exists>y\<Colon>'B. p = (x, y))"
+  by auto
+
+lemma [import_const SND]:
+  "snd = (\<lambda>p\<Colon>'A \<times> 'B. SOME y\<Colon>'B. \<exists>x\<Colon>'A. p = (x, y))"
+  by auto
+
+(*lemma [import_const CURRY]:
+  "curry = (\<lambda>(f\<Colon>'A \<times> 'B \<Rightarrow> 'C) x y. f (x, y))"
+  using curry_def .*)
+
+lemma [import_const ONE_ONE : Fun.inj]:
+  "inj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>x1 x2. f x1 = f x2 \<longrightarrow> x1 = x2)"
+  by (auto simp add: fun_eq_iff inj_on_def)
+
+lemma [import_const ONTO : Fun.surj]:
+  "surj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>y. \<exists>x. y = f x)"
+  by (auto simp add: fun_eq_iff)
+
+lemma hl_ax3: "\<exists>f\<Colon>ind \<Rightarrow> ind. inj f \<and> \<not> surj f"
+  by (rule_tac x="Suc_Rep" in exI)
+     (metis Suc_Rep_inject' injI Suc_Rep_not_Zero_Rep surjD)
+
+import_type_map num : Nat.nat
+import_const_map "_0" : Groups.zero_class.zero
+import_const_map SUC : Nat.Suc
+
+lemma NOT_SUC: "\<forall>n. Suc n \<noteq> 0"
+  by simp
+
+lemma SUC_INJ: "\<forall>m n. (Suc m = Suc n) = (m = n)"
+  by simp
+
+lemma num_INDUCTION:
+  "\<forall>P. P 0 \<and> (\<forall>n. P n \<longrightarrow> P (Suc n)) \<longrightarrow> (\<forall>n. P n)"
+  by (auto intro: nat.induct)
+
+lemma [import_const NUMERAL]: "id = (\<lambda>x :: nat. x)"
+  by auto
+
+definition [simp]: "bit0 n = 2 * n"
+
+lemma [import_const BIT0]:
+  "bit0 = (SOME fn. fn (id 0) = id 0 \<and> (\<forall>n. fn (Suc n) = Suc (Suc (fn n))))"
+  apply (auto intro!: some_equality[symmetric])
+  apply (auto simp add: fun_eq_iff)
+  apply (induct_tac x)
+  apply auto
+  done
+
+definition [import_const BIT1, simp]:
+  "bit1 = (\<lambda>x. Suc (bit0 x))"
+
+definition [simp]: "pred n = n - 1"
+
+lemma PRE[import_const PRE : HOL_Light_Maps.pred]:
+  "pred (id (0\<Colon>nat)) = id (0\<Colon>nat) \<and> (\<forall>n\<Colon>nat. pred (Suc n) = n)"
+  by simp
+
+lemma ADD[import_const "+" : Groups.plus_class.plus]:
+  "(\<forall>n :: nat. (id 0) + n = n) \<and> (\<forall>m n. (Suc m) + n = Suc (m + n))"
+  by simp
+
+lemma MULT[import_const "*" : Groups.times_class.times]:
+  "(\<forall>n :: nat. (id 0) * n = id 0) \<and> (\<forall>m n. (Suc m) * n = (m * n) + n)"
+  by simp
+
+lemma EXP[import_const EXP : Power.power_class.power]:
+  "(\<forall>m. m ^ (id 0) = id (bit1 0)) \<and> (\<forall>(m :: nat) n. m ^ (Suc n) = m * (m ^ n))"
+  by simp
+
+lemma LE[import_const "<=" : Orderings.ord_class.less_eq]:
+  "(\<forall>m :: nat. m \<le> (id 0) = (m = id 0)) \<and> (\<forall>m n. m \<le> (Suc n) = (m = Suc n \<or> m \<le> n))"
+  by auto
+
+lemma LT[import_const "<" : Orderings.ord_class.less]:
+  "(\<forall>m :: nat. m < (id 0) = False) \<and> (\<forall>m n. m < (Suc n) = (m = n \<or> m < n))"
+  by auto
+
+lemma DEF_GE[import_const ">=" : "Orderings.ord_class.greater_eq"]:
+  "(op \<ge>) = (\<lambda>x y :: nat. y \<le> x)"
+  by simp
+
+lemma DEF_GT[import_const ">" : "Orderings.ord_class.greater"]:
+  "(op >) = (\<lambda>x y :: nat. y < x)"
+  by simp
+
+lemma DEF_MAX[import_const "MAX"]:
+  "max = (\<lambda>x y :: nat. if x \<le> y then y else x)"
+  by (auto simp add: min_max.le_iff_sup fun_eq_iff)
+
+lemma DEF_MIN[import_const "MIN"]:
+  "min = (\<lambda>x y :: nat. if x \<le> y then x else y)"
+  by (auto simp add: min_max.le_iff_inf fun_eq_iff)
+
+lemma EVEN[import_const "EVEN" : "Parity.even_odd_class.even"]:
+  "even (id 0\<Colon>nat) = True \<and> (\<forall>n. even (Suc n) = (\<not> even n))"
+  by simp
+
+lemma SUB[import_const "-" : "Groups.minus_class.minus"]:
+  "(\<forall>m\<Colon>nat. m - (id 0) = m) \<and> (\<forall>m n. m - (Suc n) = pred (m - n))"
+  by simp
+
+lemma FACT[import_const "FACT" : "Fact.fact_class.fact"]:
+  "fact (id 0) = id (bit1 0) \<and> (\<forall>n. fact (Suc n) = Suc n * fact n)"
+  by simp
+
+import_const_map MOD : Divides.div_class.mod
+import_const_map DIV : Divides.div_class.div
+
+lemma DIVISION_0:
+  "\<forall>m n\<Colon>nat. if n = id 0 then m div n = id 0 \<and> m mod n = m else m = m div n * n + m mod n \<and> m mod n < n"
+  by simp
+
+lemmas [import_type sum "_dest_sum" "_mk_sum"] = type_definition_sum
+import_const_map INL : Sum_Type.Inl
+import_const_map INR : Sum_Type.Inr
+
+lemma sum_INDUCT:
+  "\<forall>P. (\<forall>a. P (Inl a)) \<and> (\<forall>a. P (Inr a)) \<longrightarrow> (\<forall>x. P x)"
+  by (auto intro: sum.induct)
+
+lemma sum_RECURSION:
+  "\<forall>Inl' Inr'. \<exists>fn. (\<forall>a. fn (Inl a) = Inl' a) \<and> (\<forall>a. fn (Inr a) = Inr' a)"
+  by (intro allI, rule_tac x="sum_case Inl' Inr'" in exI) auto
+
+lemma OUTL[import_const "OUTL" : "Sum_Type.Projl"]:
+  "Sum_Type.Projl (Inl x) = x"
+  by simp
+
+lemma OUTR[import_const "OUTR" : "Sum_Type.Projr"]:
+  "Sum_Type.Projr (Inr y) = y"
+  by simp
+
+import_type_map list : List.list
+import_const_map NIL : List.list.Nil
+import_const_map CONS : List.list.Cons
+
+lemma list_INDUCT:
+  "\<forall>P\<Colon>'A list \<Rightarrow> bool. P [] \<and> (\<forall>a0 a1. P a1 \<longrightarrow> P (a0 # a1)) \<longrightarrow> (\<forall>x. P x)"
+  using list.induct by auto
+
+lemma list_RECURSION:
+ "\<forall>nil' cons'. \<exists>fn\<Colon>'A list \<Rightarrow> 'Z. fn [] = nil' \<and> (\<forall>(a0\<Colon>'A) a1\<Colon>'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))"
+  by (intro allI, rule_tac x="list_rec nil' cons'" in exI) auto
+
+lemma HD[import_const HD : List.hd]:
+  "hd ((h\<Colon>'A) # t) = h"
+  by simp
+
+lemma TL[import_const TL : List.tl]:
+  "tl ((h\<Colon>'A) # t) = t"
+  by simp
+
+lemma APPEND[import_const APPEND : List.append]:
+  "(\<forall>l\<Colon>'A list. [] @ l = l) \<and> (\<forall>(h\<Colon>'A) t l. (h # t) @ l = h # t @ l)"
+  by simp
+
+lemma REVERSE[import_const REVERSE : List.rev]:
+  "rev [] = ([] :: 'A list) \<and> rev ((x\<Colon>'A) # l) = rev l @ [x]"
+  by simp
+
+lemma LENGTH[import_const LENGTH : List.length]:
+  "length [] = id 0 \<and> (\<forall>(h\<Colon>'A) t. length (h # t) = Suc (length t))"
+  by simp
+
+lemma MAP[import_const MAP : List.map]:
+  "(\<forall>f\<Colon>'A \<Rightarrow> 'B. map f [] = []) \<and>
+       (\<forall>(f\<Colon>'A \<Rightarrow> 'B) h t. map f (h # t) = f h # map f t)"
+  by simp
+
+lemma LAST[import_const LAST : List.last]:
+  "last ((h\<Colon>'A) # t) = (if t = [] then h else last t)"
+  by simp
+
+lemma BUTLAST[import_const BUTLAST : List.butlast]:
+    "butlast [] = ([] :: 't18337 list) \<and>
+     butlast ((h :: 't18337) # t) = (if t = [] then [] else h # butlast t)"
+  by simp
+
+lemma REPLICATE[import_const REPLICATE : List.replicate]:
+  "replicate (id (0\<Colon>nat)) (x\<Colon>'t18358) = [] \<and>
+   replicate (Suc n) x = x # replicate n x"
+  by simp
+
+lemma NULL[import_const NULL : List.null]:
+  "List.null ([]\<Colon>'t18373 list) = True \<and> List.null ((h\<Colon>'t18373) # t) = False"
+  unfolding null_def by simp
+
+lemma ALL[import_const ALL : List.list_all]:
+  "list_all (P\<Colon>'t18393 \<Rightarrow> bool) [] = True \<and>
+  list_all P (h # t) = (P h \<and> list_all P t)"
+  by simp
+
+lemma EX[import_const EX : List.list_ex]:
+  "list_ex (P\<Colon>'t18414 \<Rightarrow> bool) [] = False \<and>
+  list_ex P (h # t) = (P h \<or> list_ex P t)"
+  by simp
+
+lemma ITLIST[import_const ITLIST : List.foldr]:
+  "foldr (f\<Colon>'t18437 \<Rightarrow> 't18436 \<Rightarrow> 't18436) [] b = b \<and>
+  foldr f (h # t) b = f h (foldr f t b)"
+  by simp
+
+lemma ALL2_DEF[import_const ALL2 : List.list_all2]:
+  "list_all2 (P\<Colon>'t18495 \<Rightarrow> 't18502 \<Rightarrow> bool) [] (l2\<Colon>'t18502 list) = (l2 = []) \<and>
+  list_all2 P ((h1\<Colon>'t18495) # (t1\<Colon>'t18495 list)) l2 =
+  (if l2 = [] then False else P h1 (hd l2) \<and> list_all2 P t1 (tl l2))"
+  by simp (induct_tac l2, simp_all)
+
+lemma FILTER[import_const FILTER : List.filter]:
+  "filter (P\<Colon>'t18680 \<Rightarrow> bool) [] = [] \<and>
+  filter P ((h\<Colon>'t18680) # t) = (if P h then h # filter P t else filter P t)"
+  by simp
+
+lemma ZIP[import_const ZIP : List.zip]:
+ "zip [] [] = ([] :: ('t18824 \<times> 't18825) list) \<and>
+  zip ((h1\<Colon>'t18849) # t1) ((h2\<Colon>'t18850) # t2) = (h1, h2) # zip t1 t2"
+  by simp
+
+lemma WF[import_const WF : Wellfounded.wfP]:
+  "wfP u \<longleftrightarrow> (\<forall>P. (\<exists>x :: 'A. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))"
+proof (intro allI iffI impI wfI_min[to_pred], elim exE wfE_min[to_pred])
+  fix x :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and xa :: "'a" and Q
+  assume a: "xa \<in> Q"
+  assume "\<forall>P. Ex P \<longrightarrow> (\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y))"
+  then have "Ex (\<lambda>x. x \<in> Q) \<longrightarrow> (\<exists>xa. (\<lambda>x. x \<in> Q) xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> (\<lambda>x. x \<in> Q) y))" by auto
+  then show "\<exists>z\<in>Q. \<forall>y. x y z \<longrightarrow> y \<notin> Q" using a by auto
+next
+  fix x P and xa :: 'A and z
+  assume "P xa" "z \<in> {a. P a}" "\<And>y. x y z \<Longrightarrow> y \<notin> {a. P a}"
+  then show "\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y)" by auto
+qed auto
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/Import_Setup.thy	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,33 @@
+(*  Title:      HOL/Import/Import_Setup.thy
+    Author:     Cezary Kaliszyk, University of Innsbruck
+    Author:     Alexander Krauss, QAware GmbH
+*)
+
+header {* Importer machinery and required theorems *}
+
+theory Import_Setup
+imports "~~/src/HOL/Parity" "~~/src/HOL/Fact"
+keywords
+    "import_type_map" :: thy_decl and "import_const_map" :: thy_decl and
+    "import_file" :: thy_decl
+uses "import_data.ML" ("import_rule.ML")
+begin
+
+lemma light_ex_imp_nonempty:
+  "P t \<Longrightarrow> \<exists>x. x \<in> Collect P"
+  by auto
+
+lemma typedef_hol2hollight:
+  assumes a: "type_definition Rep Abs (Collect P)"
+  shows "Abs (Rep a) = a \<and> P r = (Rep (Abs r) = r)"
+  by (metis type_definition.Rep_inverse type_definition.Abs_inverse
+      type_definition.Rep a mem_Collect_eq)
+
+lemma ext2:
+  "(\<And>x. f x = g x) \<Longrightarrow> f = g"
+  by auto
+
+use "import_rule.ML"
+
+end
+
--- a/src/HOL/Import/Importer.thy	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,236 +0,0 @@
-(*  Title:      HOL/Import/Importer.thy
-    Author:     Sebastian Skalberg, TU Muenchen
-*)
-
-theory Importer
-imports Main
-keywords
-  "import_segment" "import_theory" "end_import" "setup_theory" "end_setup"
-  "setup_dump" "append_dump" "flush_dump" "ignore_thms" "type_maps"
-  "def_maps" "thm_maps" "const_renames" "const_moves" "const_maps" :: thy_decl and ">"
-uses "shuffler.ML" "import_rews.ML" ("proof_kernel.ML") ("replay.ML") ("import.ML")
-begin
-
-setup {* Shuffler.setup #> importer_setup *}
-
-parse_ast_translation smarter_trueprop_parsing
-
-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
-
-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"
-
-defs
-  ONE_ONE_DEF: "ONE_ONE f == ALL x y. f x = f y --> x = y"
-
-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 & ~(surj f))"
-proof (rule exI,safe)
-  show "inj Suc_Rep"
-    by (rule injI) (rule Suc_Rep_inject)
-next
-  assume "surj Suc_Rep"
-  hence "ALL y. EX x. y = Suc_Rep x"
-    by (simp add: 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 b: "Rep x = Rep y"
-    from td have "x = Abs (Rep x)"
-      by auto
-    also have "Abs (Rep x) = Abs (Rep y)"
-      by (simp add: b)
-    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 "proof_kernel.ML"
-use "replay.ML"
-use "import.ML"
-
-setup Import.setup
-
-end
-
--- a/src/HOL/Import/README	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-
-ATTENTION!  This is largely outdated.  The hint to PROOF_DIRS might be
-everything which is still relevant.
-
-All the files in this directory (except this README, Importer.thy, and
-ROOT.ML) are automatically generated.  Edit the files in
-../Generate-HOL and run "isabelle make HOL-Complex-Generate-HOL" in
-~~/src/HOL, if something needs to be changed.
-
-To build the logic in this directory, simply do a "isabelle make
-HOL-Import-HOL" in ~~/src/HOL.
-
-Note that the quick_and_dirty flag is on as default for this
-directory, which means that the original external proofs are not consulted
-at all.  If a real replay of the external proofs is desired, get and
-unpack the external proof objects to somewhere on your harddisk, and set
-the variable PROOF_DIRS to the directory where the directory "hol4"
-is.  Now edit the ROOT.ML file to unset the quick_and_dirty flag and
-do "isabelle make HOL-Import-HOL" in ~~/src/HOL.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/ROOT.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,5 @@
+use_thy "HOL_Light_Maps";
+
+if getenv "HOL_LIGHT_DUMP" <> ""
+then use_thy "HOL_Light_Import"
+else ()
--- a/src/HOL/Import/import.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,280 +0,0 @@
-(*  Title:      HOL/Import/import.ML
-    Author:     Sebastian Skalberg (TU Muenchen)
-*)
-
-signature IMPORT =
-sig
-    val debug      : bool Unsynchronized.ref
-    val import_tac : Proof.context -> string * string -> tactic
-    val setup      : theory -> theory
-end
-
-structure ImportData = Theory_Data
-(
-  type T = ProofKernel.thm option array option  (* FIXME mutable array !??!!! *)
-  val empty = NONE
-  val extend = I
-  fun merge _ = NONE
-)
-
-structure Import : IMPORT =
-struct
-
-val debug = Unsynchronized.ref false
-fun message s = if !debug then writeln s else ()
-
-fun import_tac ctxt (thyname, thmname) =
-    if ! quick_and_dirty
-    then Skip_Proof.cheat_tac (Proof_Context.theory_of ctxt)
-    else
-     fn th =>
-        let
-            val thy = Proof_Context.theory_of ctxt
-            val prem = hd (prems_of th)
-            val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
-            val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
-            val int_thms = case ImportData.get thy of
-                               NONE => fst (Replay.setup_int_thms thyname thy)
-                             | SOME a => a
-            val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
-            val importerthm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
-            val thm = snd (ProofKernel.to_isa_thm importerthm)
-            val rew = ProofKernel.rewrite_importer_term (concl_of thm) thy
-            val thm = Thm.equal_elim rew thm
-            val prew = ProofKernel.rewrite_importer_term prem thy
-            val prem' = #2 (Logic.dest_equals (prop_of prew))
-            val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm)
-            val thm = ProofKernel.disambiguate_frees thm
-            val _ = message ("Disambiguate: " ^ Display.string_of_thm ctxt 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' = Thm.equal_elim (Thm.symmetric prew) thm
-                    val res = Thm.bicompose true (false,thm',0) 1 th
-                in
-                    res
-                end
-              | NONE => (message "import: set_prop didn't succeed"; no_tac th)
-        end
-
-val setup = Method.setup @{binding import}
-  (Scan.lift (Args.name -- Args.name) >>
-    (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
-  "import external theorem"
-
-(* Parsers *)
-
-val import_segment = Parse.name >> set_import_segment
-
-
-val import_theory = (Parse.name -- Parse.name) >> (fn (location, thyname) =>
-                               fn thy =>
-                                  thy |> set_generating_thy thyname
-                                      |> Sign.add_path thyname
-                                      |> add_dump ("setup_theory " ^ quote location ^ " " ^ thyname))
-
-fun end_import toks =
-    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_out (should_ignore thyname thy) facts
-                fun replay thy = Replay.import_thms thyname int_thms thmnames thy
-            in
-                thy |> clear_import_thy
-                    |> set_segment thyname segname
-                    |> set_used_names facts
-                    |> replay 
-                    |> clear_used_names
-                    |> export_importer_pending
-                    |> Sign.parent_path
-                    |> dump_import_thy thyname
-                    |> add_dump ";end_setup"
-            end) toks
-
-val ignore_thms = Scan.repeat1 Parse.name
-                               >> (fn ignored =>
-                                   fn thy =>
-                                      let
-                                          val thyname = get_import_thy thy
-                                      in
-                                          Library.foldl (fn (thy,thmname) => ignore_importer thyname thmname thy) (thy,ignored)
-                                      end)
-
-val thm_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
-                            >> (fn thmmaps =>
-                                fn thy =>
-                                   let
-                                       val thyname = get_import_thy thy
-                                   in
-                                       Library.foldl (fn (thy,(hol,isa)) => add_importer_mapping thyname hol isa thy) (thy,thmmaps)
-                                   end)
-
-val type_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
-                             >> (fn typmaps =>
-                                 fn thy =>
-                                    let
-                                        val thyname = get_import_thy thy
-                                    in
-                                        Library.foldl (fn (thy,(hol,isa)) => add_importer_type_mapping thyname hol false isa thy) (thy,typmaps)
-                                    end)
-
-val def_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
-                            >> (fn defmaps =>
-                                fn thy =>
-                                   let
-                                       val thyname = get_import_thy thy
-                                   in
-                                       Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
-                                   end)
-
-val const_renames = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.name)
-                                 >> (fn renames =>
-                                     fn thy =>
-                                        let
-                                            val thyname = get_import_thy thy
-                                        in
-                                            Library.foldl (fn (thy,(hol,isa)) => add_importer_const_renaming thyname hol isa thy) (thy,renames)
-                                        end)
-                                                                                                                                      
-val const_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
-                              >> (fn constmaps =>
-                                  fn thy =>
-                                     let
-                                         val thyname = get_import_thy thy
-                                     in
-                                         Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol false isa thy
-                                                 | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
-                                     end)
-
-val const_moves = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
-                               >> (fn constmaps =>
-                                   fn thy =>
-                                      let
-                                          val thyname = get_import_thy thy
-                                      in
-                                          Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol true isa thy
-                                                  | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol true isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
-                                      end)
-
-fun import_thy location s =
-    let
-        val src_location = Path.append (Path.explode location) (Path.basic (s ^ ".imp"))
-        val is = TextIO.openIn (File.platform_path src_location)
-        val inp = TextIO.inputAll is
-        val _ = TextIO.closeIn is
-        val orig_source = Source.of_string inp
-        val symb_source = Symbol.source orig_source
-        val lexes =
-          (Scan.make_lexicon
-            (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
-                  Scan.empty_lexicon)
-        val token_source = Token.source {do_recover = NONE} (K lexes) Position.start symb_source
-        val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
-        val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
-        val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
-        val def_mapsP = Parse.$$$ "def_maps" |-- def_maps
-        val const_mapsP = Parse.$$$ "const_maps" |-- const_maps
-        val const_movesP = Parse.$$$ "const_moves" |-- const_moves
-        val const_renamesP = Parse.$$$ "const_renames" |-- const_renames
-        val ignore_thmsP = Parse.$$$ "ignore_thms" |-- ignore_thms
-        val thm_mapsP = Parse.$$$ "thm_maps" |-- thm_maps
-        val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
-        val importP = Parse.$$$ "import" |-- Scan.repeat parser --| Parse.$$$ "end"
-        fun apply [] thy = thy
-          | apply (f::fs) thy = apply fs (f thy)
-    in
-        apply (set_replaying_thy s :: Sign.add_path s :: fst (importP token_list))
-    end
-
-fun create_int_thms thyname thy =
-    if ! quick_and_dirty
-    then thy
-    else
-        case ImportData.get thy of
-            NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
-          | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
-
-fun clear_int_thms thy =
-    if ! quick_and_dirty
-    then thy
-    else
-        case ImportData.get thy of
-            NONE => error "Import data already cleared - forgotten a setup_theory?"
-          | SOME _ => ImportData.put NONE thy
-
-val setup_theory = (Parse.name -- Parse.name)
-                       >>
-                       (fn (location, thyname) =>
-                        fn thy =>
-                           (case Importer_DefThy.get thy of
-                                NoImport => thy
-                              | Generating _ => error "Currently generating a theory!"
-                              | Replaying _ => thy |> clear_import_thy)
-                               |> import_thy location thyname
-                               |> create_int_thms thyname)
-
-fun end_setup toks =
-    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
-                    |> clear_int_thms
-                    |> Sign.parent_path
-            end) toks
-
-val set_dump  = Parse.string -- Parse.string   >> setup_dump
-fun fl_dump toks  = Scan.succeed flush_dump toks
-val append_dump = (Parse.verbatim || Parse.string) >> add_dump
-
-val _ =
-  (Outer_Syntax.command @{command_spec "import_segment"} "set import segment name"
-     (import_segment >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "import_theory"} "set default external theory name"
-     (import_theory >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "end_import"} "end external import"
-     (end_import >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "setup_theory"} "setup external theory replaying"
-     (setup_theory >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "end_setup"} "end external setup"
-     (end_setup >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "setup_dump"} "setup the dump file name"
-     (set_dump >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "append_dump"} "add text to dump file"
-     (append_dump >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "flush_dump"} "write the dump to file"
-     (fl_dump >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "ignore_thms"}
-     "theorems to ignore in next external theory import"
-     (ignore_thms >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "type_maps"}
-     "map external type names to existing Isabelle/HOL type names"
-     (type_maps >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "def_maps"}
-     "map external constant names to their primitive definitions"
-     (def_maps >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "thm_maps"}
-     "map external theorem names to existing Isabelle/HOL theorem names"
-     (thm_maps >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "const_renames"}
-     "rename external const names"
-     (const_renames >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "const_moves"}
-     "rename external const names to other external constants"
-     (const_moves >> Toplevel.theory);
-   Outer_Syntax.command @{command_spec "const_maps"}
-     "map external const names to existing Isabelle/HOL const names"
-     (const_maps >> Toplevel.theory));
-
-end
-
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import_data.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,125 @@
+(*  Title:      HOL/Import/import_data.ML
+    Author:     Cezary Kaliszyk, University of Innsbruck
+    Author:     Alexander Krauss, QAware GmbH
+
+Importer data.
+*)
+
+signature IMPORT_DATA =
+sig
+  val get_const_map : string -> theory -> string option
+  val get_typ_map : string -> theory -> string option
+  val get_const_def : string -> theory -> thm option
+  val get_typ_def : string -> theory -> thm option
+
+  val add_const_map : string -> string -> theory -> theory
+  val add_typ_map : string -> string -> theory -> theory
+  val add_const_def : string -> thm -> string option -> theory -> theory
+  val add_typ_def : string -> string -> string -> thm -> theory -> theory
+end
+
+structure Import_Data: IMPORT_DATA =
+struct
+
+structure Data = Theory_Data
+(
+  type T = {const_map: string Symtab.table, ty_map: string Symtab.table,
+            const_def: thm Symtab.table, ty_def: thm Symtab.table}
+  val empty = {const_map = Symtab.empty, ty_map = Symtab.empty,
+               const_def = Symtab.empty, ty_def = Symtab.empty}
+  val extend = I
+  fun merge
+   ({const_map = cm1, ty_map = tm1, const_def = cd1, ty_def = td1},
+    {const_map = cm2, ty_map = tm2, const_def = cd2, ty_def = td2}) : T =
+    {const_map = Symtab.merge (K true) (cm1, cm2), ty_map = Symtab.merge (K true) (tm1, tm2),
+     const_def = Symtab.merge (K true) (cd1, cd2), ty_def = Symtab.merge (K true) (td1, td2)
+    }
+)
+
+fun get_const_map s thy = Symtab.lookup (#const_map (Data.get thy)) s
+
+fun get_typ_map s thy = Symtab.lookup (#ty_map (Data.get thy)) s
+
+fun get_const_def s thy = Symtab.lookup (#const_def (Data.get thy)) s
+
+fun get_typ_def s thy = Symtab.lookup (#ty_def (Data.get thy)) s
+
+fun add_const_map s1 s2 thy =
+  Data.map (fn {const_map, ty_map, const_def, ty_def} =>
+    {const_map = (Symtab.update (s1, s2) const_map), ty_map = ty_map,
+     const_def = const_def, ty_def = ty_def}) thy
+
+fun add_typ_map s1 s2 thy =
+  Data.map (fn {const_map, ty_map, const_def, ty_def} =>
+    {const_map = const_map, ty_map = (Symtab.update (s1, s2) ty_map),
+     const_def = const_def, ty_def = ty_def}) thy
+
+fun add_const_def s th name_opt thy =
+  let
+    val th = Thm.legacy_freezeT th
+    val name = case name_opt of
+         NONE => (fst o dest_Const o fst o HOLogic.dest_eq o
+           HOLogic.dest_Trueprop o prop_of) th
+       | SOME n => n
+    val thy' = add_const_map s name thy
+  in
+    Data.map (fn {const_map, ty_map, const_def, ty_def} =>
+      {const_map = const_map, ty_map = ty_map,
+       const_def = (Symtab.update (s, th) const_def), ty_def = ty_def}) thy'
+  end
+
+fun add_typ_def tyname absname repname th thy =
+  let
+    val th = Thm.legacy_freezeT th
+    val (l, _) = dest_comb (HOLogic.dest_Trueprop (prop_of th))
+    val (l, abst) = dest_comb l
+    val (_, rept) = dest_comb l
+    val (absn, _) = dest_Const abst
+    val (repn, _) = dest_Const rept
+    val nty = domain_type (fastype_of rept)
+    val ntyn = fst (dest_Type nty)
+    val thy2 = add_typ_map tyname ntyn thy
+    val thy3 = add_const_map absname absn thy2
+    val thy4 = add_const_map repname repn thy3
+  in
+    Data.map (fn {const_map, ty_map, const_def, ty_def} =>
+      {const_map = const_map, ty_map = ty_map,
+       const_def = const_def, ty_def = (Symtab.update (tyname, th) ty_def)}) thy4
+  end
+
+val parser = Parse.name -- Scan.option (Parse.$$$ ":" |-- Parse.xname)
+fun add_const_attr (s1, s2) = Thm.declaration_attribute
+  (fn th => Context.mapping (add_const_def s1 th s2) (fn x => x))
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup (Binding.name "import_const") (Scan.lift parser >> add_const_attr)
+  "Declare a theorem as an equality that maps the given constant"))
+
+val parser = Parse.name -- (Parse.name -- Parse.name)
+fun add_typ_attr (tyname, (absname, repname)) = Thm.declaration_attribute
+  (fn th => Context.mapping (add_typ_def tyname absname repname th) (fn x => x))
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup (Binding.name "import_type") (Scan.lift parser >> add_typ_attr)
+  "Declare a type_definion theorem as a map for an imported type abs rep"))
+
+val type_map = (Parse.name --| Parse.$$$ ":" -- Parse.xname) >>
+  (fn (ty_name1, ty_name2) => add_typ_map ty_name1 ty_name2)
+val _ = Outer_Syntax.command @{command_spec "import_type_map"}
+  "Map external type name to existing Isabelle/HOL type name"
+  (type_map >> Toplevel.theory)
+
+val const_map = (Parse.name --| Parse.$$$ ":" -- Parse.xname) >>
+  (fn (cname1, cname2) => add_const_map cname1 cname2)
+val _ = Outer_Syntax.command @{command_spec "import_const_map"}
+  "Map external const name to existing Isabelle/HOL const name"
+  (const_map >> Toplevel.theory)
+
+(* Initial type and constant maps, for types and constants that are not
+   defined, which means their definitions do not appear in the proof dump *)
+val _ = Context.>> (Context.map_theory (
+  add_typ_map "bool" @{type_name bool} o
+  add_typ_map "fun" @{type_name fun} o
+  add_typ_map "ind" @{type_name ind} o
+  add_const_map "=" @{const_name HOL.eq} o
+  add_const_map "@" @{const_name "Eps"}))
+
+end
--- a/src/HOL/Import/import_rews.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,633 +0,0 @@
-(*  Title:      HOL/Import/import_rews.ML
-    Author:     Sebastian Skalberg (TU Muenchen)
-*)
-
-structure StringPair = Table(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 Importer_DefThy = Theory_Data
-(
-  type T = ImportStatus
-  val empty = NoImport
-  val extend = I
-  fun merge (NoImport, NoImport) = NoImport
-    | merge _ = (warning "Import status set during merge"; NoImport)
-)
-
-structure ImportSegment = Theory_Data
-(
-  type T = string
-  val empty = ""
-  val extend = 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"
-)
-
-val get_import_segment = ImportSegment.get
-val set_import_segment = ImportSegment.put
-
-structure Importer_UNames = Theory_Data
-(
-  type T = string list
-  val empty = []
-  val extend = I
-  fun merge ([], []) = []
-    | merge _ = error "Used names not empty during merge"
-)
-
-structure Importer_Dump = Theory_Data
-(
-  type T = string * string * string list
-  val empty = ("","",[])
-  val extend = I
-  fun merge (("","",[]),("","",[])) = ("","",[])
-    | merge _ = error "Dump data not empty during merge"
-)
-
-structure Importer_Moves = Theory_Data
-(
-  type T = string Symtab.table
-  val empty = Symtab.empty
-  val extend = I
-  fun merge data = Symtab.merge (K true) data
-)
-
-structure Importer_Imports = Theory_Data
-(
-  type T = string Symtab.table
-  val empty = Symtab.empty
-  val extend = I
-  fun merge data = Symtab.merge (K true) data
-)
-
-fun get_segment2 thyname thy =
-  Symtab.lookup (Importer_Imports.get thy) thyname
-
-fun set_segment thyname segname thy =
-    let
-        val imps = Importer_Imports.get thy
-        val imps' = Symtab.update_new (thyname,segname) imps
-    in
-        Importer_Imports.put imps' thy
-    end
-
-structure Importer_CMoves = Theory_Data
-(
-  type T = string Symtab.table
-  val empty = Symtab.empty
-  val extend = I
-  fun merge data = Symtab.merge (K true) data
-)
-
-structure Importer_Maps = Theory_Data
-(
-  type T = string option StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_Thms = Theory_Data
-(
-  type T = holthm StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_ConstMaps = Theory_Data
-(
-  type T = (bool * string * typ option) StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_Rename = Theory_Data
-(
-  type T = string StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_DefMaps = Theory_Data
-(
-  type T = string StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_TypeMaps = Theory_Data
-(
-  type T = (bool * string) StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_Pending = Theory_Data
-(
-  type T = ((term * term) list * thm) StringPair.table
-  val empty = StringPair.empty
-  val extend = I
-  fun merge data = StringPair.merge (K true) data
-)
-
-structure Importer_Rewrites = Theory_Data
-(
-  type T = thm list
-  val empty = []
-  val extend = I
-  val merge = Thm.merge_thms
-)
-
-val importer_debug = Unsynchronized.ref false
-fun message s = if !importer_debug then writeln s else ()
-
-fun add_importer_rewrite (Context.Theory thy, th) =
-    let
-        val _ = message "Adding external rewrite"
-        val th1 = th RS @{thm eq_reflection}
-                  handle THM _ => th
-        val current_rews = Importer_Rewrites.get thy
-        val new_rews = insert Thm.eq_thm th1 current_rews
-        val updated_thy  = Importer_Rewrites.put new_rews thy
-    in
-        (Context.Theory updated_thy,th1)
-    end
-  | add_importer_rewrite (context, th) = (context,
-      (th RS @{thm eq_reflection} handle THM _ => th)
-    );
-
-fun ignore_importer bthy bthm thy =
-    let
-        val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
-        val curmaps = Importer_Maps.get thy
-        val newmaps = StringPair.update_new ((bthy,bthm),NONE) curmaps
-        val upd_thy = Importer_Maps.put newmaps thy
-    in
-        upd_thy
-    end
-
-val opt_get_output_thy = #2 o Importer_Dump.get
-
-fun get_output_thy thy =
-    case #2 (Importer_Dump.get thy) of
-        "" => error "No theory file being output"
-      | thyname => thyname
-
-val get_output_dir = #1 o Importer_Dump.get
-
-fun add_importer_move bef aft thy =
-    let
-        val curmoves = Importer_Moves.get thy
-        val newmoves = Symtab.update_new (bef, aft) curmoves
-    in
-        Importer_Moves.put newmoves thy
-    end
-
-fun get_importer_move bef thy =
-  Symtab.lookup (Importer_Moves.get thy) bef
-
-fun follow_name thmname thy =
-    let
-        val moves = Importer_Moves.get thy
-        fun F thmname =
-            case Symtab.lookup moves thmname of
-                SOME name => F name
-              | NONE => thmname
-    in
-        F thmname
-    end
-
-fun add_importer_cmove bef aft thy =
-    let
-        val curmoves = Importer_CMoves.get thy
-        val newmoves = Symtab.update_new (bef, aft) curmoves
-    in
-        Importer_CMoves.put newmoves thy
-    end
-
-fun get_importer_cmove bef thy =
-  Symtab.lookup (Importer_CMoves.get thy) bef
-
-fun follow_cname thmname thy =
-    let
-        val moves = Importer_CMoves.get thy
-        fun F thmname =
-            case Symtab.lookup moves thmname of
-                SOME name => F name
-              | NONE => thmname
-    in
-        F thmname
-    end
-
-fun add_importer_mapping bthy bthm isathm thy =
-    let 
-       (* val _ = writeln ("Before follow_name: "^isathm)  *)
-      val isathm = follow_name isathm thy
-       (* val _ = writeln ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)*)
-        val curmaps = Importer_Maps.get thy
-        val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps
-        val upd_thy = Importer_Maps.put newmaps thy
-    in
-        upd_thy
-    end;
-
-fun get_importer_type_mapping bthy tycon thy =
-    let
-        val maps = Importer_TypeMaps.get thy
-    in
-        StringPair.lookup maps (bthy,tycon)
-    end
-
-fun get_importer_mapping bthy bthm thy =
-    let
-        val curmaps = Importer_Maps.get thy
-    in
-        StringPair.lookup curmaps (bthy,bthm)
-    end;
-
-fun add_importer_const_mapping bthy bconst internal isaconst thy =
-    let
-        val thy = case opt_get_output_thy thy of
-                      "" => thy
-                    | output_thy => if internal
-                                    then add_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
-                                    else thy
-        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-        val curmaps = Importer_ConstMaps.get thy
-        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,NONE)) curmaps
-        val upd_thy = Importer_ConstMaps.put newmaps thy
-    in
-        upd_thy
-    end;
-
-fun add_importer_const_renaming bthy bconst newname thy =
-    let
-        val currens = Importer_Rename.get thy
-        val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
-        val newrens = StringPair.update_new ((bthy,bconst),newname) currens
-        val upd_thy = Importer_Rename.put newrens thy
-    in
-        upd_thy
-    end;
-
-fun get_importer_const_renaming bthy bconst thy =
-    let
-        val currens = Importer_Rename.get thy
-    in
-        StringPair.lookup currens (bthy,bconst)
-    end;
-
-fun get_importer_const_mapping bthy bconst thy =
-    let
-        val bconst = case get_importer_const_renaming bthy bconst thy of
-                         SOME name => name
-                       | NONE => bconst
-        val maps = Importer_ConstMaps.get thy
-    in
-        StringPair.lookup maps (bthy,bconst)
-    end
-
-fun add_importer_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_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
-                                    else thy
-        val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-        val curmaps = Importer_ConstMaps.get thy
-        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,SOME typ)) curmaps
-        val upd_thy = Importer_ConstMaps.put newmaps thy
-    in
-        upd_thy
-    end;
-
-fun add_importer_type_mapping bthy bconst internal isaconst thy =
-    let
-        val curmaps = Importer_TypeMaps.get thy
-        val _ = writeln ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
-        val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst)) curmaps
-               (* FIXME avoid handle x *)
-               handle x => let val (_, isaconst') = the (StringPair.lookup curmaps (bthy, bconst)) in
-                      warning ("couldn't map type "^bthy^"."^bconst^" to "^isaconst^": already mapped to "^isaconst'); raise x end
-        val upd_thy = Importer_TypeMaps.put newmaps thy
-    in
-        upd_thy
-    end;
-
-fun add_importer_pending bthy bthm hth thy =
-    let
-        val thmname = Sign.full_bname thy bthm
-        val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
-        val curpend = Importer_Pending.get thy
-        val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
-        val upd_thy = Importer_Pending.put newpend thy
-        val thy' = case opt_get_output_thy upd_thy of
-                       "" => add_importer_mapping bthy bthm thmname upd_thy
-                     | output_thy =>
-                       let
-                           val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
-                       in
-                           upd_thy |> add_importer_move thmname new_thmname
-                                   |> add_importer_mapping bthy bthm new_thmname
-                       end
-    in
-        thy'
-    end;
-
-fun get_importer_theorem thyname thmname thy =
-  let
-    val isathms = Importer_Thms.get thy
-  in
-    StringPair.lookup isathms (thyname,thmname)
-  end;
-
-fun add_importer_theorem thyname thmname hth =
-  let
-    val _ = message ("Adding external theorem " ^ thyname ^ "." ^ thmname)
-  in
-    Importer_Thms.map (StringPair.update_new ((thyname, thmname), hth))
-  end;
-
-fun export_importer_pending thy =
-  let
-    val rews = Importer_Rewrites.get thy;
-    val pending = Importer_Pending.get thy;
-    fun process ((bthy,bthm), hth as (_,thm)) thy =
-      let
-        val thm1 = rewrite_rule (map (Thm.transfer thy) rews) (Thm.transfer thy thm);
-        val thm2 = Drule.export_without_context thm1;
-      in
-        thy
-        |> Global_Theory.store_thm (Binding.name bthm, thm2)
-        |> snd
-        |> add_importer_theorem bthy bthm hth
-      end;
-  in
-    thy
-    |> StringPair.fold process pending
-    |> Importer_Pending.put StringPair.empty
-  end;
-
-fun setup_dump (dir,thyname) thy =
-    Importer_Dump.put (dir,thyname,["(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)"]) thy
-
-fun add_dump str thy =
-    let
-        val (dir,thyname,curdump) = Importer_Dump.get thy
-    in
-        Importer_Dump.put (dir,thyname,str::curdump) thy
-    end
-
-fun flush_dump thy =
-    let
-        val (dir,thyname,dumpdata) = Importer_Dump.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
-        Importer_Dump.put ("","",[]) thy
-    end
-
-fun set_generating_thy thyname thy =
-    case Importer_DefThy.get thy of
-        NoImport => Importer_DefThy.put (Generating thyname) thy
-      | _ => error "Import already in progess"
-
-fun set_replaying_thy thyname thy =
-    case Importer_DefThy.get thy of
-        NoImport => Importer_DefThy.put (Replaying thyname) thy
-      | _ => error "Import already in progess"
-
-fun clear_import_thy thy =
-    case Importer_DefThy.get thy of
-        NoImport => error "No import in progress"
-      | _ => Importer_DefThy.put NoImport thy
-
-fun get_generating_thy thy =
-    case Importer_DefThy.get thy of
-        Generating thyname => thyname
-      | _ => error "No theory being generated"
-
-fun get_replaying_thy thy =
-    case Importer_DefThy.get thy of
-        Replaying thyname => thyname
-      | _ => error "No theory being replayed"
-
-fun get_import_thy thy =
-    case Importer_DefThy.get thy of
-        Replaying thyname => thyname
-      | Generating thyname => thyname
-      | _ => error "No theory being imported"
-
-fun should_ignore thyname thy thmname =
-    case get_importer_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 input_thy = Context.theory_name thy
-        val import_segment = get_import_segment 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.fold
-      (fn ((bthy, bthm), v) => fn (ign, map) =>
-        if bthy = thyname
-        then case v
-         of NONE => (bthm :: ign, map)
-          | SOME w => (ign, (bthm, w) :: map)
-        else (ign, map)) (Importer_Maps.get thy) ([],[]);
-    fun mk init = StringPair.fold
-      (fn ((bthy, bthm), v) => if bthy = thyname then cons (bthm, v) else I) init [];
-    val constmaps = mk (Importer_ConstMaps.get thy);
-    val constrenames = mk (Importer_Rename.get thy);
-    val typemaps = mk (Importer_TypeMaps.get thy);
-    val defmaps = mk (Importer_DefMaps.get thy);
-        fun new_name internal isa =
-            if internal
-            then
-                let
-                    val paths = Long_Name.explode isa
-                    val i = drop (length paths - 2) paths
-                in
-                    case i of
-                        [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
-                      | _ => error "import_rews.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,(_,isa,opt_ty)) =>
-                        (out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
-                         case opt_ty of
-                             SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy 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"
-
-        fun gen2replay in_thy out_thy s = 
-            let
-                val ss = Long_Name.explode s
-            in
-                if (hd ss = in_thy) then 
-                    Long_Name.implode (out_thy::(tl ss))
-                else
-                    s
-            end 
-
-        val _ = if null mapped
-                then ()
-                else out "thm_maps"
-        val _ = app (fn (hol,isa) => out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (gen2replay input_thy output_thy 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 = Importer_UNames.get thy
-    in
-        case unames of
-            [] => Importer_UNames.put names thy
-          | _ => error "import_rews.set_used_names called on initialized data!"
-    end
-
-val clear_used_names = Importer_UNames.put [];
-
-fun get_defmap thyname const thy =
-    let
-        val maps = Importer_DefMaps.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 = Importer_DefMaps.get thy
-        val maps' = StringPair.update_new ((thyname,const),defname) maps
-        val thy' = Importer_DefMaps.put maps' thy
-    in
-        thy'
-    end
-
-fun get_defname thyname name thy =
-    let
-        val maps = Importer_DefMaps.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 = Importer_UNames.get thy
-                val defname = Thm.def_name name
-                val pdefname = name ^ "_primdef"
-            in
-                if not (member (op =) used defname)
-                then F defname                 (* name_def *)
-                else if not (member (op =) used pdefname)
-                then F pdefname                (* name_primdef *)
-                else F (singleton (Name.variant_list used) pdefname) (* last resort *)
-            end
-    end
-
-local
-    fun handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "=="}, _],_,_]] = x
-      | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax all}, _],_]] = x
-      | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "==>"}, _],_,_]] = x
-      | handle_meta [x] = Ast.Appl [Ast.Constant @{const_syntax Trueprop}, x]
-      | handle_meta _ = error "import_rews error: Trueprop not applied to single argument"
-in
-val smarter_trueprop_parsing = [(@{const_syntax Trueprop},handle_meta)]
-end
-
-val importer_setup =
-  add_importer_type_mapping "min" "bool" false @{type_name bool}
-  #> add_importer_type_mapping "min" "fun" false "fun"
-  #> add_importer_type_mapping "min" "ind" false @{type_name ind}
-  #> add_importer_const_mapping "min" "=" false @{const_name HOL.eq}
-  #> add_importer_const_mapping "min" "==>" false @{const_name HOL.implies}
-  #> add_importer_const_mapping "min" "@" false @{const_name "Eps"}
-  #> Attrib.setup @{binding import_rew}
-    (Scan.succeed (Thm.mixed_attribute add_importer_rewrite)) "external rewrite rule";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import_rule.ML	Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,442 @@
+(*  Title:      HOL/Import/import_rule.ML
+    Author:     Cezary Kaliszyk, University of Innsbruck
+    Author:     Alexander Krauss, QAware GmbH
+
+Importer proof rules and processing of lines and files.
+
+Based on earlier code by Steven Obua and Sebastian Skalberg.
+*)
+
+signature IMPORT_RULE =
+sig
+  val beta : cterm -> thm
+  val eq_mp : thm -> thm -> thm
+  val comb : thm -> thm -> thm
+  val trans : thm -> thm -> thm
+  val deduct : thm -> thm -> thm
+  val conj1 : thm -> thm
+  val conj2 : thm -> thm
+  val refl : cterm -> thm
+  val abs : cterm -> thm -> thm
+  val mdef : string -> theory -> thm
+  val def : string -> cterm -> theory -> thm * theory
+  val mtydef : string -> theory -> thm
+  val tydef :
+    string -> string -> string -> cterm -> cterm -> thm -> theory -> thm * theory
+  val inst_type : (ctyp * ctyp) list -> thm -> theory -> thm
+  val inst : (cterm * cterm) list -> thm -> thm
+
+  type state
+  val init_state : state
+  val process_line : string -> (theory * state) -> (theory * state)
+  val process_file : Path.T -> theory -> theory
+end
+
+structure Import_Rule: IMPORT_RULE =
+struct
+
+val init_state = ((Inttab.empty, 0), (Inttab.empty, 0), (Inttab.empty, 0))
+
+type state = (ctyp Inttab.table * int) * (cterm Inttab.table * int) * (thm Inttab.table * int)
+
+fun implies_elim_all th = implies_elim_list th (map Thm.assume (cprems_of th))
+
+fun meta_mp th1 th2 =
+  let
+    val th1a = implies_elim_all th1
+    val th1b = Thm.implies_intr (strip_imp_concl (cprop_of th2)) th1a
+    val th2a = implies_elim_all th2
+    val th3 = Thm.implies_elim th1b th2a
+  in
+    implies_intr_hyps th3
+  end
+
+fun meta_eq_to_obj_eq th =
+  let
+    val (tml, tmr) = Thm.dest_binop (strip_imp_concl (cprop_of th))
+    val cty = ctyp_of_term tml
+    val i = Drule.instantiate' [SOME cty] [SOME tml, SOME tmr]
+      @{thm meta_eq_to_obj_eq}
+  in
+    Thm.implies_elim i th
+  end
+
+fun beta ct = meta_eq_to_obj_eq (Thm.beta_conversion false ct)
+
+fun eq_mp th1 th2 =
+  let
+    val (tm1l, tm1r) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th1)))
+    val i1 = Drule.instantiate' [] [SOME tm1l, SOME tm1r] @{thm iffD1}
+    val i2 = meta_mp i1 th1
+  in
+    meta_mp i2 th2
+  end
+
+fun comb th1 th2 =
+  let
+    val t1c = Thm.dest_arg (strip_imp_concl (cprop_of th1))
+    val t2c = Thm.dest_arg (strip_imp_concl (cprop_of th2))
+    val (cf, cg) = Thm.dest_binop t1c
+    val (cx, cy) = Thm.dest_binop t2c
+    val [fd, fr] = Thm.dest_ctyp (ctyp_of_term cf)
+    val i1 = Drule.instantiate' [SOME fd, SOME fr]
+      [SOME cf, SOME cg, SOME cx, SOME cy] @{thm cong}
+    val i2 = meta_mp i1 th1
+  in
+    meta_mp i2 th2
+  end
+
+fun trans th1 th2 =
+  let
+    val t1c = Thm.dest_arg (strip_imp_concl (cprop_of th1))
+    val t2c = Thm.dest_arg (strip_imp_concl (cprop_of th2))
+    val (r, s) = Thm.dest_binop t1c
+    val (_, t) = Thm.dest_binop t2c
+    val ty = ctyp_of_term r
+    val i1 = Drule.instantiate' [SOME ty] [SOME r, SOME s, SOME t] @{thm trans}
+    val i2 = meta_mp i1 th1
+  in
+    meta_mp i2 th2
+  end
+
+fun deduct th1 th2 =
+  let
+    val th1c = strip_imp_concl (cprop_of th1)
+    val th2c = strip_imp_concl (cprop_of th2)
+    val th1a = implies_elim_all th1
+    val th2a = implies_elim_all th2
+    val th1b = Thm.implies_intr th2c th1a
+    val th2b = Thm.implies_intr th1c th2a
+    val i = Drule.instantiate' []
+      [SOME (Thm.dest_arg th1c), SOME (Thm.dest_arg th2c)] @{thm iffI}
+    val i1 = Thm.implies_elim i (Thm.assume (cprop_of th2b))
+    val i2 = Thm.implies_elim i1 th1b
+    val i3 = Thm.implies_intr (cprop_of th2b) i2
+    val i4 = Thm.implies_elim i3 th2b
+  in
+    implies_intr_hyps i4
+  end
+
+fun conj1 th =
+  let
+    val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th)))
+    val i = Drule.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct1}
+  in
+    meta_mp i th
+  end
+
+fun conj2 th =
+  let
+    val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th)))
+    val i = Drule.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct2}
+  in
+    meta_mp i th
+  end
+
+fun refl ctm =
+  let
+    val cty = Thm.ctyp_of_term ctm
+  in
+    Drule.instantiate' [SOME cty] [SOME ctm] @{thm refl}
+  end
+
+fun abs cv th =
+  let
+    val th1 = implies_elim_all th
+    val (tl, tr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th1)))
+    val (ll, lr) = (Thm.lambda cv tl, Thm.lambda cv tr)
+    val (al, ar) = (Thm.apply ll cv, Thm.apply lr cv)
+    val bl = beta al
+    val br = meta_eq_to_obj_eq (Thm.symmetric (Thm.beta_conversion false ar))
+    val th2 = trans (trans bl th1) br
+    val th3 = implies_elim_all th2
+    val th4 = Thm.forall_intr cv th3
+    val i = Drule.instantiate' [SOME (ctyp_of_term cv), SOME (ctyp_of_term tl)]
+      [SOME ll, SOME lr] @{thm ext2}
+  in
+    meta_mp i th4
+  end
+
+fun freezeT thm =
+  let
+    val tvars = Term.add_tvars (prop_of thm) []
+    val tfrees = map (fn ((t, _), s) => TFree (t, s)) tvars
+    val tvars = map TVar tvars
+    val thy = Thm.theory_of_thm thm
+    fun inst ty = ctyp_of thy ty
+  in
+    Thm.instantiate ((map inst tvars ~~ map inst tfrees), []) thm
+  end
+
+fun def' constname rhs thy =
+  let
+    val rhs = term_of rhs
+    val typ = type_of rhs
+    val thy1 = Sign.add_consts_i [(Binding.name constname, typ, NoSyn)] thy
+    val eq = Logic.mk_equals (Const (Sign.intern_const thy1 constname, typ), rhs)
+    val (thms, thy2) = Global_Theory.add_defs false
+      [((Binding.suffix_name "_hldef" (Binding.name constname), eq), [])] thy1
+    val def_thm = freezeT (hd thms)
+  in
+    (meta_eq_to_obj_eq def_thm, thy2)
+  end
+
+fun mdef name thy =
+  case Import_Data.get_const_def name thy of
+    SOME th => th
+  | NONE => error ("constant mapped but no definition: " ^ name)
+
+fun def constname rhs thy =
+  case Import_Data.get_const_def constname thy of
+    SOME _ =>
+      let
+        val () = warning ("Const mapped but def provided: " ^ constname)
+      in
+        (mdef constname thy, thy)
+      end
+  | NONE => def' constname rhs thy
+
+fun typedef_hollight th thy =
+  let
+    val (th_s, cn) = Thm.dest_comb (Thm.dest_arg (cprop_of th))
+    val (th_s, abst) = Thm.dest_comb th_s
+    val rept = Thm.dest_arg th_s
+    val P = Thm.dest_arg cn
+    val [nty, oty] = Thm.dest_ctyp (ctyp_of_term rept)
+  in
+    Drule.instantiate' [SOME nty, SOME oty] [SOME rept, SOME abst, SOME P,
+      SOME (cterm_of thy (Free ("a", typ_of nty))),
+      SOME (cterm_of thy (Free ("r", typ_of oty)))] @{thm typedef_hol2hollight}
+  end
+
+fun tydef' tycname abs_name rep_name cP ct td_th thy =
+  let
+    val ctT = ctyp_of_term ct
+    val nonempty = Drule.instantiate' [SOME ctT] [SOME cP, SOME ct] @{thm light_ex_imp_nonempty}
+    val th2 = meta_mp nonempty td_th
+    val c = case concl_of th2 of
+        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+      | _ => error "type_introduction: bad type definition theorem"
+    val tfrees = Term.add_tfrees c []
+    val tnames = sort_strings (map fst tfrees)
+    val ((_, typedef_info), thy') =
+     Typedef.add_typedef_global false NONE
+       (Binding.name tycname, map (rpair dummyS) tnames, NoSyn) c
+       (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
+    val aty = #abs_type (#1 typedef_info)
+    val th = freezeT (#type_definition (#2 typedef_info))
+    val (th_s, _) = Thm.dest_comb (Thm.dest_arg (cprop_of th))
+    val (th_s, abst) = Thm.dest_comb th_s
+    val rept = Thm.dest_arg th_s
+    val [nty, oty] = Thm.dest_ctyp (ctyp_of_term rept)
+    val typedef_th =
+       Drule.instantiate'
+          [SOME nty, SOME oty]
+          [SOME rept, SOME abst, SOME cP, SOME (cterm_of thy' (Free ("a", aty))),
+             SOME (cterm_of thy' (Free ("r", typ_of ctT)))]
+          @{thm typedef_hol2hollight}
+    val th4 = meta_mp typedef_th (#type_definition (#2 typedef_info))
+  in
+    (th4, thy')
+  end
+
+fun mtydef name thy =
+  case Import_Data.get_typ_def name thy of
+    SOME thn => meta_mp (typedef_hollight thn thy) thn
+  | NONE => error ("type mapped but no tydef thm registered: " ^ name)
+
+fun tydef tycname abs_name rep_name P t td_th thy =
+  case Import_Data.get_typ_def tycname thy of
+    SOME _ =>
+      let
+        val () = warning ("Type mapped but proofs provided: " ^ tycname)
+      in
+        (mtydef tycname thy, thy)
+      end
+  | NONE => tydef' tycname abs_name rep_name P t td_th thy
+
+fun inst_type lambda th thy =
+  let
+    fun assoc _ [] = error "assoc"
+      | assoc x ((x',y)::rest) = if x = x' then y else assoc x rest
+    val lambda = map (fn (a, b) => (typ_of a, b)) lambda
+    val tys_before = Term.add_tfrees (prop_of th) []
+    val th1 = Thm.varifyT_global th
+    val tys_after = Term.add_tvars (prop_of th1) []
+    val tyinst = map2 (fn bef => fn iS =>
+       (case try (assoc (TFree bef)) lambda of
+              SOME cty => (ctyp_of thy (TVar iS), cty)
+            | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
+       )) tys_before tys_after
+  in
+    Thm.instantiate (tyinst,[]) th1
+  end
+
+fun inst sigma th =
+  let
+    val (dom, rng) = ListPair.unzip (rev sigma)
+  in
+    th |> forall_intr_list dom
+       |> forall_elim_list rng
+  end
+
+fun transl_dotc #"." = "dot"
+  | transl_dotc c = Char.toString c
+val transl_dot = String.translate transl_dotc
+
+fun transl_qmc #"?" = "t"
+  | transl_qmc c = Char.toString c
+val transl_qm = String.translate transl_qmc
+
+fun getconstname s thy =
+  case Import_Data.get_const_map s thy of
+      SOME s => s
+    | NONE => Sign.full_name thy (Binding.name (transl_dot s))
+fun gettyname s thy =
+  case Import_Data.get_typ_map s thy of
+    SOME s => s
+  | NONE => Sign.full_name thy (Binding.name s)
+
+fun get (map, no) s =
+  case Int.fromString s of
+    NONE => error "Import_Rule.get: not a number"
+  | SOME i => (case Inttab.lookup map (Int.abs i) of
+      NONE => error "Import_Rule.get: lookup failed"
+    | SOME res => (res, (if i < 0 then Inttab.delete (Int.abs i) map else map, no)))
+
+fun getty i (thy, (tyi, tmi, thi)) = let val (i, tyi) = (get tyi i) in (i, (thy, (tyi, tmi, thi))) end
+fun gettm i (thy, (tyi, tmi, thi)) = let val (i, tmi) = (get tmi i) in (i, (thy, (tyi, tmi, thi))) end
+fun getth i (thy, (tyi, tmi, thi)) = let val (i, thi) = (get thi i) in (i, (thy, (tyi, tmi, thi))) end
+fun set (map, no) v = (Inttab.update_new (no + 1, v) map, no + 1)
+fun setty v (thy, (tyi, tmi, thi)) = (thy, (set tyi v, tmi, thi))
+fun settm v (thy, (tyi, tmi, thi)) = (thy, (tyi, set tmi v, thi))
+fun setth v (thy, (tyi, tmi, thi)) = (thy, (tyi, tmi, set thi v))
+
+fun last_thm (_, _, (map, no)) =
+  case Inttab.lookup map no of
+    NONE => error "Import_Rule.last_thm: lookup failed"
+  | SOME thm => thm
+
+fun listLast (h1 :: (h2 :: t)) = apfst (fn t => h1 :: h2 :: t) (listLast t)
+  | listLast [p] = ([], p)
+  | listLast [] = error "listLast: empty"
+
+fun pairList (h1 :: (h2 :: t)) = ((h1, h2) :: pairList t)
+  | pairList [] = []
+  | pairList _ = error "pairList: odd list length"
+
+fun store_thm binding thm thy =
+  let
+    val thm = Drule.export_without_context_open thm
+    val tvs = Term.add_tvars (prop_of thm) []
+    val tns = map (fn (_, _) => "'") tvs
+    val nms = fst (fold_map Name.variant tns (Variable.names_of (Proof_Context.init_global thy)))
+    val vs = map TVar ((nms ~~ (map (snd o fst) tvs)) ~~ (map snd tvs))
+    val cvs = map (ctyp_of thy) vs
+    val ctvs = map (ctyp_of thy) (map TVar tvs)
+    val thm' = Thm.instantiate ((ctvs ~~ cvs), []) thm
+  in
+    snd (Global_Theory.add_thm ((binding, thm'), []) thy)
+  end
+
+fun process_line str tstate =
+  let
+    fun process tstate (#"R", [t]) = gettm t tstate |>> refl |-> setth
+      | process tstate (#"B", [t]) = gettm t tstate |>> beta |-> setth
+      | process tstate (#"1", [th]) = getth th tstate |>> conj1 |-> setth
+      | process tstate (#"2", [th]) = getth th tstate |>> conj2 |-> setth
+      | process tstate (#"H", [t]) =
+          gettm t tstate |>> Thm.apply @{cterm Trueprop} |>> Thm.trivial |-> setth
+      | process tstate (#"A", [_, t]) =
+          gettm t tstate |>> Thm.apply @{cterm Trueprop} |>> Skip_Proof.make_thm_cterm |-> setth
+      | process tstate (#"C", [th1, th2]) =
+          getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => comb t1 t2) |-> setth
+      | process tstate (#"T", [th1, th2]) =
+          getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => trans t1 t2) |-> setth
+      | process tstate (#"E", [th1, th2]) =
+          getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => eq_mp t1 t2) |-> setth
+      | process tstate (#"D", [th1, th2]) =
+          getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => deduct t1 t2) |-> setth
+      | process tstate (#"L", [t, th]) =
+          gettm t tstate ||>> (fn ti => getth th ti) |>> (fn (tm, th) => abs tm th) |-> setth
+      | process (thy, state) (#"M", [s]) =
+          let
+            val ctxt = Variable.set_body false (Proof_Context.init_global thy)
+            val thm = freezeT (Global_Theory.get_thm thy s)
+            val ((_, [th']), _) = Variable.import true [thm] ctxt
+          in
+            setth th' (thy, state)
+          end
+      | process (thy, state) (#"Q", l) =
+          let
+            val (tys, th) = listLast l
+            val (th, tstate) = getth th (thy, state)
+            val (tys, tstate) = fold_map getty tys tstate
+          in
+            setth (inst_type (pairList tys) th thy) tstate
+          end
+      | process tstate (#"S", l) =
+          let
+            val (tms, th) = listLast l
+            val (th, tstate) = getth th tstate
+            val (tms, tstate) = fold_map gettm tms tstate
+          in
+            setth (inst (pairList tms) th) tstate
+          end
+      | process tstate (#"F", [name, t]) =
+          let
+            val (tm, (thy, state)) = gettm t tstate
+            val (th, thy) = def (transl_dot name) tm thy
+          in
+            setth th (thy, state)
+          end
+      | process (thy, state) (#"F", [name]) = setth (mdef name thy) (thy, state)
+      | process tstate (#"Y", [name, absname, repname, t1, t2, th]) =
+          let
+            val (th, tstate) = getth th tstate
+            val (t1, tstate) = gettm t1 tstate
+            val (t2, (thy, state)) = gettm t2 tstate
+            val (th, thy) = tydef name absname repname t1 t2 th thy
+          in
+            setth th (thy, state)
+          end
+      | process (thy, state) (#"Y", [name, _, _]) = setth (mtydef name thy) (thy, state)
+      | process (thy, state) (#"t", [n]) =
+          setty (ctyp_of thy (TFree ("'" ^ (transl_qm n), ["HOL.type"]))) (thy, state)
+      | process (thy, state) (#"a", n :: l) =
+          fold_map getty l (thy, state) |>>
+            (fn tys => ctyp_of thy (Type (gettyname n thy, map typ_of tys))) |-> setty
+      | process (thy, state) (#"v", [n, ty]) =
+          getty ty (thy, state) |>> (fn ty => cterm_of thy (Free (transl_dot n, typ_of ty))) |-> settm
+      | process (thy, state) (#"c", [n, ty]) =
+          getty ty (thy, state) |>> (fn ty => cterm_of thy (Const (getconstname n thy, typ_of ty))) |-> settm
+      | process tstate (#"f", [t1, t2]) =
+          gettm t1 tstate ||>> gettm t2 |>> (fn (t1, t2) => Thm.apply t1 t2) |-> settm
+      | process tstate (#"l", [t1, t2]) =
+          gettm t1 tstate ||>> gettm t2 |>> (fn (t1, t2) => Thm.lambda t1 t2) |-> settm
+      | process (thy, state) (#"+", [s]) =
+          let
+            val _ = tracing ("Noting: " ^ s)
+          in
+            (store_thm (Binding.name (transl_dot s)) (last_thm state) thy, state)
+          end
+      | process _ (c, _) = error ("process: unknown command: " ^ String.implode [c])
+
+    fun parse_line s =
+        case String.tokens (fn x => (x = #"\n" orelse x = #" ")) s of
+          [] => error "parse_line: empty"
+        | h :: t => (case String.explode h of
+            [] => error "parse_line: empty command"
+          | sh :: st => (sh, (String.implode st) :: t))
+  in
+    process tstate (parse_line str)
+  end
+
+fun process_file path thy =
+  (thy, init_state) |> File.fold_lines process_line path |> fst
+
+val _ = Outer_Syntax.command @{command_spec "import_file"}
+  "Import a recorded proof file" (Parse.path >> process_file >> Toplevel.theory)
+
+
+end
--- a/src/HOL/Import/proof_kernel.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2084 +0,0 @@
-(*  Title:      HOL/Import/proof_kernel.ML
-    Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
-*)
-
-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 option
-    val disambiguate_frees : Thm.thm -> Thm.thm
-    val debug : bool Unsynchronized.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_importer_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 to_hol_thm : Thm.thm -> thm
-
-    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
-
-    val prin : term -> unit
-    val protect_factname : string -> string
-    val replay_protect_varname : string -> string -> unit
-    val replay_add_dump : string -> theory -> theory
-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
-
-fun hthm2thm (HOLThm (_, th)) = th
-
-fun to_hol_thm th = HOLThm ([], th)
-
-val replay_add_dump = add_dump
-fun add_dump s thy = replay_add_dump s thy
-
-datatype proof_info
-  = Info of {disk_info: (string * string) option Unsynchronized.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)
-
-
-(* Compatibility. *)
-
-val string_of_mixfix = Pretty.string_of o Mixfix.pretty_mixfix;
-
-fun mk_syn thy c =
-  if Lexicon.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
-  else Delimfix (Syntax_Ext.escape c)
-
-fun quotename c =
-  if Lexicon.is_identifier c andalso not (Keyword.is_keyword c) then c else quote c
-
-exception SMART_STRING
-
-fun no_vars context tm =
-  let
-    val ctxt = Variable.set_body false context;
-    val ([tm'], _) = Variable.import_terms true [tm] ctxt;
-  in tm' end
-
-fun smart_string_of_cterm ctxt0 ct =
-    let
-        val ctxt = ctxt0
-          |> Config.put show_brackets false
-          |> Config.put show_all_types false
-          |> Config.put show_types false
-          |> Config.put show_sorts false;
-        val {t,T,...} = rep_cterm ct
-        (* Hack to avoid parse errors with Trueprop *)
-        val t' = HOLogic.dest_Trueprop t
-                 handle TERM _ => t
-        val tn = no_vars ctxt t'
-        fun match u =
-            let val _ = Thm.match (ct, cterm_of (Proof_Context.theory_of ctxt) u) in true end
-            handle Pattern.MATCH => false
-        fun G 0 f ctxt x = f ctxt x
-          | G 1 f ctxt x = f (Config.put show_types true ctxt) x
-          | G 2 f ctxt x = G 1 f (Config.put show_sorts true ctxt) x
-          | G 3 f ctxt x = G 2 f (Config.put show_all_types true ctxt) x
-          | G 4 f ctxt x = G 3 f (Config.put show_brackets true ctxt) x
-          | G _ _ _ _ = raise SMART_STRING
-        fun F n =
-            let
-                val str = G n Syntax.string_of_term ctxt tn
-                val _ = warning (@{make_string} (n, str))
-                val u = Syntax.parse_term ctxt str
-                val u = if t = t' then u else HOLogic.mk_Trueprop u
-                val u = Syntax.check_term ctxt (Type.constraint T u)
-            in
-                if match u
-                then quote str
-                else F (n+1)
-            end
-            handle ERROR _ => F (n + 1)
-              | SMART_STRING =>
-                  let val _ =
-                    warning ("smart_string failed for: "^ G 0 Syntax.string_of_term ctxt (term_of ct))
-                  in quote (G 2 Syntax.string_of_term ctxt tn) end
-    in
-      Print_Mode.setmp [] F 0
-    end
-
-fun smart_string_of_thm ctxt = smart_string_of_cterm ctxt o cprop_of
-
-fun prth th = writeln (Print_Mode.setmp [] Display.string_of_thm_without_context th);
-val topctxt = ML_Context.the_local_context ();
-fun prin t = writeln (Print_Mode.setmp []
-  (fn () => Syntax.string_of_term topctxt t) ());
-fun pth (HOLThm(_,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 = Unsynchronized.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 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
-infix mem;
-fun i mem L =
-    let fun itr [] = false
-          | itr (a::rst) = i=a orelse itr rst
-    in itr L end;
-
-infix union;
-fun [] union S = S
-  | S union [] = S
-  | (a::rst) union S2 = rst union (insert (op =) a S2);
-
-fun implode_subst [] = []
-  | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
-  | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
-
-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
-
-(* Actual code. *)
-
-fun get_segment thyname l = (Lib.assoc "s" l
-                             handle PK _ => thyname)
-val get_name : (string * string) list -> string = Lib.assoc "n"
-
-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 xml_to_import_xml (XML.Elem ((n, l), ts)) = Elem (n, l, map xml_to_import_xml ts)
-  | xml_to_import_xml (XML.Text _) = raise XML "Incorrect proof file: text";
-
-val type_of = Term.type_of
-
-val propT = Type("prop",[])
-
-fun mk_defeq name rhs thy =
-    let
-        val ty = type_of rhs
-    in
-        Logic.mk_equals (Const(Sign.intern_const 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 thy name,ty),rhs)
-    end
-
-fun intern_const_name thyname const thy =
-    case get_importer_const_mapping thyname const thy of
-        SOME (_,cname,_) => cname
-      | NONE => (case get_importer_const_renaming thyname const thy of
-                     SOME cname => Sign.intern_const thy (thyname ^ "." ^ cname)
-                   | NONE => Sign.intern_const thy (thyname ^ "." ^ const))
-
-fun intern_type_name thyname const thy =
-    case get_importer_type_mapping thyname const thy of
-        SOME (_,cname) => cname
-      | NONE => Sign.intern_const 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 mk_thy_const thy Thy Nam Ty = Const(intern_const_name Thy Nam thy,Ty)
-
-local
-  fun get_const sg _ name =
-    (case Sign.const_type sg name of
-      SOME ty => Const (name, ty)
-    | NONE => raise ERR "get_type" (name ^ ": No such constant"))
-in
-fun prim_mk_const thy Thy Nam =
-    let
-      val name = intern_const_name Thy Nam thy
-      val cmaps = Importer_ConstMaps.get thy
-    in
-      case StringPair.lookup cmaps (Thy,Nam) of
-        SOME(_,_,SOME ty) => Const(name,ty)
-      | _ => get_const thy Thy name
-    end
-end
-
-fun mk_comb(f,a) = f $ 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
-
-val protected_varnames = Unsynchronized.ref (Symtab.empty:string Symtab.table)
-val invented_isavar = Unsynchronized.ref 0
-
-fun innocent_varname s = Lexicon.is_identifier s andalso not (String.isPrefix "u_" s)
-
-fun valid_boundvarname s =
-  can (fn () => Syntax.read_term_global @{theory Main} ("SOME "^s^". True")) ();
-
-fun valid_varname s =
-  can (fn () => Syntax.read_term_global @{theory Main} s) ();
-
-fun protect_varname s =
-    if innocent_varname s andalso valid_varname s then s else
-    case Symtab.lookup (!protected_varnames) s of
-      SOME t => t
-    | NONE =>
-      let
-          val _ = Unsynchronized.inc invented_isavar
-          val t = "u_" ^ string_of_int (!invented_isavar)
-          val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
-      in
-          t
-      end
-
-exception REPLAY_PROTECT_VARNAME of string*string*string
-
-fun replay_protect_varname s t =
-        case Symtab.lookup (!protected_varnames) s of
-          SOME t' => raise REPLAY_PROTECT_VARNAME (s, t, t')
-        | NONE =>
-          let
-              val _ = Unsynchronized.inc invented_isavar
-              val t = "u_" ^ string_of_int (!invented_isavar)
-              val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
-          in
-              ()
-          end
-
-fun protect_boundvarname s = if innocent_varname s andalso valid_boundvarname s then s else "u"
-
-fun mk_lambda (v as Free (x, T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
-  | mk_lambda (v as Var ((x, _), T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
-  | mk_lambda v t = raise TERM ("lambda", [v, t]);
-
-fun replacestr x y s =
-let
-  val xl = raw_explode x
-  val yl = raw_explode y
-  fun isprefix [] _ = true
-    | isprefix (x::xs) (y::ys) = if x = y then isprefix xs ys else false
-    | isprefix _ _ = false
-  fun isp s = isprefix xl s
-  fun chg s = yl@(List.drop (s, List.length xl))
-  fun r [] = []
-    | r (S as (s::ss)) = if isp S then r (chg S) else s::(r ss)
-in
-  implode(r (raw_explode s))
-end
-
-fun protect_factname s = replacestr "." "_dot_" s
-fun unprotect_factname s = replacestr "_dot_" "." s
-
-val ty_num_prefix = "N_"
-
-fun startsWithDigit s = Char.isDigit (hd (String.explode s))
-
-fun protect_tyname tyn =
-  let
-    val tyn' =
-      if String.isPrefix ty_num_prefix tyn then raise (ERR "protect_ty_name" ("type name '"^tyn^"' is reserved")) else
-      (if startsWithDigit tyn then ty_num_prefix^tyn else tyn)
-  in
-    tyn'
-  end
-
-fun protect_constname tcn = tcn
- (* if tcn = ".." then "dotdot"
-  else if tcn = "==" then "eqeq"
-  else tcn*)
-
-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)
-                        (protect_tyname (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)
-                        (protect_tyname (get_name atts))
-                        (map gtfx tys)
-          | gtfx _ = raise ERR "get_type" "Bad type"
-    in
-        gtfx
-    end
-
-fun input_types _ (Elem("tylist",[("i",i)],xtys)) =
-    let
-        val types = Array.array(the (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 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 = protect_constname(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_lambda 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 _ _ (Elem("tmlist",[("i",i)],xtms)) =
-    let
-        val terms = Array.array(the (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 path = space_explode ":" (getenv "IMPORTER_PROOFS")
-        fun find [] = NONE
-          | find (p::ps) =
-            (let
-                 val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
-             in
-                 if OS.FileSys.isDir dir
-                 then SOME dir
-                 else find ps
-             end) handle OS.SysErr _ => find ps
-    in
-        Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
-    end
-
-fun proof_file_name thyname thmname thy =
-    let
-        val path = case get_proof_dir thyname thy of
-                       SOME p => p
-                     | NONE => error "Cannot find proof files"
-        val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
-    in
-        OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = (unprotect_factname thmname), ext = SOME "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 = protect_factname (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,protect_constname 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,protect_tyname name,protect_constname abs_name,protect_constname rep_name,P,t,x2p p))
-            end
-          | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
-            mk_proof (PTyDef(seg,protect_tyname name,x2p p))
-          | x2p (Elem("poracle",[],chldr)) =
-            let
-                val (oracles,terms) = List.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
-                val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | _ => raise ERR "x2p" "bad oracle") oracles
-                val (c,asl) = case terms of
-                                  [] => raise ERR "x2p" "Bad oracle description"
-                                | (hd::tl) => (hd,tl)
-                val tg = fold_rev (Tag.merge o Tag.read) 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 _ = raise ERR "x2p" "Bad proof"
-    in
-        x2p prf
-    end
-
-fun import_proof_concl thyname thmname thy =
-    let
-        val is = TextIO.openIn(proof_file_name thyname thmname thy)
-        val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll is))
-        val _ = TextIO.closeIn is
-    in
-        case proof_xml of
-            Elem("proof",[],xtypes::xterms::_::rest) =>
-            let
-                val types = TypeNet.input_types thyname xtypes
-                val terms = TermNet.input_terms thyname types xterms
-                fun f xtm thy = TermNet.get_term_from_xml thy thyname types terms xtm
-            in
-                case rest of
-                    [] => NONE
-                  | [xtm] => SOME (f xtm)
-                  | _ => raise ERR "import_proof" "Bad argument list"
-            end
-          | _ => raise ERR "import_proof" "Bad proof"
-    end
-
-fun import_proof thyname thmname thy =
-    let
-        val is = TextIO.openIn(proof_file_name thyname thmname thy)
-        val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll 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 = Thm.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 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 thy = theory_of_thm th
-    val pp = Thm.reflexive (cterm_of thy (Const(@{const_name Trueprop},HOLogic.boolT-->propT)))
-in
-val not_elim_thm = Thm.combination pp th
-end
-
-val not_intro_thm = Thm.symmetric not_elim_thm
-val abs_thm = @{thm ext}
-val trans_thm = @{thm trans}
-val symmetry_thm = @{thm sym}
-val eqmp_thm = @{thm iffD1}
-val eqimp_thm = @{thm Importer.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 =
-    Library.foldl (fn (th,p) => Thm.implies_elim th (Thm.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 (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(@{const_name All},_) $ Abs(oldname,_,_)) =>
-                              (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
-                            | tpc $ (Const(@{const_name All},_) $ _) => (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
-
-fun rearrange sg tm th =
-    let
-        val tm' = Envir.beta_eta_contract tm
-        fun find []      n = Thm.permute_prems 0 1 (Thm.implies_intr (Thm.cterm_of sg tm) th)
-          | find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
-                             then Thm.permute_prems n 1 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
-
-(* Code for disambiguating variablenames (wrt. types) *)
-
-val disamb_info_empty = {vars=[],rens=[]}
-
-fun rens_of { vars = _, rens = rens } = rens
-
-fun disamb_term_from info tm = (info, tm)
-
-fun has_ren (HOLThm _) = false
-
-fun disamb_thm_from info (HOLThm (_,thm)) = (info, thm)
-
-fun disamb_terms_from info tms = (info, tms)
-
-fun disamb_thms_from info hthms = (info, map hthm2thm hthms)
-
-fun disamb_term tm   = disamb_term_from disamb_info_empty tm
-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 _ (hth as HOLThm _) = hth
-
-(* End of disambiguating code *)
-
-fun disambiguate_frees thm =
-    let
-      fun ERR s = error ("Drule.disambiguate_frees: "^s)
-      val ct = cprop_of thm
-      val t = term_of ct
-      val thy = theory_of_cterm ct
-      val frees = Misc_Legacy.term_frees t
-      val freenames = Term.add_free_names t []
-      val is_old_name = member (op =) freenames
-      fun name_of (Free (n, _)) = n
-        | name_of _ = ERR "name_of"
-      fun new_name' bump map n =
-          let val n' = n^bump in
-            if is_old_name n' orelse Symtab.lookup map n' <> NONE then
-              new_name' (Symbol.bump_string bump) map n
-            else
-              n'
-          end
-      val new_name = new_name' "a"
-      fun replace_name n' (Free (_, t)) = Free (n', t)
-        | replace_name _ _ = ERR "replace_name"
-      (* map: old or fresh name -> old free,
-         invmap: old free which has fresh name assigned to it -> fresh name *)
-      fun dis v (mapping as (map,invmap)) =
-          let val n = name_of v in
-            case Symtab.lookup map n of
-              NONE => (Symtab.update (n, v) map, invmap)
-            | SOME v' =>
-              if v=v' then
-                mapping
-              else
-                let val n' = new_name map n in
-                  (Symtab.update (n', v) map,
-                   Termtab.update (v, n') invmap)
-                end
-          end
-    in
-      if (length freenames = length frees) then
-        thm
-      else
-        let
-          val (_, invmap) =
-              fold dis frees (Symtab.empty, Termtab.empty)
-          fun make_subst (oldfree, newname) (intros, elims) =
-              (cterm_of thy oldfree :: intros,
-               cterm_of thy (replace_name newname oldfree) :: elims)
-          val (intros, elims) = fold make_subst (Termtab.dest invmap) ([], [])
-        in
-          forall_elim_list elims (forall_intr_list intros thm)
-        end
-    end
-
-val debug = Unsynchronized.ref false
-
-fun if_debug f x = if !debug then f x else ()
-val message = if_debug writeln
-
-fun get_importer_thm thyname thmname thy =
-    case get_importer_theorem thyname thmname thy of
-        SOME hth => SOME (HOLThm hth)
-      | NONE =>
-        let
-            val pending = Importer_Pending.get thy
-        in
-            case StringPair.lookup pending (thyname,thmname) of
-                SOME hth => SOME (HOLThm hth)
-              | NONE => NONE
-        end
-
-fun non_trivial_term_consts t = fold_aterms
-  (fn Const (c, _) =>
-      if c = @{const_name Trueprop} orelse c = @{const_name All}
-        orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
-      then I else insert (op =) c
-    | _ => I) t [];
-
-fun split_name str =
-    let
-        val sub = Substring.full str
-        val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
-        val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
-    in
-        if not (idx = "") andalso u = "_"
-        then SOME (newstr, the (Int.fromString idx))
-        else NONE
-    end
-    handle _ => NONE  (* FIXME avoid handle _ *)
-
-fun rewrite_importer_term t thy =
-    let
-        val import_rews1 = map (Thm.transfer thy) (Importer_Rewrites.get thy)
-        val importerss = Simplifier.global_context thy empty_ss addsimps import_rews1
-    in
-        Thm.transfer thy (Simplifier.full_rewrite importerss (cterm_of thy t))
-    end
-
-fun get_isabelle_thm thyname thmname importerconc thy =
-    let
-        val (info,importerconc') = disamb_term importerconc
-        val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') 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 importerconc';
-                 message "Modified conclusion:";
-                 if_debug prin isaconc)
-
-        fun mk_res th = HOLThm (rens_of info, Thm.equal_elim i2h_conc th)
-    in
-        case get_importer_mapping thyname thmname thy of
-            SOME (SOME thmname) =>
-            let
-                val th1 = (SOME (Global_Theory.get_thm thy thmname)
-                           handle ERROR _ =>
-                                  (case split_name thmname of
-                                       SOME (listname,idx) => (SOME (nth (Global_Theory.get_thms thy listname) (idx - 1))
-                                                               handle _ => NONE)  (* FIXME avoid handle _ *)
-                                     | 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 (commas cs))
-                val pot_thms = Shuffler.find_potential thy isaconc
-                val _ = message (string_of_int (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_importer_theorem thyname thmname args
-                                        |> add_importer_mapping thyname thmname isaname
-                    in
-                        (thy',SOME hth)
-                    end
-                  | NONE => (thy,NONE)
-            end
-    end
-    handle e =>
-      if Exn.is_interrupt e then reraise e
-      else
-        (if_debug (fn () =>
-            writeln ("Exception in get_isabelle_thm:\n" ^ ML_Compiler.exn_message e)) ();
-          (thy,NONE))
-
-fun get_isabelle_thm_and_warn thyname thmname importerconc thy =
-  let
-    val (a, b) = get_isabelle_thm thyname thmname importerconc thy
-    fun warn () =
-        let
-            val (_,importerconc') = disamb_term importerconc
-            val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') thy)
-        in
-            case concl_of i2h_conc of
-                Const("==",_) $ lhs $ _ =>
-                (warning ("Failed lookup of theorem '"^thmname^"':");
-                 writeln "Original conclusion:";
-                 prin importerconc';
-                 writeln "Modified conclusion:";
-                 prin lhs)
-              | _ => ()
-        end
-  in
-      case b of
-          NONE => (warn () handle _ => (); (a,b)) (* FIXME avoid handle _ *)
-        | _ => (a, b)
-  end
-
-fun get_thm thyname thmname thy =
-    case get_importer_thm thyname thmname thy of
-        SOME hth => (thy,SOME hth)
-      | NONE => ((case import_proof_concl thyname thmname thy of
-                      SOME f => get_isabelle_thm_and_warn thyname thmname (f thy) thy
-                    | NONE => (message "No conclusion"; (thy,NONE)))
-                 handle IO.Io _ => (message "IO exception"; (thy,NONE))
-                      | PK _ => (message "PK exception"; (thy,NONE)))
-
-fun rename_const thyname thy name =
-    case get_importer_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_and_warn 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 (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
-        val rew = rewrite_importer_term (concl_of th) thy
-        val th  = Thm.equal_elim rew th
-        val thy' = add_importer_pending thyname thmname args thy
-        val th = disambiguate_frees th
-        val th = Object_Logic.rulify th
-        val thy2 =
-          if gen_output
-          then
-            add_dump ("lemma " ^ (quotename thmname) ^ ": " ^
-                (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^ "\n  by (import " ^
-                thyname ^ " " ^ (quotename 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 ctm = Thm.cterm_of thy 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 ctm = Thm.cterm_of thy (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(_,th)) thy =
-    let
-        val _ = message "INST_TYPE:"
-        val _ = if_debug pth hth
-        val tys_before = Misc_Legacy.add_term_tfrees (prop_of th,[])
-        val th1 = Thm.varifyT_global th
-        val tys_after = Misc_Legacy.add_term_tvars (prop_of th1,[])
-        val tyinst = map (fn (bef, iS) =>
-                             (case try (Lib.assoc (TFree bef)) lambda of
-                                  SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
-                                | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
-                             ))
-                         (zip tys_before tys_after)
-        val res = Drule.instantiate_normalize (tyinst,[]) th1
-        val hth = HOLThm([],res)
-        val res = norm_hthm thy 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 (sdom,srng) = ListPair.unzip (rev sigma)
-        val th = hthm2thm hth
-        val th1 = mk_INST (map (cterm_of thy) sdom) (map (cterm_of thy) srng) th
-        val res = HOLThm([],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 thy =
-    let
-        val (f,g) = case concl_of th1 of
-                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
-                      | _ => raise ERR "mk_COMB" "First theorem not an equality"
-        val (x,y) = case concl_of th2 of
-                        _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
-                      | _ => raise ERR "mk_COMB" "Second theorem not an equality"
-        val fty = type_of f
-        val (fd,fr) = Term.dest_funT fty
-        val comb_thm' = Drule.instantiate'
-                            [SOME (ctyp_of thy fd),SOME (ctyp_of thy fr)]
-                            [SOME (cterm_of thy f),SOME (cterm_of thy g),
-                             SOME (cterm_of thy x),SOME (cterm_of thy 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 cctxt = cterm_of thy ctxt'
-        fun subst th [] = th
-          | subst th (rew::rews) = subst (mk_COMB th rew thy) 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 th1 = norm_hyps th1
-        val th2 = norm_hyps th2
-        val (l,r) = case concl_of th of
-                        _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
-                      | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
-        val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
-        val th2' = rearrange thy (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 ct = Thm.cterm_of thy 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 ct = Thm.cterm_of thy 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 cwit = cterm_of thy wit'
-        val cty = ctyp_of_term cwit
-        val a = case ex' of
-                    (Const(@{const_name Ex},_) $ a) => a
-                  | _ => raise ERR "EXISTS" "Argument not existential"
-        val ca = cterm_of thy 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 (Thm.implies_elim th (Thm.assume p))
-          | strip _ _ _ = raise ERR "CHOOSE" "strip error"
-        val cv = cterm_of thy v'
-        val th2 = norm_hyps th2
-        val cvty = ctyp_of_term cv
-        val c = HOLogic.dest_Trueprop (concl_of th2)
-        val cc = cterm_of thy c
-        val a = case concl_of th1 of
-                    _ $ (Const(@{const_name Ex},_) $ a) => a
-                  | _ => raise ERR "CHOOSE" "Conclusion not existential"
-        val ca = cterm_of (theory_of_thm th1) a
-        val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
-        val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
-        val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
-        val th23 = beta_eta_thm (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 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 ctm = Thm.cterm_of thy 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 res = HOLThm(rens_of info,mk_COMB th1 th2 thy)
-        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 ct = cterm_of thy tm'
-        val th1 = rearrange thy (HOLogic.mk_Trueprop (Const(@{const_name Not},HOLogic.boolT-->HOLogic.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 thy =
-    let
-        val cv = cterm_of thy v
-        val th1 = implies_elim_all (beta_eta_thm th)
-        val (f,g) = case concl_of th1 of
-                        _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
-                      | _ => raise ERR "mk_ABS" "Bad conclusion"
-        val (fd,fr) = Term.dest_funT (type_of f)
-        val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
-        val th2 = Thm.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 res = HOLThm(rens_of info',mk_ABS v' th thy)
-        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 th1 =
-            case copt of
-                SOME (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 (Type(name,tys)) =
-                        Type(name,map (inst_type ty1 ty2) tys)
-                in
-                    fold_rev (fn v => fn th =>
-                              let
-                                  val cdom = fst (Term.dest_funT (fst (Term.dest_funT cty)))
-                                  val vty  = type_of v
-                                  val newcty = inst_type cdom vty cty
-                                  val cc = cterm_of thy (Const(cname,newcty))
-                              in
-                                  mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy
-                              end) vlist' th
-                end
-              | SOME _ => raise ERR "GEN_ABS" "Bad constant"
-              | NONE =>
-                fold_rev (fn v => fn th => mk_ABS v th thy) 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 th1 = implies_elim_all (beta_eta_thm th)
-        val a = case concl_of th1 of
-                    _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
-                  | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
-        val ca = cterm_of thy a
-        val th2 = Thm.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 th1 = implies_elim_all (beta_eta_thm th)
-        val a = case concl_of th1 of
-                    _ $ (Const(@{const_name Not},_) $ a) => a
-                  | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
-        val ca = cterm_of thy a
-        val th2 = Thm.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 th1 = beta_eta_thm th
-        val th2 = implies_elim_all th1
-        val th3 = Thm.implies_intr (cterm_of thy (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 = space_implode " "
-
-fun new_definition thyname constname rhs thy =
-    let
-        val constname = rename_const thyname thy constname
-        val redeclared = is_some (Sign.const_type thy (Sign.intern_const thy constname));
-        val _ = warning ("Introducing constant " ^ constname)
-        val (thmname,thy) = get_defname thyname constname thy
-        val (_,rhs') = disamb_term rhs
-        val ctype = type_of rhs'
-        val csyn = mk_syn thy constname
-        val thy1 = case Importer_DefThy.get thy of
-                       Replaying _ => thy
-                     | _ => Sign.add_consts_i [(Binding.name constname,ctype,csyn)] thy
-        val eq = mk_defeq constname rhs' thy1
-        val (thms, thy2) = Global_Theory.add_defs false [((Binding.name 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 thy' constname
-        val thy'' = add_importer_const_mapping thyname constname true fullcname thy'
-        val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
-        val rew = rewrite_importer_term eq thy''
-        val crhs = cterm_of thy'' (#2 (Logic.dest_equals (prop_of rew)))
-        val thy22 =
-          if Thm.def_name constname = thmname andalso not redeclared andalso csyn = NoSyn
-          then
-              let
-                  val ctxt = Syntax.init_pretty_global thy''
-                  val p1 = quotename constname
-                  val p2 = Syntax.string_of_typ ctxt ctype
-                  val p3 = string_of_mixfix csyn
-                  val p4 = smart_string_of_cterm ctxt crhs
-              in
-                add_dump ("definition\n  " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ " where\n  " ^ p4) thy''
-              end
-          else
-              let val ctxt = Syntax.init_pretty_global thy'' in
-                add_dump ("consts\n  " ^ quotename constname ^ " :: \"" ^
-                  Syntax.string_of_typ ctxt ctype ^
-                  "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n  " ^
-                  quotename thmname ^ ": " ^ smart_string_of_cterm ctxt crhs) thy''
-              end
-        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_bname thy22 thmname
-        val thy22' = case opt_get_output_thy thy22 of
-                         "" => add_importer_mapping thyname thmname fullname thy22
-                       | output_thy =>
-                         let
-                             val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
-                         in
-                             thy22 |> add_importer_move fullname moved_thmname
-                                   |> add_importer_mapping thyname thmname moved_thmname
-                         end
-        val _ = message "new_definition:"
-        val _ = if_debug pth hth
-    in
-        (thy22',hth)
-    end
-
-fun new_specification thyname thmname names hth thy =
-    case Importer_DefThy.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 " ^ commas names)
-            val (HOLThm(rens,th)) = norm_hthm thy hth
-            val thy1 = case Importer_DefThy.get thy of
-                           Replaying _ => thy
-                         | _ =>
-                           let
-                               fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
-                                 | dest_eta_abs body =
-                                   let
-                                       val (dT,_) = Term.dest_funT (type_of body)
-                                   in
-                                       ("x",dT,body $ Bound 0)
-                                   end
-                                   handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
-                               fun dest_exists (Const(@{const_name Ex},_) $ abody) =
-                                   dest_eta_abs abody
-                                 | dest_exists _ =
-                                   raise ERR "new_specification" "Bad existential formula"
-
-                               val (consts,_) = Library.foldl (fn ((cs,ex),cname) =>
-                                                          let
-                                                              val (_,cT,p) = dest_exists ex
-                                                          in
-                                                              ((cname,cT,mk_syn thy cname)::cs,p)
-                                                          end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
-                               val str = Library.foldl (fn (acc, (c, T, csyn)) =>
-                                   acc ^ "\n  " ^ quotename c ^ " :: \"" ^
-                                   Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
-                               val thy' = add_dump str thy
-                           in
-                               Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy'
-                           end
-
-            val thy1 = fold_rev (fn name => fn thy =>
-                                snd (get_defname thyname name thy)) names thy1
-            fun new_name name = fst (get_defname thyname name thy1)
-            val names' = map (fn name => (new_name name,name,false)) names
-            val (thy',res) = Choice_Specification.add_specification NONE
-                                 names'
-                                 (thy1,th)
-            val res' = Thm.unvarify_global res
-            val hth = HOLThm(rens,res')
-            val rew = rewrite_importer_term (concl_of res') thy'
-            val th  = Thm.equal_elim rew res'
-            fun handle_const (name,thy) =
-                let
-                    val defname = Thm.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') = fold_rev (fn name => fn (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 (Syntax.init_pretty_global thy') 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
-
-fun new_axiom name _ _ = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
-
-fun to_isa_thm (hth as HOLThm(_,th)) =
-    let
-        val (HOLThm args) = norm_hthm (theory_of_thm th) hth
-    in
-        apsnd Thm.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 Importer_DefThy.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 thy hth
-            val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
-            val c = case concl_of th2 of
-                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
-                      | _ => raise ERR "new_type_definition" "Bad type definition theorem"
-            val tfrees = Misc_Legacy.term_tfrees c
-            val tnames = map fst tfrees
-            val tsyn = mk_syn thy tycname
-            val ((_, typedef_info), thy') =
-              Typedef.add_typedef_global false (SOME (Binding.name thmname))
-                (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c NONE (rtac th2 1) thy
-
-            val th3 = (#type_definition (#2 typedef_info)) RS typedef_hol2hol4
-
-            val fulltyname = Sign.intern_type thy' tycname
-            val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
-
-            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th3))
-            val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
-                    else ()
-            val thy4 = add_importer_pending thyname thmname args thy''
-
-            val rew = rewrite_importer_term (concl_of td_th) thy4
-            val th  = Thm.equal_elim rew (Thm.transfer thy4 td_th)
-            val c   = case HOLogic.dest_Trueprop (prop_of th) of
-                          Const(@{const_name Ex},exT) $ P =>
-                          let
-                              val PT = domain_type exT
-                          in
-                              Const (@{const_name 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 "(" ^ commas tnames ^ ") "
-            val proc_prop =
-              smart_string_of_cterm
-                (Syntax.init_pretty_global thy4
-                  |> not (null tnames) ? Config.put show_all_types true)
-            val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " "
-                                 ^ (string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
-
-            val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2importer [OF type_definition_" ^ tycname ^ "]") thy5
-
-            val _ = message "RESULT:"
-            val _ = if_debug pth hth'
-        in
-            (thy6,hth')
-        end
-
-fun add_dump_syntax thy name =
-    let
-      val n = quotename name
-      val syn = string_of_mixfix (mk_syn thy name)
-    in
-      add_dump ("syntax\n  "^n^" :: _ "^syn) thy
-    end
-
-fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
-    case Importer_DefThy.get thy of
-        Replaying _ => (thy,
-          HOLThm([], Global_Theory.get_thm thy (thmname^"_@intern")) handle ERROR _ => 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 thy hth
-            val tT = type_of t
-            val light_nonempty' =
-                Drule.instantiate' [SOME (ctyp_of thy tT)]
-                                   [SOME (cterm_of thy P),
-                                    SOME (cterm_of thy t)] light_nonempty
-            val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
-            val c = case concl_of th2 of
-                        _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
-                      | _ => raise ERR "type_introduction" "Bad type definition theorem"
-            val tfrees = Misc_Legacy.term_tfrees c
-            val tnames = sort_strings (map fst tfrees)
-            val tsyn = mk_syn thy tycname
-            val ((_, typedef_info), thy') =
-              Typedef.add_typedef_global false NONE
-                (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c
-                (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
-            val fulltyname = Sign.intern_type thy' tycname
-            val aty = Type (fulltyname, map mk_vartype tnames)
-            val typedef_hol2hollight' =
-                Drule.instantiate'
-                    [SOME (ctyp_of thy' aty), SOME (ctyp_of thy' tT)]
-                    [NONE, NONE, NONE, SOME (cterm_of thy' (Free ("a", aty))), SOME (cterm_of thy' (Free ("r", tT)))]
-                    typedef_hol2hollight
-            val th4 = (#type_definition (#2 typedef_info)) RS typedef_hol2hollight'
-            val _ = null (Thm.fold_terms Term.add_tvars th4 []) orelse
-              raise ERR "type_introduction" "no type variables expected any more"
-            val _ = null (Thm.fold_terms Term.add_vars th4 []) orelse
-              raise ERR "type_introduction" "no term variables expected any more"
-            val _ = message ("step 3: thyname="^thyname^", tycname="^tycname^", fulltyname="^fulltyname)
-            val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
-            val _ = message "step 4"
-            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th4))
-            val thy4 = add_importer_pending thyname thmname args thy''
-
-            val P' = P (* why !? #2 (Logic.dest_equals (concl_of (rewrite_importer_term P thy4))) *)
-            val c   =
-                let
-                    val PT = type_of P'
-                in
-                    Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P'
-                end
-
-            val tnames_string = if null tnames
-                                then ""
-                                else "(" ^ commas tnames ^ ") "
-            val proc_prop =
-              smart_string_of_cterm
-                (Syntax.init_pretty_global thy4
-                  |> not (null tnames) ? Config.put show_all_types true)
-            val thy = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^
-              " = " ^ (proc_prop (cterm_of thy4 c)) ^ " " ^
-              (string_of_mixfix tsyn) ^ " morphisms "^
-              (quote rep_name)^" "^(quote abs_name)^"\n"^
-              ("  apply (rule light_ex_imp_nonempty[where t="^
-              (proc_prop (cterm_of thy4 t))^"])\n"^
-              ("  by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
-            val str_aty = Syntax.string_of_typ_global thy aty
-            val thy = add_dump_syntax thy rep_name
-            val thy = add_dump_syntax thy abs_name
-            val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
-              " = typedef_hol2hollight \n"^
-              "  [where a=\"a :: "^str_aty^"\" and r=r" ^
-              " ,\n   OF "^(quotename ("type_definition_" ^ tycname)) ^ "]") thy
-            val _ = message "RESULT:"
-            val _ = if_debug pth hth'
-        in
-            (thy,hth')
-        end
-end
-
-val prin = prin
-
-end
--- a/src/HOL/Import/replay.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,344 +0,0 @@
-(*  Title:      HOL/Import/replay.ML
-    Author:     Sebastian Skalberg (TU Muenchen)
-*)
-
-structure Replay =  (* FIXME proper signature *)
-struct
-
-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 = ProofKernel.REFL tm thy
-          | rp (PInstT(p,lambda)) thy =
-            let
-                val (thy',th) = rp' p thy
-            in
-                ProofKernel.INST_TYPE lambda th thy'
-            end
-          | rp (PSubst(prfs,ctxt,prf)) thy =
-            let
-                val (thy',ths) = fold_rev (fn p => fn (thy, ths) =>
-                                           let
-                                               val (thy',th) = rp' p thy
-                                           in
-                                               (thy',th::ths)
-                                           end) prfs (thy,[])
-                val (thy'',th) = rp' prf thy'
-            in
-                ProofKernel.SUBST ths ctxt th thy''
-            end
-          | rp (PAbs(prf,v)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.ABS v th thy'
-            end
-          | rp (PDisch(prf,tm)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.DISCH tm th thy'
-            end
-          | rp (PMp(prf1,prf2)) thy =
-            let
-                val (thy1,th1) = rp' prf1 thy
-                val (thy2,th2) = rp' prf2 thy1
-            in
-                ProofKernel.MP th1 th2 thy2
-            end
-          | rp (PHyp tm) thy = ProofKernel.ASSUME tm thy
-          | rp (PDef(seg,name,rhs)) thy =
-            (case ProofKernel.get_def seg name rhs thy of
-                 (thy',SOME res) => (thy',res)
-               | (thy',NONE) => 
-                 if seg = thyname
-                 then ProofKernel.new_definition seg name rhs thy'
-                 else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
-          | rp (POracle _) thy = NY "ORACLE"
-          | rp (PSpec(prf,tm)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.SPEC tm th thy'
-            end
-          | rp (PInst(prf,theta)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.INST theta th thy'
-            end
-          | rp (PGen(prf,v)) thy =
-            let
-                val (thy',th) = rp' prf thy
-                val p = ProofKernel.GEN v th thy'
-            in
-                p
-            end
-          | rp (PGenAbs(prf,opt,vl)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.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
-                ProofKernel.IMP_ANTISYM th1 th2 thy2
-            end
-          | rp (PSym prf) thy =
-            let
-                val (thy1,th) = rp' prf thy
-            in
-                ProofKernel.SYM th thy1
-            end
-          | rp (PTrans(prf1,prf2)) thy =
-            let
-                val (thy1,th1) = rp' prf1 thy
-                val (thy2,th2) = rp' prf2 thy1
-            in
-                ProofKernel.TRANS th1 th2 thy2
-            end
-          | rp (PComb(prf1,prf2)) thy =
-            let
-                val (thy1,th1) = rp' prf1 thy
-                val (thy2,th2) = rp' prf2 thy1
-            in
-                ProofKernel.COMB th1 th2 thy2
-            end
-          | rp (PEqMp(prf1,prf2)) thy =
-            let
-                val (thy1,th1) = rp' prf1 thy
-                val (thy2,th2) = rp' prf2 thy1
-            in
-                ProofKernel.EQ_MP th1 th2 thy2
-            end
-          | rp (PEqImp prf) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.EQ_IMP_RULE th thy'
-            end
-          | rp (PExists(prf,ex,wit)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.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
-                ProofKernel.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
-                ProofKernel.CONJ th1 th2 thy2
-            end
-          | rp (PConjunct1 prf) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.CONJUNCT1 th thy'
-            end
-          | rp (PConjunct2 prf) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.CONJUNCT2 th thy'
-            end
-          | rp (PDisj1(prf,tm)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.DISJ1 th tm thy'
-            end
-          | rp (PDisj2(prf,tm)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.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
-                ProofKernel.DISJ_CASES th th1 th2 thy2
-            end
-          | rp (PNotI prf) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.NOT_INTRO th thy'
-            end
-          | rp (PNotE prf) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.NOT_ELIM th thy'
-            end
-          | rp (PContr(prf,tm)) thy =
-            let
-                val (thy',th) = rp' prf thy
-            in
-                ProofKernel.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 ProofKernel.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 (_, prf) = import_proof thyname' thmname thy'
-                                                    val prf = prf thy'
-                                                    val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
-                                                    val _ = writeln ("Successfully finished replaying "^thmname^" !")
-                                                in
-                                                    case content_of prf of
-                                                        PTmSpec _ => (thy',th)
-                                                      | PTyDef  _ => (thy',th)
-                                                      | PTyIntro _ => (thy',th)
-                                                      | _ => ProofKernel.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 ProofKernel.get_axiom thyname name thy of
-                            (thy',SOME res) => (thy',res)
-                          | (thy',NONE) => ProofKernel.new_axiom name c thy')
-                  | PTmSpec(seg,names,prf') =>
-                    let
-                        val (thy',th) = rp' prf' thy
-                    in
-                        ProofKernel.new_specification seg thmname names th thy'
-                    end
-                  | PTyDef(seg,name,prf') =>
-                    let
-                        val (thy',th) = rp' prf' thy
-                    in
-                        ProofKernel.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
-                        ProofKernel.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
-                    end
-                  | _ => rp pc thy
-            end
-    in
-        rp' prf thy
-    end
-
-fun setup_int_thms thyname thy =
-    let
-        val fname =
-            case ProofKernel.get_proof_dir thyname thy of
-                SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
-              | NONE => error "Cannot find proof files"
-        val is = TextIO.openIn fname
-        val (num_int_thms,facts) =
-            let
-                fun get_facts facts =
-                    case TextIO.inputLine is of
-                        NONE => (case facts of
-                                   i::facts => (the (Int.fromString i),map ProofKernel.protect_factname (rev facts))
-                                 | _ => raise ERR "replay_thm" "Bad facts.lst file")
-                      | SOME 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 prf = mk_proof PDisk
-                val _ = set_disk_info_of prf thyname thmname
-                val _ = writeln ("Replaying "^thmname^" ...")
-                val p = fst (replay_proof int_thms thyname thmname prf thy)
-            in
-                p
-            end
-    in
-        replay_fact (thmname,thy)
-    end
-
-fun import_thms thyname int_thms thmnames thy =
-    let
-        fun replay_fact thmname thy = 
-            let
-                val prf = mk_proof PDisk        
-                val _ = set_disk_info_of prf thyname thmname
-                val _ = writeln ("Replaying "^thmname^" ...")
-                val p = fst (replay_proof int_thms thyname thmname prf thy)
-            in
-                p
-            end
-        val res_thy = fold replay_fact thmnames thy
-    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 prf = mk_proof PDisk        
-                val _ = set_disk_info_of prf thyname thmname        
-                val _ = writeln ("Replaying "^thmname^" ...")
-                val p = fst (replay_proof int_thms thyname thmname prf thy)
-            in 
-                p
-            end
-    in
-        replay_fact (thmname,thy)
-    end
-
-end
--- a/src/HOL/Import/shuffler.ML	Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,585 +0,0 @@
-(*  Title:      HOL/Import/shuffler.ML
-    Author:     Sebastian Skalberg, TU Muenchen
-
-Package for proving two terms equal by normalizing (hence the
-"shuffler" name).  Uses the simplifier for the normalization.
-*)
-
-signature Shuffler =
-sig
-    val debug      : bool Unsynchronized.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: Proof.context -> bool -> (string * thm) list -> int -> tactic
-    val shuffle_tac: Proof.context -> thm list -> int -> tactic
-    val search_tac : Proof.context -> int -> tactic
-
-    val print_shuffles: theory -> unit
-
-    val add_shuffle_rule: thm -> theory -> theory
-    val shuffle_attr: attribute
-
-    val setup      : theory -> theory
-end
-
-structure Shuffler : Shuffler =
-struct
-
-val debug = Unsynchronized.ref false
-
-fun if_debug f x = if !debug then f x else ()
-val message = if_debug writeln
-
-val string_of_thm = Print_Mode.setmp [] Display.string_of_thm_without_context;
-
-structure ShuffleData = Theory_Data
-(
-  type T = thm list
-  val empty = []
-  val extend = I
-  val merge = Thm.merge_thms
-)
-
-fun print_shuffles thy =
-  Pretty.writeln (Pretty.big_list "Shuffle theorems:"
-    (map (Display.pretty_thm_global thy) (ShuffleData.get thy)))
-
-val weaken =
-    let
-        val cert = cterm_of Pure.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 cPQ = cert PQ
-        val cPPQ = cert PPQ
-        val th1 = Thm.assume cPQ |> implies_intr_list [cPQ,cP]
-        val th3 = Thm.assume cP
-        val th4 = implies_elim_list (Thm.assume cPPQ) [th3,th3]
-                                    |> implies_intr_list [cPPQ,cP]
-    in
-        Thm.equal_intr th4 th1 |> Drule.export_without_context
-    end
-
-val imp_comm =
-    let
-        val cert = cterm_of Pure.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 (Thm.assume cPQR) [Thm.assume cP,Thm.assume cQ]
-                                    |> implies_intr_list [cPQR,cQ,cP]
-        val th2 = implies_elim_list (Thm.assume cQPR) [Thm.assume cQ,Thm.assume cP]
-                                    |> implies_intr_list [cQPR,cP,cQ]
-    in
-        Thm.equal_intr th1 th2 |> Drule.export_without_context
-    end
-
-val all_comm =
-    let
-        val cert = cterm_of Pure.thy
-        val xT = TFree("'a",[])
-        val yT = TFree("'b",[])
-        val x = Free("x",xT)
-        val y = Free("y",yT)
-        val P = Free("P",xT-->yT-->propT)
-        val lhs = Logic.all x (Logic.all y (P $ x $ y))
-        val rhs = Logic.all y (Logic.all x (P $ x $ y))
-        val cl = cert lhs
-        val cr = cert rhs
-        val cx = cert x
-        val cy = cert y
-        val th1 = Thm.assume cr
-                         |> forall_elim_list [cy,cx]
-                         |> forall_intr_list [cx,cy]
-                         |> Thm.implies_intr cr
-        val th2 = Thm.assume cl
-                         |> forall_elim_list [cx,cy]
-                         |> forall_intr_list [cy,cx]
-                         |> Thm.implies_intr cl
-    in
-        Thm.equal_intr th1 th2 |> Drule.export_without_context
-    end
-
-val equiv_comm =
-    let
-        val cert = cterm_of Pure.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  = Thm.assume ctu |> Thm.symmetric |> Thm.implies_intr ctu
-        val th2  = Thm.assume cut |> Thm.symmetric |> Thm.implies_intr cut
-    in
-        Thm.equal_intr th1 th2 |> Drule.export_without_context
-    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 thy xv yv t =
-    let
-        val lhs = Logic.list_all ([xv,yv],t)
-        val rhs = Logic.list_all ([yv,xv],swap_bound 0 t)
-        val rew = Logic.mk_equals (lhs,rhs)
-        val init = Thm.trivial (cterm_of thy rew)
-    in
-        all_comm RS init
-    end
-
-fun quant_rewrite thy _ (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 thy (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    = Thm.reflexive (cterm_of thy t)
-                         val newt_th = Thm.reflexive (cterm_of thy newt)
-                     in
-                         SOME (Thm.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 = Misc_Legacy.term_tvars t
-        val tfree_names = Misc_Legacy.add_term_tfree_names(t,[])
-        val (type_inst,_) =
-            fold (fn (w as (v,_), S) => fn (inst, used) =>
-                      let
-                          val v' = singleton (Name.variant_list used) v
-                      in
-                          ((w,TFree(v',S))::inst,v'::used)
-                      end)
-                  tvars ([], tfree_names)
-        val t' = subst_TVars type_inst t
-    in
-        (t', map (fn (w,TFree(v,S)) => (v,TVar(w,S))
-                  | _ => error "Internal error in Shuffler.freeze_thaw") type_inst)
-    end
-
-fun inst_tfrees thy [] thm = thm
-  | inst_tfrees thy ((name,U)::rest) thm =
-    let
-        val cU = ctyp_of thy U
-        val tfrees = Misc_Legacy.add_term_tfrees (prop_of thm,[])
-        val (rens, thm') = Thm.varifyT_global'
-    (remove (op = o apsnd fst) name tfrees) thm
-        val mid =
-            case rens of
-                [] => thm'
-              | [((_, S), idx)] => Drule.instantiate_normalize
-            ([(ctyp_of thy (TVar (idx, S)), cU)], []) thm'
-              | _ => error "Shuffler.inst_tfrees internal error"
-    in
-        inst_tfrees thy 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 thy _ origt =
-    let
-        val (typet,Tinst) = freeze_thaw_term origt
-        val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
-        val final = inst_tfrees thy 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 (cat_lines
-                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
-                      Syntax.string_of_term_global thy origt,
-                      Syntax.string_of_term_global thy lhs,
-                      Syntax.string_of_term_global thy typet,
-                      Syntax.string_of_term_global thy t] @
-                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
-                else ()
-            end
-    in
-        case t of
-            Const("all",_) $ (Abs(x,xT,Const("==",_) $ P $ Q)) =>
-            (if eta_redex P andalso eta_redex Q
-              then
-                  let
-                      val cert = cterm_of thy
-                      val v = Free (singleton (Name.variant_list (Term.add_free_names t [])) "v", xT)
-                      val cv = cert v
-                      val ct = cert t
-                      val th = (Thm.assume ct)
-                                   |> Thm.forall_elim cv
-                                   |> Thm.abstract_rule x cv
-                      val ext_th = Thm.eta_conversion (cert (Abs(x,xT,P)))
-                      val th' = Thm.transitive (Thm.symmetric ext_th) th
-                      val cu = cert (prop_of th')
-                      val uth = Thm.combination (Thm.assume cu) (Thm.reflexive cv)
-                      val uth' = (Thm.beta_conversion false (cert (Abs(x,xT,Q) $ v)))
-                                     |> Thm.transitive uth
-                                     |> Thm.forall_intr cv
-                                     |> Thm.implies_intr cu
-                      val rew_th = Thm.equal_intr (th' |> Thm.implies_intr ct) uth'
-                      val res = final rew_th
-                  in
-                       SOME res
-                  end
-              else NONE)
-          | _ => NONE
-       end
-
-fun eta_expand thy _ origt =
-    let
-        val (typet,Tinst) = freeze_thaw_term origt
-        val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
-        val final = inst_tfrees thy 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 (cat_lines
-                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
-                      Syntax.string_of_term_global thy origt,
-                      Syntax.string_of_term_global thy lhs,
-                      Syntax.string_of_term_global thy typet,
-                      Syntax.string_of_term_global thy t] @
-                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
-                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,_]) =>
-                      let
-                          val cert = cterm_of thy
-                          val vname = singleton (Name.variant_list (Term.add_free_names t [])) "v"
-                          val v = Free(vname,aT)
-                          val cv = cert v
-                          val ct = cert t
-                          val th1 = (Thm.combination (Thm.assume ct) (Thm.reflexive cv))
-                                        |> Thm.forall_intr cv
-                                        |> Thm.implies_intr ct
-                          val concl = cert (concl_of th1)
-                          val th2 = (Thm.assume concl)
-                                        |> Thm.forall_elim cv
-                                        |> Thm.abstract_rule vname cv
-                          val (lhs,rhs) = Logic.dest_equals (prop_of th2)
-                          val elhs = Thm.eta_conversion (cert lhs)
-                          val erhs = Thm.eta_conversion (cert rhs)
-                          val th2' = Thm.transitive
-                                         (Thm.transitive (Thm.symmetric elhs) th2)
-                                         erhs
-                          val res = Thm.equal_intr th1 (th2' |> Thm.implies_intr concl)
-                          val res' = final res
-                      in
-                          SOME res'
-                      end
-                    | _ => NONE)
-            else NONE
-          | _ => error ("Bad eta_expand argument" ^ Syntax.string_of_term_global thy t)
-    end;
-
-fun mk_tfree s = TFree("'"^s,[])
-fun mk_free s t = Free (s,t)
-val xT = mk_tfree "a"
-val yT = mk_tfree "b"
-val x = Free ("x", xT)
-val y = Free ("y", yT)
-val P  = mk_free "P" (xT-->yT-->propT)
-val Q  = mk_free "Q" (xT-->yT)
-val R  = mk_free "R" (xT-->yT)
-in
-
-fun quant_simproc thy = Simplifier.simproc_global_i
-                           thy
-                           "Ordered rewriting of nested quantifiers"
-                           [Logic.all x (Logic.all y (P $ x $ y))]
-                           quant_rewrite
-fun eta_expand_simproc thy = Simplifier.simproc_global_i
-                         thy
-                         "Smart eta-expansion by equivalences"
-                         [Logic.mk_equals(Q,R)]
-                         eta_expand
-fun eta_contract_simproc thy = Simplifier.simproc_global_i
-                         thy
-                         "Smart handling of eta-contractions"
-                         [Logic.all x (Logic.mk_equals (Q $ x, R $ x))]
-                         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 thy 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(_,xT,t),idx) =
-            let
-                val x' = "x" ^ string_of_int idx
-                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 thy t
-        val ct' = cterm_of thy t'
-        val res = Thm.transitive (Thm.reflexive ct) (Thm.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 norms = ShuffleData.get thy
-        val ss = Simplifier.global_context thy empty_ss
-          addsimps (map (Thm.transfer thy) norms)
-          addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy]
-        fun chain f th =
-            let
-                val rhs = Thm.rhs_of th
-            in
-                Thm.transitive th (f rhs)
-            end
-        val th =
-            t |> disamb_bound thy
-              |> chain (Simplifier.full_rewrite ss)
-              |> chain Thm.eta_conversion
-              |> Thm.strip_shyps
-        val _ = message ("norm_term: " ^ (string_of_thm th))
-    in
-        th
-    end
-
-
-(* Closes a theorem with respect to free and schematic variables (does
-not touch type variables, though). *)
-
-fun close_thm th =
-    let
-        val thy = Thm.theory_of_thm th
-        val c = prop_of th
-        val vars = Misc_Legacy.add_term_frees (c, Misc_Legacy.add_term_vars(c,[]))
-    in
-        Drule.forall_intr_list (map (cterm_of thy) vars) th
-    end
-
-
-(* Normalizes a theorem's conclusion using norm_term. *)
-
-fun norm_thm thy th =
-    let
-        val c = prop_of th
-    in
-        Thm.equal_elim (norm_term thy c) th
-    end
-
-(* make_equal thy t u tries to construct the theorem t == u under the
-signature thy.  If it succeeds, SOME (t == u) is returned, otherwise
-NONE is returned. *)
-
-fun make_equal thy t u =
-    let
-        val t_is_t' = norm_term thy t
-        val u_is_u' = norm_term thy u
-        val th = Thm.transitive t_is_t' (Thm.symmetric u_is_u')
-        val _ = message ("make_equal: SOME " ^ (string_of_thm th))
-    in
-        SOME th
-    end
-    handle THM _ => (message "make_equal: NONE";NONE)
-
-fun match_consts ignore t (* th *) =
-    let
-        fun add_consts (Const (c, _), cs) =
-            if member (op =) ignore c
-            then cs
-            else insert (op =) c 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) =>
-        let
-            val th_consts = add_consts(prop_of th,[])
-        in
-            eq_set (op =) (t_consts, th_consts)
-        end
-    end
-
-val collect_ignored = fold_rev (fn thm => fn cs =>
-  let
-    val (lhs, rhs) = Logic.dest_equals (prop_of thm);
-    val consts_lhs = Term.add_const_names lhs [];
-    val consts_rhs = Term.add_const_names rhs [];
-    val ignore_lhs = subtract (op =) consts_rhs consts_lhs;
-    val ignore_rhs = subtract (op =) consts_lhs consts_rhs;
-  in
-    fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
-  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 vars = Misc_Legacy.add_term_frees (t, Misc_Legacy.add_term_vars (t,[]))
-        val closed_t = fold_rev Logic.all vars t
-        val rew_th = norm_term thy closed_t
-        val rhs = Thm.rhs_of rew_th
-
-        fun process [] = NONE
-          | process ((name,th)::thms) =
-            let
-                val norm_th = Thm.varifyT_global (norm_thm thy (close_thm (Thm.transfer thy th)))
-                val triv_th = Thm.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 (Thm.bicompose false (*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 = Thm.equal_elim (Thm.symmetric rew_th) mod_th
-                    in
-                        message ("Shuffler.set_prop succeeded by " ^ name);
-                        SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)
-                    end
-                  | NONE => process thms
-            end
-            handle THM _ => process thms
-    in
-        fn thms =>
-           case process thms of
-               res as SOME (_,th) => if (prop_of th) aconv t
-                                        then res
-                                        else error "Internal error in set_prop"
-             | NONE => NONE
-    end
-
-fun find_potential thy t =
-    let
-        val shuffles = ShuffleData.get thy
-        val ignored = collect_ignored shuffles []
-        val all_thms =
-          map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (Global_Theory.facts_of thy)))
-    in
-        filter (match_consts ignored t) all_thms
-    end
-
-fun gen_shuffle_tac ctxt search thms = SUBGOAL (fn (t, i) =>
-    let
-        val thy = Proof_Context.theory_of ctxt
-        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)
-    end)
-
-fun shuffle_tac ctxt thms =
-  gen_shuffle_tac ctxt false (map (pair "") thms);
-
-fun search_tac ctxt =
-  gen_shuffle_tac ctxt true (map (pair "premise") (Assumption.all_prems_of ctxt));
-
-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
-
-val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I);
-
-val setup =
-  Method.setup @{binding shuffle_tac}
-    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (shuffle_tac ctxt ths)))
-    "solve goal by shuffling terms around" #>
-  Method.setup @{binding search_tac}
-    (Scan.succeed (SIMPLE_METHOD' o search_tac)) "search for suitable theorems" #>
-  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.setup @{binding shuffle_rule} (Scan.succeed shuffle_attr) "declare rule for shuffler";
-
-end
--- a/src/HOL/IsaMakefile	Sun Apr 01 09:12:03 2012 +0200
+++ b/src/HOL/IsaMakefile	Sun Apr 01 14:50:47 2012 +0200
@@ -82,8 +82,7 @@
     # ^ this is the sort position
 
 generate: \
-  HOL-HOL4-Generate \
-  HOL-HOL_Light-Generate
+  HOL-HOL4-Generate
 
 benchmark: \
   HOL-Datatype_Benchmark \
@@ -558,24 +557,17 @@
 ## Import sessions
 
 IMPORTER_BASIC_DEPENDENCIES = \
-  Import/Importer.thy \
-  Import/shuffler.ML \
-  Import/import_rews.ML \
-  Import/proof_kernel.ML \
-  Import/replay.ML \
-  Import/import.ML
+  Import/HOL4/Importer.thy \
+  Import/HOL4/shuffler.ML \
+  Import/HOL4/import_rews.ML \
+  Import/HOL4/proof_kernel.ML \
+  Import/HOL4/replay.ML \
+  Import/HOL4/import.ML
 
 IMPORTER_HOL4_COMPATIBILITY_DEDEPENDENCIES = \
   Import/HOL4/ROOT.ML \
   Import/HOL4/Compatibility.thy
 
-IMPORTER_HOL_LIGHT_COMPATIBILITY_DEDEPENDENCIES = \
-  Import/HOL_Light/ROOT.ML \
-  Import/HOL_Light/HOLLightInt.thy \
-  Import/HOL_Light/HOLLightList.thy \
-  Import/HOL_Light/HOLLightReal.thy \
-  Import/HOL_Light/Compatibility.thy
-
 HOL-HOL4: $(LOG)/HOL-HOL4.gz
 
 $(LOG)/HOL-HOL4.gz: $(OUT)/HOL \
@@ -586,9 +578,13 @@
 HOL-HOL_Light: $(LOG)/HOL-HOL_Light.gz
 
 $(LOG)/HOL-HOL_Light.gz: $(OUT)/HOL \
-  $(IMPORTER_BASIC_DEPENDENCIES) \
-  $(IMPORTER_HOL_LIGHT_COMPATIBILITY_DEDEPENDENCIES)
-	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Import/HOL_Light
+  Import/ROOT.ML \
+  Import/Import_Setup.thy \
+  Import/import_data.ML \
+  Import/import_rule.ML \
+  Import/HOL_Light_Maps.thy \
+  Import/HOL_Light_Import.thy
+	@cd Import; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-HOL_Light
 
 HOL-HOL4-Imported: $(LOG)/HOL-HOL4-Imported.gz
 
@@ -652,17 +648,6 @@
   Import/HOL4/Generated/word_num.imp
 	@$(ISABELLE_TOOL) usedir -f imported.ML -s HOL4-Imported -p 0 $(OUT)/HOL Import/HOL4
 
-HOL-HOL_Light-Imported: $(LOG)/HOL-HOL_Light-Imported.gz
-
-$(LOG)/HOL-HOL_Light-Imported.gz: $(OUT)/HOL \
-  $(IMPORTER_BASIC_DEPENDENCIES) \
-  $(IMPORTER_HOL_LIGHT_COMPATIBILITY_DEDEPENDENCIES) \
-  Import/HOL_Light/imported.ML \
-  Import/HOL_Light/Imported.thy \
-  Import/HOL_Light/Generated/HOLLight.thy \
-  Import/HOL_Light/Generated/hollight.imp
-	@$(ISABELLE_TOOL) usedir -f imported.ML -s HOL_Light-Imported -p 0 $(OUT)/HOL Import/HOL_Light
-
 HOL-HOL4-Generate: $(LOG)/Import-HOL4-Generate.gz
 
 $(LOG)/Import-HOL4-Generate.gz: $(OUT)/HOL \
@@ -677,17 +662,6 @@
   Import/HOL4/Template/GenHOL4Word32.thy
 	@$(ISABELLE_TOOL) usedir -f generate.ML -s HOL4-Generate -p 0 $(OUT)/HOL Import/HOL4
 
-HOL-HOL_Light-Generate: $(LOG)/Import-HOL_Light-Generate.gz
-
-$(LOG)/Import-HOL_Light-Generate.gz: $(OUT)/HOL \
-  $(IMPORTER_BASIC_DEPENDENCIES) \
-  $(IMPORTER_HOL_LIGHT_COMPATIBILITY_DEDEPENDENCIES) \
-  Import/HOL_Light/generate.ML \
-  Import/HOL_Light/Generate.thy \
-  Import/HOL_Light/Template/GenHOLLight.thy
-	@$(ISABELLE_TOOL) usedir -f generate.ML -s HOL_Light-Generate -p 0 $(OUT)/HOL Import/HOL_Light
-
-
 ## HOL-Number_Theory
 
 HOL-Number_Theory: HOL $(LOG)/HOL-Number_Theory.gz