src/HOL/Import/proof_kernel.ML
author wenzelm
Fri, 08 Apr 2011 16:34:14 +0200
changeset 43162 b1f544c84040
parent 43160 2074b31650e6
child 43235 8c674b3b8e44
permissions -rw-r--r--
discontinued special treatment of structure Lexicon;
     1 (*  Title:      HOL/Import/proof_kernel.ML
     2     Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
     3 *)
     4 
     5 signature ProofKernel =
     6 sig
     7     type hol_type
     8     type tag
     9     type term
    10     type thm
    11     type ('a,'b) subst
    12 
    13     type proof_info
    14     datatype proof = Proof of proof_info * proof_content
    15          and proof_content
    16            = PRefl of term
    17            | PInstT of proof * (hol_type,hol_type) subst
    18            | PSubst of proof list * term * proof
    19            | PAbs of proof * term
    20            | PDisch of proof * term
    21            | PMp of proof * proof
    22            | PHyp of term
    23            | PAxm of string * term
    24            | PDef of string * string * term
    25            | PTmSpec of string * string list * proof
    26            | PTyDef of string * string * proof
    27            | PTyIntro of string * string * string * string * term * term * proof
    28            | POracle of tag * term list * term
    29            | PDisk
    30            | PSpec of proof * term
    31            | PInst of proof * (term,term) subst
    32            | PGen of proof * term
    33            | PGenAbs of proof * term option * term list
    34            | PImpAS of proof * proof
    35            | PSym of proof
    36            | PTrans of proof * proof
    37            | PComb of proof * proof
    38            | PEqMp of proof * proof
    39            | PEqImp of proof
    40            | PExists of proof * term * term
    41            | PChoose of term * proof * proof
    42            | PConj of proof * proof
    43            | PConjunct1 of proof
    44            | PConjunct2 of proof
    45            | PDisj1 of proof * term
    46            | PDisj2 of proof * term
    47            | PDisjCases of proof * proof * proof
    48            | PNotI of proof
    49            | PNotE of proof
    50            | PContr of proof * term
    51 
    52     exception PK of string * string
    53 
    54     val get_proof_dir: string -> theory -> string option
    55     val disambiguate_frees : Thm.thm -> Thm.thm
    56     val debug : bool Unsynchronized.ref
    57     val disk_info_of : proof -> (string * string) option
    58     val set_disk_info_of : proof -> string -> string -> unit
    59     val mk_proof : proof_content -> proof
    60     val content_of : proof -> proof_content
    61     val import_proof : string -> string -> theory -> (theory -> term) option * (theory -> proof)
    62 
    63     val rewrite_hol4_term: Term.term -> theory -> Thm.thm
    64 
    65     val type_of : term -> hol_type
    66 
    67     val get_thm  : string -> string         -> theory -> (theory * thm option)
    68     val get_def  : string -> string -> term -> theory -> (theory * thm option)
    69     val get_axiom: string -> string         -> theory -> (theory * thm option)
    70 
    71     val store_thm : string -> string -> thm -> theory -> theory * thm
    72 
    73     val to_isa_thm : thm -> (term * term) list * Thm.thm
    74     val to_isa_term: term -> Term.term
    75     val to_hol_thm : Thm.thm -> thm
    76 
    77     val REFL : term -> theory -> theory * thm
    78     val ASSUME : term -> theory -> theory * thm
    79     val INST_TYPE : (hol_type,hol_type) subst -> thm -> theory -> theory * thm
    80     val INST : (term,term)subst -> thm -> theory -> theory * thm
    81     val EQ_MP : thm -> thm -> theory -> theory * thm
    82     val EQ_IMP_RULE : thm -> theory -> theory * thm
    83     val SUBST : thm list -> term -> thm -> theory -> theory * thm
    84     val DISJ_CASES : thm -> thm -> thm -> theory -> theory * thm
    85     val DISJ1: thm -> term -> theory -> theory * thm
    86     val DISJ2: term -> thm -> theory -> theory * thm
    87     val IMP_ANTISYM: thm -> thm -> theory -> theory * thm
    88     val SYM : thm -> theory -> theory * thm
    89     val MP : thm -> thm -> theory -> theory * thm
    90     val GEN : term -> thm -> theory -> theory * thm
    91     val CHOOSE : term -> thm -> thm -> theory -> theory * thm
    92     val EXISTS : term -> term -> thm -> theory -> theory * thm
    93     val ABS : term -> thm -> theory -> theory * thm
    94     val GEN_ABS : term option -> term list -> thm -> theory -> theory * thm
    95     val TRANS : thm -> thm -> theory -> theory * thm
    96     val CCONTR : term -> thm -> theory -> theory * thm
    97     val CONJ : thm -> thm -> theory -> theory * thm
    98     val CONJUNCT1: thm -> theory -> theory * thm
    99     val CONJUNCT2: thm -> theory -> theory * thm
   100     val NOT_INTRO: thm -> theory -> theory * thm
   101     val NOT_ELIM : thm -> theory -> theory * thm
   102     val SPEC : term -> thm -> theory -> theory * thm
   103     val COMB : thm -> thm -> theory -> theory * thm
   104     val DISCH: term -> thm -> theory -> theory * thm
   105 
   106     val type_introduction: string -> string -> string -> string -> string -> term * term -> thm -> theory -> theory * thm
   107 
   108     val new_definition : string -> string -> term -> theory -> theory * thm
   109     val new_specification : string -> string -> string list -> thm -> theory -> theory * thm
   110     val new_type_definition : string -> string -> string -> thm -> theory -> theory * thm
   111     val new_axiom : string -> term -> theory -> theory * thm
   112 
   113     val prin : term -> unit
   114     val protect_factname : string -> string
   115     val replay_protect_varname : string -> string -> unit
   116     val replay_add_dump : string -> theory -> theory
   117 end
   118 
   119 structure ProofKernel :> ProofKernel =
   120 struct
   121 type hol_type = Term.typ
   122 type term = Term.term
   123 datatype tag = Tag of string list
   124 type ('a,'b) subst = ('a * 'b) list
   125 datatype thm = HOLThm of (Term.term * Term.term) list * Thm.thm
   126 
   127 fun hthm2thm (HOLThm (_, th)) = th
   128 
   129 fun to_hol_thm th = HOLThm ([], th)
   130 
   131 val replay_add_dump = add_dump
   132 fun add_dump s thy = (ImportRecorder.add_dump s; replay_add_dump s thy)
   133 
   134 datatype proof_info
   135   = Info of {disk_info: (string * string) option Unsynchronized.ref}
   136 
   137 datatype proof = Proof of proof_info * proof_content
   138      and proof_content
   139        = PRefl of term
   140        | PInstT of proof * (hol_type,hol_type) subst
   141        | PSubst of proof list * term * proof
   142        | PAbs of proof * term
   143        | PDisch of proof * term
   144        | PMp of proof * proof
   145        | PHyp of term
   146        | PAxm of string * term
   147        | PDef of string * string * term
   148        | PTmSpec of string * string list * proof
   149        | PTyDef of string * string * proof
   150        | PTyIntro of string * string * string * string * term * term * proof
   151        | POracle of tag * term list * term
   152        | PDisk
   153        | PSpec of proof * term
   154        | PInst of proof * (term,term) subst
   155        | PGen of proof * term
   156        | PGenAbs of proof * term option * term list
   157        | PImpAS of proof * proof
   158        | PSym of proof
   159        | PTrans of proof * proof
   160        | PComb of proof * proof
   161        | PEqMp of proof * proof
   162        | PEqImp of proof
   163        | PExists of proof * term * term
   164        | PChoose of term * proof * proof
   165        | PConj of proof * proof
   166        | PConjunct1 of proof
   167        | PConjunct2 of proof
   168        | PDisj1 of proof * term
   169        | PDisj2 of proof * term
   170        | PDisjCases of proof * proof * proof
   171        | PNotI of proof
   172        | PNotE of proof
   173        | PContr of proof * term
   174 
   175 exception PK of string * string
   176 fun ERR f mesg = PK (f,mesg)
   177 
   178 
   179 (* Compatibility. *)
   180 
   181 val string_of_mixfix = Pretty.string_of o Mixfix.pretty_mixfix;
   182 
   183 fun mk_syn thy c =
   184   if Lexicon.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
   185   else Delimfix (Syntax_Ext.escape c)
   186 
   187 fun quotename c =
   188   if Lexicon.is_identifier c andalso not (Keyword.is_keyword c) then c else quote c
   189 
   190 fun simple_smart_string_of_cterm ctxt0 ct =
   191     let
   192         val ctxt = ctxt0
   193           |> Config.put show_brackets false
   194           |> Config.put show_all_types true
   195           |> Config.put show_sorts true
   196           |> Config.put Syntax.ambiguity_enabled true;
   197         val {t,T,...} = rep_cterm ct
   198         (* Hack to avoid parse errors with Trueprop *)
   199         val ct  = cterm_of (Thm.theory_of_cterm ct) (HOLogic.dest_Trueprop t)
   200                            handle TERM _ => ct
   201     in
   202         quote (Print_Mode.setmp [] (Syntax.string_of_term ctxt o Thm.term_of) ct)
   203     end
   204 
   205 exception SMART_STRING
   206 
   207 fun smart_string_of_cterm ctxt0 ct =
   208     let
   209         val ctxt = ctxt0 |> Config.put Syntax.ambiguity_enabled false;
   210         val {t,T,...} = rep_cterm ct
   211         (* Hack to avoid parse errors with Trueprop *)
   212         val ct  = cterm_of (Thm.theory_of_cterm ct) (HOLogic.dest_Trueprop t)
   213                    handle TERM _ => ct
   214         fun match u = t aconv u
   215         fun G 0 f ctxt x = f (Config.put show_types true (Config.put show_sorts true ctxt)) x
   216           | G 1 f ctxt x = G 0 f (Config.put show_brackets true ctxt) x
   217           | G 2 f ctxt x = G 0 f (Config.put show_all_types true ctxt) x
   218           | G 3 f ctxt x = G 2 f (Config.put show_brackets true ctxt) x
   219           | G _ _ _ _ = raise SMART_STRING
   220         fun F n =
   221             let
   222                 val str =
   223                   G n Syntax.string_of_term (Config.put show_brackets false ctxt) (term_of ct)
   224                 val u = Syntax.parse_term ctxt str
   225                   |> Type.constraint T |> Syntax.check_term ctxt
   226             in
   227                 if match u
   228                 then quote str
   229                 else F (n+1)
   230             end
   231             handle ERROR mesg => F (n + 1)
   232               | SMART_STRING =>
   233                   error ("smart_string failed for: "^ G 0 Syntax.string_of_term ctxt (term_of ct))
   234     in
   235       Print_Mode.setmp [] F 0
   236     end
   237     handle ERROR mesg => simple_smart_string_of_cterm ctxt0 ct
   238 
   239 fun smart_string_of_thm ctxt = smart_string_of_cterm ctxt o cprop_of
   240 
   241 fun prth th = writeln (Print_Mode.setmp [] Display.string_of_thm_without_context th);
   242 fun prin t = writeln (Print_Mode.setmp []
   243   (fn () => Syntax.string_of_term (ML_Context.the_local_context ()) t) ());
   244 fun pth (HOLThm(ren,thm)) =
   245     let
   246         (*val _ = writeln "Renaming:"
   247         val _ = app (fn(v,w) => (prin v; writeln " -->"; prin w)) ren*)
   248         val _ = prth thm
   249     in
   250         ()
   251     end
   252 
   253 fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
   254 fun mk_proof p = Proof(Info{disk_info = Unsynchronized.ref NONE},p)
   255 fun content_of (Proof(_,p)) = p
   256 
   257 fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
   258     disk_info := SOME(thyname,thmname)
   259 
   260 structure Lib =
   261 struct
   262 
   263 fun assoc x =
   264     let
   265         fun F [] = raise PK("Lib.assoc","Not found")
   266           | F ((x',y)::rest) = if x = x'
   267                                then y
   268                                else F rest
   269     in
   270         F
   271     end
   272 infix mem;
   273 fun i mem L =
   274     let fun itr [] = false
   275           | itr (a::rst) = i=a orelse itr rst
   276     in itr L end;
   277 
   278 infix union;
   279 fun [] union S = S
   280   | S union [] = S
   281   | (a::rst) union S2 = rst union (insert (op =) a S2);
   282 
   283 fun implode_subst [] = []
   284   | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
   285   | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
   286 
   287 end
   288 open Lib
   289 
   290 structure Tag =
   291 struct
   292 val empty_tag = Tag []
   293 fun read name = Tag [name]
   294 fun merge (Tag tag1) (Tag tag2) = Tag (Lib.union(tag1,tag2))
   295 end
   296 
   297 (* Actual code. *)
   298 
   299 fun get_segment thyname l = (Lib.assoc "s" l
   300                              handle PK _ => thyname)
   301 val get_name : (string * string) list -> string = Lib.assoc "n"
   302 
   303 local
   304     open LazyScan
   305     infix 7 |-- --|
   306     infix 5 :-- -- ^^
   307     infix 3 >>
   308     infix 0 ||
   309 in
   310 exception XML of string
   311 
   312 datatype xml = Elem of string * (string * string) list * xml list
   313 datatype XMLtype = XMLty of xml | FullType of hol_type
   314 datatype XMLterm = XMLtm of xml | FullTerm of term
   315 
   316 fun pair x y = (x,y)
   317 
   318 fun scan_id toks =
   319     let
   320         val (x,toks2) = one Char.isAlpha toks
   321         val (xs,toks3) = any Char.isAlphaNum toks2
   322     in
   323         (String.implode (x::xs),toks3)
   324     end
   325 
   326 fun scan_string str c =
   327     let
   328         fun F [] toks = (c,toks)
   329           | F (c::cs) toks =
   330             case LazySeq.getItem toks of
   331                 SOME(c',toks') =>
   332                 if c = c'
   333                 then F cs toks'
   334                 else raise SyntaxError
   335               | NONE => raise SyntaxError
   336     in
   337         F (String.explode str)
   338     end
   339 
   340 local
   341     val scan_entity =
   342         (scan_string "amp;" #"&")
   343             || scan_string "quot;" #"\""
   344             || scan_string "gt;" #">"
   345             || scan_string "lt;" #"<"
   346             || scan_string "apos;" #"'"
   347 in
   348 fun scan_nonquote toks =
   349     case LazySeq.getItem toks of
   350         SOME (c,toks') =>
   351         (case c of
   352              #"\"" => raise SyntaxError
   353            | #"&" => scan_entity toks'
   354            | c => (c,toks'))
   355       | NONE => raise SyntaxError
   356 end
   357 
   358 val scan_string = $$ #"\"" |-- repeat scan_nonquote --| $$ #"\"" >>
   359                      String.implode
   360 
   361 val scan_attribute = scan_id -- $$ #"=" |-- scan_string
   362 
   363 val scan_start_of_tag = $$ #"<" |-- scan_id --
   364                            repeat ($$ #" " |-- scan_attribute)
   365 
   366 fun scan_end_of_tag toks = ($$ #"/" |-- $$ #">" |-- succeed []) toks
   367 
   368 val scan_end_tag = $$ #"<" |-- $$ #"/" |-- scan_id --| $$ #">"
   369 
   370 fun scan_children id = $$ #">" |-- repeat scan_tag -- scan_end_tag >>
   371                        (fn (chldr,id') => if id = id'
   372                                           then chldr
   373                                           else raise XML "Tag mismatch")
   374 and scan_tag toks =
   375     let
   376         val ((id,atts),toks2) = scan_start_of_tag toks
   377         val (chldr,toks3) = (scan_children id || scan_end_of_tag) toks2
   378     in
   379         (Elem (id,atts,chldr),toks3)
   380     end
   381 end
   382 
   383 val type_of = Term.type_of
   384 
   385 val propT = Type("prop",[])
   386 
   387 fun mk_defeq name rhs thy =
   388     let
   389         val ty = type_of rhs
   390     in
   391         Logic.mk_equals (Const(Sign.intern_const thy name,ty),rhs)
   392     end
   393 
   394 fun mk_teq name rhs thy =
   395     let
   396         val ty = type_of rhs
   397     in
   398         HOLogic.mk_eq (Const(Sign.intern_const thy name,ty),rhs)
   399     end
   400 
   401 fun intern_const_name thyname const thy =
   402     case get_hol4_const_mapping thyname const thy of
   403         SOME (_,cname,_) => cname
   404       | NONE => (case get_hol4_const_renaming thyname const thy of
   405                      SOME cname => Sign.intern_const thy (thyname ^ "." ^ cname)
   406                    | NONE => Sign.intern_const thy (thyname ^ "." ^ const))
   407 
   408 fun intern_type_name thyname const thy =
   409     case get_hol4_type_mapping thyname const thy of
   410         SOME (_,cname) => cname
   411       | NONE => Sign.intern_const thy (thyname ^ "." ^ const)
   412 
   413 fun mk_vartype name = TFree(name,["HOL.type"])
   414 fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args)
   415 
   416 val mk_var = Free
   417 
   418 fun mk_thy_const thy Thy Nam Ty = Const(intern_const_name Thy Nam thy,Ty)
   419 
   420 local
   421   fun get_const sg thyname name =
   422     (case Sign.const_type sg name of
   423       SOME ty => Const (name, ty)
   424     | NONE => raise ERR "get_type" (name ^ ": No such constant"))
   425 in
   426 fun prim_mk_const thy Thy Nam =
   427     let
   428       val name = intern_const_name Thy Nam thy
   429       val cmaps = HOL4ConstMaps.get thy
   430     in
   431       case StringPair.lookup cmaps (Thy,Nam) of
   432         SOME(_,_,SOME ty) => Const(name,ty)
   433       | _ => get_const thy Thy name
   434     end
   435 end
   436 
   437 fun mk_comb(f,a) = f $ a
   438 
   439 (* Needed for HOL Light *)
   440 fun protect_tyvarname s =
   441     let
   442         fun no_quest s =
   443             if Char.contains s #"?"
   444             then String.translate (fn #"?" => "q_" | c => Char.toString c) s
   445             else s
   446         fun beg_prime s =
   447             if String.isPrefix "'" s
   448             then s
   449             else "'" ^ s
   450     in
   451         s |> no_quest |> beg_prime
   452     end
   453 
   454 val protected_varnames = Unsynchronized.ref (Symtab.empty:string Symtab.table)
   455 val invented_isavar = Unsynchronized.ref 0
   456 
   457 fun innocent_varname s = Lexicon.is_identifier s andalso not (String.isPrefix "u_" s)
   458 
   459 fun valid_boundvarname s =
   460   can (fn () => Syntax.read_term_global @{theory Main} ("SOME "^s^". True")) ();
   461 
   462 fun valid_varname s =
   463   can (fn () => Syntax.read_term_global @{theory Main} s) ();
   464 
   465 fun protect_varname s =
   466     if innocent_varname s andalso valid_varname s then s else
   467     case Symtab.lookup (!protected_varnames) s of
   468       SOME t => t
   469     | NONE =>
   470       let
   471           val _ = Unsynchronized.inc invented_isavar
   472           val t = "u_" ^ string_of_int (!invented_isavar)
   473           val _ = ImportRecorder.protect_varname s t
   474           val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
   475       in
   476           t
   477       end
   478 
   479 exception REPLAY_PROTECT_VARNAME of string*string*string
   480 
   481 fun replay_protect_varname s t =
   482         case Symtab.lookup (!protected_varnames) s of
   483           SOME t' => raise REPLAY_PROTECT_VARNAME (s, t, t')
   484         | NONE =>
   485           let
   486               val _ = Unsynchronized.inc invented_isavar
   487               val t = "u_" ^ string_of_int (!invented_isavar)
   488               val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
   489           in
   490               ()
   491           end
   492 
   493 fun protect_boundvarname s = if innocent_varname s andalso valid_boundvarname s then s else "u"
   494 
   495 fun mk_lambda (v as Free (x, T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
   496   | mk_lambda (v as Var ((x, _), T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
   497   | mk_lambda v t = raise TERM ("lambda", [v, t]);
   498 
   499 fun replacestr x y s =
   500 let
   501   val xl = raw_explode x
   502   val yl = raw_explode y
   503   fun isprefix [] ys = true
   504     | isprefix (x::xs) (y::ys) = if x = y then isprefix xs ys else false
   505     | isprefix _ _ = false
   506   fun isp s = isprefix xl s
   507   fun chg s = yl@(List.drop (s, List.length xl))
   508   fun r [] = []
   509     | r (S as (s::ss)) = if isp S then r (chg S) else s::(r ss)
   510 in
   511   implode(r (raw_explode s))
   512 end
   513 
   514 fun protect_factname s = replacestr "." "_dot_" s
   515 fun unprotect_factname s = replacestr "_dot_" "." s
   516 
   517 val ty_num_prefix = "N_"
   518 
   519 fun startsWithDigit s = Char.isDigit (hd (String.explode s))
   520 
   521 fun protect_tyname tyn =
   522   let
   523     val tyn' =
   524       if String.isPrefix ty_num_prefix tyn then raise (ERR "protect_ty_name" ("type name '"^tyn^"' is reserved")) else
   525       (if startsWithDigit tyn then ty_num_prefix^tyn else tyn)
   526   in
   527     tyn'
   528   end
   529 
   530 fun protect_constname tcn = tcn
   531  (* if tcn = ".." then "dotdot"
   532   else if tcn = "==" then "eqeq"
   533   else tcn*)
   534 
   535 structure TypeNet =
   536 struct
   537 
   538 fun get_type_from_index thy thyname types is =
   539     case Int.fromString is of
   540         SOME i => (case Array.sub(types,i) of
   541                        FullType ty => ty
   542                      | XMLty xty =>
   543                        let
   544                            val ty = get_type_from_xml thy thyname types xty
   545                            val _  = Array.update(types,i,FullType ty)
   546                        in
   547                            ty
   548                        end)
   549       | NONE => raise ERR "get_type_from_index" "Bad index"
   550 and get_type_from_xml thy thyname types =
   551     let
   552         fun gtfx (Elem("tyi",[("i",iS)],[])) =
   553                  get_type_from_index thy thyname types iS
   554           | gtfx (Elem("tyc",atts,[])) =
   555             mk_thy_type thy
   556                         (get_segment thyname atts)
   557                         (protect_tyname (get_name atts))
   558                         []
   559           | gtfx (Elem("tyv",[("n",s)],[])) = mk_vartype (protect_tyvarname s)
   560           | gtfx (Elem("tya",[],(Elem("tyc",atts,[]))::tys)) =
   561             mk_thy_type thy
   562                         (get_segment thyname atts)
   563                         (protect_tyname (get_name atts))
   564                         (map gtfx tys)
   565           | gtfx _ = raise ERR "get_type" "Bad type"
   566     in
   567         gtfx
   568     end
   569 
   570 fun input_types thyname (Elem("tylist",[("i",i)],xtys)) =
   571     let
   572         val types = Array.array(the (Int.fromString i),XMLty (Elem("",[],[])))
   573         fun IT _ [] = ()
   574           | IT n (xty::xtys) =
   575             (Array.update(types,n,XMLty xty);
   576              IT (n+1) xtys)
   577         val _ = IT 0 xtys
   578     in
   579         types
   580     end
   581   | input_types _ _ = raise ERR "input_types" "Bad type list"
   582 end
   583 
   584 structure TermNet =
   585 struct
   586 
   587 fun get_term_from_index thy thyname types terms is =
   588     case Int.fromString is of
   589         SOME i => (case Array.sub(terms,i) of
   590                        FullTerm tm => tm
   591                      | XMLtm xtm =>
   592                        let
   593                            val tm = get_term_from_xml thy thyname types terms xtm
   594                            val _  = Array.update(terms,i,FullTerm tm)
   595                        in
   596                            tm
   597                        end)
   598       | NONE => raise ERR "get_term_from_index" "Bad index"
   599 and get_term_from_xml thy thyname types terms =
   600     let
   601         fun get_type [] = NONE
   602           | get_type [ty] = SOME (TypeNet.get_type_from_xml thy thyname types ty)
   603           | get_type _ = raise ERR "get_term" "Bad type"
   604 
   605         fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
   606             mk_var(protect_varname name,TypeNet.get_type_from_index thy thyname types tyi)
   607           | gtfx (Elem("tmc",atts,[])) =
   608             let
   609                 val segment = get_segment thyname atts
   610                 val name = protect_constname(get_name atts)
   611             in
   612                 mk_thy_const thy segment name (TypeNet.get_type_from_index thy thyname types (Lib.assoc "t" atts))
   613                 handle PK _ => prim_mk_const thy segment name
   614             end
   615           | gtfx (Elem("tma",[("f",tmf),("a",tma)],[])) =
   616             let
   617                 val f = get_term_from_index thy thyname types terms tmf
   618                 val a = get_term_from_index thy thyname types terms tma
   619             in
   620                 mk_comb(f,a)
   621             end
   622           | gtfx (Elem("tml",[("x",tmx),("a",tma)],[])) =
   623             let
   624                 val x = get_term_from_index thy thyname types terms tmx
   625                 val a = get_term_from_index thy thyname types terms tma
   626             in
   627                 mk_lambda x a
   628             end
   629           | gtfx (Elem("tmi",[("i",iS)],[])) =
   630             get_term_from_index thy thyname types terms iS
   631           | gtfx (Elem(tag,_,_)) =
   632             raise ERR "get_term" ("Not a term: "^tag)
   633     in
   634         gtfx
   635     end
   636 
   637 fun input_terms thyname types (Elem("tmlist",[("i",i)],xtms)) =
   638     let
   639         val terms = Array.array(the (Int.fromString i), XMLtm (Elem("",[],[])))
   640 
   641         fun IT _ [] = ()
   642           | IT n (xtm::xtms) =
   643             (Array.update(terms,n,XMLtm xtm);
   644              IT (n+1) xtms)
   645         val _ = IT 0 xtms
   646     in
   647         terms
   648     end
   649   | input_terms _ _ _ = raise ERR "input_terms" "Bad term list"
   650 end
   651 
   652 fun get_proof_dir (thyname:string) thy =
   653     let
   654         val import_segment =
   655             case get_segment2 thyname thy of
   656                 SOME seg => seg
   657               | NONE => get_import_segment thy
   658         val path = space_explode ":" (getenv "HOL4_PROOFS")
   659         fun find [] = NONE
   660           | find (p::ps) =
   661             (let
   662                  val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
   663              in
   664                  if OS.FileSys.isDir dir
   665                  then SOME dir
   666                  else find ps
   667              end) handle OS.SysErr _ => find ps
   668     in
   669         Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
   670     end
   671 
   672 fun proof_file_name thyname thmname thy =
   673     let
   674         val path = case get_proof_dir thyname thy of
   675                        SOME p => p
   676                      | NONE => error "Cannot find proof files"
   677         val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
   678     in
   679         OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = (unprotect_factname thmname), ext = SOME "prf"}}
   680     end
   681 
   682 fun xml_to_proof thyname types terms prf thy =
   683     let
   684         val xml_to_hol_type = TypeNet.get_type_from_xml thy thyname types
   685         val xml_to_term = TermNet.get_term_from_xml thy thyname types terms
   686 
   687         fun index_to_term is =
   688             TermNet.get_term_from_index thy thyname types terms is
   689 
   690         fun x2p (Elem("prefl",[("i",is)],[])) = mk_proof (PRefl (index_to_term is))
   691           | x2p (Elem("pinstt",[],p::lambda)) =
   692             let
   693                 val p = x2p p
   694                 val lambda = implode_subst (map xml_to_hol_type lambda)
   695             in
   696                 mk_proof (PInstT(p,lambda))
   697             end
   698           | x2p (Elem("psubst",[("i",is)],prf::prfs)) =
   699             let
   700                 val tm = index_to_term is
   701                 val prf = x2p prf
   702                 val prfs = map x2p prfs
   703             in
   704                 mk_proof (PSubst(prfs,tm,prf))
   705             end
   706           | x2p (Elem("pabs",[("i",is)],[prf])) =
   707             let
   708                 val p = x2p prf
   709                 val t = index_to_term is
   710             in
   711                 mk_proof (PAbs (p,t))
   712             end
   713           | x2p (Elem("pdisch",[("i",is)],[prf])) =
   714             let
   715                 val p = x2p prf
   716                 val t = index_to_term is
   717             in
   718                 mk_proof (PDisch (p,t))
   719             end
   720           | x2p (Elem("pmp",[],[prf1,prf2])) =
   721             let
   722                 val p1 = x2p prf1
   723                 val p2 = x2p prf2
   724             in
   725                 mk_proof (PMp(p1,p2))
   726             end
   727           | x2p (Elem("phyp",[("i",is)],[])) = mk_proof (PHyp (index_to_term is))
   728           | x2p (Elem("paxiom",[("n",n),("i",is)],[])) =
   729             mk_proof (PAxm(n,index_to_term is))
   730           | x2p (Elem("pfact",atts,[])) =
   731             let
   732                 val thyname = get_segment thyname atts
   733                 val thmname = protect_factname (get_name atts)
   734                 val p = mk_proof PDisk
   735                 val _  = set_disk_info_of p thyname thmname
   736             in
   737                 p
   738             end
   739           | x2p (Elem("pdef",[("s",seg),("n",name),("i",is)],[])) =
   740             mk_proof (PDef(seg,protect_constname name,index_to_term is))
   741           | x2p (Elem("ptmspec",[("s",seg)],p::names)) =
   742             let
   743                 val names = map (fn Elem("name",[("n",name)],[]) => name
   744                                   | _ => raise ERR "x2p" "Bad proof (ptmspec)") names
   745             in
   746                 mk_proof (PTmSpec(seg,names,x2p p))
   747             end
   748           | x2p (Elem("ptyintro",[("s",seg),("n",name),("a",abs_name),("r",rep_name)],[xP,xt,p])) =
   749             let
   750                 val P = xml_to_term xP
   751                 val t = xml_to_term xt
   752             in
   753                 mk_proof (PTyIntro(seg,protect_tyname name,protect_constname abs_name,protect_constname rep_name,P,t,x2p p))
   754             end
   755           | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
   756             mk_proof (PTyDef(seg,protect_tyname name,x2p p))
   757           | x2p (xml as Elem("poracle",[],chldr)) =
   758             let
   759                 val (oracles,terms) = List.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
   760                 val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | xml => raise ERR "x2p" "bad oracle") oracles
   761                 val (c,asl) = case terms of
   762                                   [] => raise ERR "x2p" "Bad oracle description"
   763                                 | (hd::tl) => (hd,tl)
   764                 val tg = fold_rev (Tag.merge o Tag.read) ors Tag.empty_tag
   765             in
   766                 mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c))
   767             end
   768           | x2p (Elem("pspec",[("i",is)],[prf])) =
   769             let
   770                 val p = x2p prf
   771                 val tm = index_to_term is
   772             in
   773                 mk_proof (PSpec(p,tm))
   774             end
   775           | x2p (Elem("pinst",[],p::theta)) =
   776             let
   777                 val p = x2p p
   778                 val theta = implode_subst (map xml_to_term theta)
   779             in
   780                 mk_proof (PInst(p,theta))
   781             end
   782           | x2p (Elem("pgen",[("i",is)],[prf])) =
   783             let
   784                 val p = x2p prf
   785                 val tm = index_to_term is
   786             in
   787                 mk_proof (PGen(p,tm))
   788             end
   789           | x2p (Elem("pgenabs",[],prf::tms)) =
   790             let
   791                 val p = x2p prf
   792                 val tml = map xml_to_term tms
   793             in
   794                 mk_proof (PGenAbs(p,NONE,tml))
   795             end
   796           | x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
   797             let
   798                 val p = x2p prf
   799                 val tml = map xml_to_term tms
   800             in
   801                 mk_proof (PGenAbs(p,SOME (index_to_term is),tml))
   802             end
   803           | x2p (Elem("pimpas",[],[prf1,prf2])) =
   804             let
   805                 val p1 = x2p prf1
   806                 val p2 = x2p prf2
   807             in
   808                 mk_proof (PImpAS(p1,p2))
   809             end
   810           | x2p (Elem("psym",[],[prf])) =
   811             let
   812                 val p = x2p prf
   813             in
   814                 mk_proof (PSym p)
   815             end
   816           | x2p (Elem("ptrans",[],[prf1,prf2])) =
   817             let
   818                 val p1 = x2p prf1
   819                 val p2 = x2p prf2
   820             in
   821                 mk_proof (PTrans(p1,p2))
   822             end
   823           | x2p (Elem("pcomb",[],[prf1,prf2])) =
   824             let
   825                 val p1 = x2p prf1
   826                 val p2 = x2p prf2
   827             in
   828                 mk_proof (PComb(p1,p2))
   829             end
   830           | x2p (Elem("peqmp",[],[prf1,prf2])) =
   831             let
   832                 val p1 = x2p prf1
   833                 val p2 = x2p prf2
   834             in
   835                 mk_proof (PEqMp(p1,p2))
   836             end
   837           | x2p (Elem("peqimp",[],[prf])) =
   838             let
   839                 val p = x2p prf
   840             in
   841                 mk_proof (PEqImp p)
   842             end
   843           | x2p (Elem("pexists",[("e",ise),("w",isw)],[prf])) =
   844             let
   845                 val p = x2p prf
   846                 val ex = index_to_term ise
   847                 val w = index_to_term isw
   848             in
   849                 mk_proof (PExists(p,ex,w))
   850             end
   851           | x2p (Elem("pchoose",[("i",is)],[prf1,prf2])) =
   852             let
   853                 val v  = index_to_term is
   854                 val p1 = x2p prf1
   855                 val p2 = x2p prf2
   856             in
   857                 mk_proof (PChoose(v,p1,p2))
   858             end
   859           | x2p (Elem("pconj",[],[prf1,prf2])) =
   860             let
   861                 val p1 = x2p prf1
   862                 val p2 = x2p prf2
   863             in
   864                 mk_proof (PConj(p1,p2))
   865             end
   866           | x2p (Elem("pconjunct1",[],[prf])) =
   867             let
   868                 val p = x2p prf
   869             in
   870                 mk_proof (PConjunct1 p)
   871             end
   872           | x2p (Elem("pconjunct2",[],[prf])) =
   873             let
   874                 val p = x2p prf
   875             in
   876                 mk_proof (PConjunct2 p)
   877             end
   878           | x2p (Elem("pdisj1",[("i",is)],[prf])) =
   879             let
   880                 val p = x2p prf
   881                 val t = index_to_term is
   882             in
   883                 mk_proof (PDisj1 (p,t))
   884             end
   885           | x2p (Elem("pdisj2",[("i",is)],[prf])) =
   886             let
   887                 val p = x2p prf
   888                 val t = index_to_term is
   889             in
   890                 mk_proof (PDisj2 (p,t))
   891             end
   892           | x2p (Elem("pdisjcases",[],[prf1,prf2,prf3])) =
   893             let
   894                 val p1 = x2p prf1
   895                 val p2 = x2p prf2
   896                 val p3 = x2p prf3
   897             in
   898                 mk_proof (PDisjCases(p1,p2,p3))
   899             end
   900           | x2p (Elem("pnoti",[],[prf])) =
   901             let
   902                 val p = x2p prf
   903             in
   904                 mk_proof (PNotI p)
   905             end
   906           | x2p (Elem("pnote",[],[prf])) =
   907             let
   908                 val p = x2p prf
   909             in
   910                 mk_proof (PNotE p)
   911             end
   912           | x2p (Elem("pcontr",[("i",is)],[prf])) =
   913             let
   914                 val p = x2p prf
   915                 val t = index_to_term is
   916             in
   917                 mk_proof (PContr (p,t))
   918             end
   919           | x2p xml = raise ERR "x2p" "Bad proof"
   920     in
   921         x2p prf
   922     end
   923 
   924 fun import_proof_concl thyname thmname thy =
   925     let
   926         val is = TextIO.openIn(proof_file_name thyname thmname thy)
   927         val (proof_xml,_) = scan_tag (LazySeq.of_instream is)
   928         val _ = TextIO.closeIn is
   929     in
   930         case proof_xml of
   931             Elem("proof",[],xtypes::xterms::prf::rest) =>
   932             let
   933                 val types = TypeNet.input_types thyname xtypes
   934                 val terms = TermNet.input_terms thyname types xterms
   935                 fun f xtm thy = TermNet.get_term_from_xml thy thyname types terms xtm
   936             in
   937                 case rest of
   938                     [] => NONE
   939                   | [xtm] => SOME (f xtm)
   940                   | _ => raise ERR "import_proof" "Bad argument list"
   941             end
   942           | _ => raise ERR "import_proof" "Bad proof"
   943     end
   944 
   945 fun import_proof thyname thmname thy =
   946     let
   947         val is = TextIO.openIn(proof_file_name thyname thmname thy)
   948         val (proof_xml,_) = scan_tag (LazySeq.of_instream is)
   949         val _ = TextIO.closeIn is
   950     in
   951         case proof_xml of
   952             Elem("proof",[],xtypes::xterms::prf::rest) =>
   953             let
   954                 val types = TypeNet.input_types thyname xtypes
   955                 val terms = TermNet.input_terms thyname types xterms
   956             in
   957                 (case rest of
   958                      [] => NONE
   959                    | [xtm] => SOME (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
   960                    | _ => raise ERR "import_proof" "Bad argument list",
   961                  xml_to_proof thyname types terms prf)
   962             end
   963           | _ => raise ERR "import_proof" "Bad proof"
   964     end
   965 
   966 fun uniq_compose m th i st =
   967     let
   968         val res = Thm.bicompose false (false,th,m) i st
   969     in
   970         case Seq.pull res of
   971             SOME (th,rest) => (case Seq.pull rest of
   972                                    SOME _ => raise ERR "uniq_compose" "Not unique!"
   973                                  | NONE => th)
   974           | NONE => raise ERR "uniq_compose" "No result"
   975     end
   976 
   977 val reflexivity_thm = @{thm refl}
   978 val substitution_thm = @{thm subst}
   979 val mp_thm = @{thm mp}
   980 val imp_antisym_thm = @{thm light_imp_as}
   981 val disch_thm = @{thm impI}
   982 val ccontr_thm = @{thm ccontr}
   983 
   984 val meta_eq_to_obj_eq_thm = @{thm meta_eq_to_obj_eq}
   985 
   986 val gen_thm = @{thm HOLallI}
   987 val choose_thm = @{thm exE}
   988 val exists_thm = @{thm exI}
   989 val conj_thm = @{thm conjI}
   990 val conjunct1_thm = @{thm conjunct1}
   991 val conjunct2_thm = @{thm conjunct2}
   992 val spec_thm = @{thm spec}
   993 val disj_cases_thm = @{thm disjE}
   994 val disj1_thm = @{thm disjI1}
   995 val disj2_thm = @{thm disjI2}
   996 
   997 local
   998     val th = @{thm not_def}
   999     val thy = theory_of_thm th
  1000     val pp = Thm.reflexive (cterm_of thy (Const(@{const_name Trueprop},HOLogic.boolT-->propT)))
  1001 in
  1002 val not_elim_thm = Thm.combination pp th
  1003 end
  1004 
  1005 val not_intro_thm = Thm.symmetric not_elim_thm
  1006 val abs_thm = @{thm ext}
  1007 val trans_thm = @{thm trans}
  1008 val symmetry_thm = @{thm sym}
  1009 val transitivity_thm = @{thm trans}
  1010 val eqmp_thm = @{thm iffD1}
  1011 val eqimp_thm = @{thm HOL4Setup.eq_imp}
  1012 val comb_thm = @{thm cong}
  1013 
  1014 (* Beta-eta normalizes a theorem (only the conclusion, not the *
  1015 hypotheses!)  *)
  1016 
  1017 fun beta_eta_thm th =
  1018     let
  1019         val th1 = Thm.equal_elim (Thm.beta_conversion true (cprop_of th))  th
  1020         val th2 = Thm.equal_elim (Thm.eta_conversion       (cprop_of th1)) th1
  1021     in
  1022         th2
  1023     end
  1024 
  1025 fun implies_elim_all th =
  1026     Library.foldl (fn (th,p) => Thm.implies_elim th (Thm.assume p)) (th,cprems_of th)
  1027 
  1028 fun norm_hyps th =
  1029     th |> beta_eta_thm
  1030        |> implies_elim_all
  1031        |> implies_intr_hyps
  1032 
  1033 fun mk_GEN v th sg =
  1034     let
  1035         val c = HOLogic.dest_Trueprop (concl_of th)
  1036         val cv = cterm_of sg v
  1037         val lc = Term.lambda v c
  1038         val clc = Thm.cterm_of sg lc
  1039         val cvty = ctyp_of_term cv
  1040         val th1 = implies_elim_all th
  1041         val th2 = beta_eta_thm (Thm.forall_intr cv th1)
  1042         val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm))
  1043         val c = prop_of th3
  1044         val vname = fst(dest_Free v)
  1045         val (cold,cnew) = case c of
  1046                               tpc $ (Const(@{const_name All},allT) $ Abs(oldname,ty,body)) =>
  1047                               (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
  1048                             | tpc $ (Const(@{const_name All},allT) $ rest) => (tpc,tpc)
  1049                             | _ => raise ERR "mk_GEN" "Unknown conclusion"
  1050         val th4 = Thm.rename_boundvars cold cnew th3
  1051         val res = implies_intr_hyps th4
  1052     in
  1053         res
  1054     end
  1055 
  1056 fun rearrange sg tm th =
  1057     let
  1058         val tm' = Envir.beta_eta_contract tm
  1059         fun find []      n = Thm.permute_prems 0 1 (Thm.implies_intr (Thm.cterm_of sg tm) th)
  1060           | find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
  1061                              then Thm.permute_prems n 1 th
  1062                              else find ps (n+1)
  1063     in
  1064         find (prems_of th) 0
  1065     end
  1066 
  1067 fun zip (x::xs) (y::ys) = (x,y)::(zip xs ys)
  1068   | zip [] [] = []
  1069   | zip _ _ = raise ERR "zip" "arguments not of same length"
  1070 
  1071 fun mk_INST dom rng th =
  1072     th |> forall_intr_list dom
  1073        |> forall_elim_list rng
  1074 
  1075 val collect_vars =
  1076     let
  1077         fun F vars (Bound _) = vars
  1078           | F vars (tm as Free _) =
  1079             if member (op =) vars tm
  1080             then vars
  1081             else (tm::vars)
  1082           | F vars (Const _) = vars
  1083           | F vars (tm1 $ tm2) = F (F vars tm1) tm2
  1084           | F vars (Abs(_,_,body)) = F vars body
  1085           | F vars (Var _) = raise ERR "collect_vars" "Schematic variable found"
  1086     in
  1087         F []
  1088     end
  1089 
  1090 (* Code for disambiguating variablenames (wrt. types) *)
  1091 
  1092 val disamb_info_empty = {vars=[],rens=[]}
  1093 
  1094 fun rens_of {vars,rens} = rens
  1095 
  1096 fun name_of_var (Free(vname,_)) = vname
  1097   | name_of_var _ = raise ERR "name_of_var" "Not a variable"
  1098 
  1099 fun disamb_term_from info tm = (info, tm)
  1100 
  1101 fun swap (x,y) = (y,x)
  1102 
  1103 fun has_ren (HOLThm _) = false
  1104 
  1105 fun prinfo {vars,rens} = (writeln "Vars:";
  1106                           app prin vars;
  1107                           writeln "Renaming:";
  1108                           app (fn(x,y)=>(prin x; writeln " -->"; prin y)) rens)
  1109 
  1110 fun disamb_thm_from info (HOLThm (_,thm)) = (info, thm)
  1111 
  1112 fun disamb_terms_from info tms = (info, tms)
  1113 
  1114 fun disamb_thms_from info hthms = (info, map hthm2thm hthms)
  1115 
  1116 fun disamb_term tm   = disamb_term_from disamb_info_empty tm
  1117 fun disamb_terms tms = disamb_terms_from disamb_info_empty tms
  1118 fun disamb_thm thm   = disamb_thm_from disamb_info_empty thm
  1119 fun disamb_thms thms = disamb_thms_from disamb_info_empty thms
  1120 
  1121 fun norm_hthm sg (hth as HOLThm _) = hth
  1122 
  1123 (* End of disambiguating code *)
  1124 
  1125 fun disambiguate_frees thm =
  1126     let
  1127       fun ERR s = error ("Drule.disambiguate_frees: "^s)
  1128       val ct = cprop_of thm
  1129       val t = term_of ct
  1130       val thy = theory_of_cterm ct
  1131       val frees = OldTerm.term_frees t
  1132       val freenames = Term.add_free_names t []
  1133       val is_old_name = member (op =) freenames
  1134       fun name_of (Free (n, _)) = n
  1135         | name_of _ = ERR "name_of"
  1136       fun new_name' bump map n =
  1137           let val n' = n^bump in
  1138             if is_old_name n' orelse Symtab.lookup map n' <> NONE then
  1139               new_name' (Symbol.bump_string bump) map n
  1140             else
  1141               n'
  1142           end
  1143       val new_name = new_name' "a"
  1144       fun replace_name n' (Free (n, t)) = Free (n', t)
  1145         | replace_name n' _ = ERR "replace_name"
  1146       (* map: old or fresh name -> old free,
  1147          invmap: old free which has fresh name assigned to it -> fresh name *)
  1148       fun dis v (mapping as (map,invmap)) =
  1149           let val n = name_of v in
  1150             case Symtab.lookup map n of
  1151               NONE => (Symtab.update (n, v) map, invmap)
  1152             | SOME v' =>
  1153               if v=v' then
  1154                 mapping
  1155               else
  1156                 let val n' = new_name map n in
  1157                   (Symtab.update (n', v) map,
  1158                    Termtab.update (v, n') invmap)
  1159                 end
  1160           end
  1161     in
  1162       if (length freenames = length frees) then
  1163         thm
  1164       else
  1165         let
  1166           val (_, invmap) =
  1167               fold dis frees (Symtab.empty, Termtab.empty)
  1168           fun make_subst (oldfree, newname) (intros, elims) =
  1169               (cterm_of thy oldfree :: intros,
  1170                cterm_of thy (replace_name newname oldfree) :: elims)
  1171           val (intros, elims) = fold make_subst (Termtab.dest invmap) ([], [])
  1172         in
  1173           forall_elim_list elims (forall_intr_list intros thm)
  1174         end
  1175     end
  1176 
  1177 val debug = Unsynchronized.ref false
  1178 
  1179 fun if_debug f x = if !debug then f x else ()
  1180 val message = if_debug writeln
  1181 
  1182 val conjE_helper = Thm.permute_prems 0 1 conjE
  1183 
  1184 fun get_hol4_thm thyname thmname thy =
  1185     case get_hol4_theorem thyname thmname thy of
  1186         SOME hth => SOME (HOLThm hth)
  1187       | NONE =>
  1188         let
  1189             val pending = HOL4Pending.get thy
  1190         in
  1191             case StringPair.lookup pending (thyname,thmname) of
  1192                 SOME hth => SOME (HOLThm hth)
  1193               | NONE => NONE
  1194         end
  1195 
  1196 fun non_trivial_term_consts t = fold_aterms
  1197   (fn Const (c, _) =>
  1198       if c = @{const_name Trueprop} orelse c = @{const_name All}
  1199         orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
  1200       then I else insert (op =) c
  1201     | _ => I) t [];
  1202 
  1203 fun match_consts t (* th *) =
  1204     let
  1205         fun add_consts (Const (c, _), cs) =
  1206             (case c of
  1207                  @{const_name HOL.eq} => insert (op =) "==" cs
  1208                | @{const_name HOL.implies} => insert (op =) "==>" cs
  1209                | @{const_name All} => cs
  1210                | "all" => cs
  1211                | @{const_name HOL.conj} => cs
  1212                | @{const_name Trueprop} => cs
  1213                | _ => insert (op =) c cs)
  1214           | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
  1215           | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
  1216           | add_consts (_, cs) = cs
  1217         val t_consts = add_consts(t,[])
  1218     in
  1219         fn th => eq_set (op =) (t_consts, add_consts (prop_of th, []))
  1220     end
  1221 
  1222 fun split_name str =
  1223     let
  1224         val sub = Substring.full str
  1225         val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
  1226         val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
  1227     in
  1228         if not (idx = "") andalso u = "_"
  1229         then SOME (newstr, the (Int.fromString idx))
  1230         else NONE
  1231     end
  1232     handle _ => NONE  (* FIXME avoid handle _ *)
  1233 
  1234 fun rewrite_hol4_term t thy =
  1235     let
  1236         val hol4rews1 = map (Thm.transfer thy) (HOL4Rewrites.get thy)
  1237         val hol4ss = Simplifier.global_context thy empty_ss
  1238           setmksimps (K single) addsimps hol4rews1
  1239     in
  1240         Thm.transfer thy (Simplifier.full_rewrite hol4ss (cterm_of thy t))
  1241     end
  1242 
  1243 fun get_isabelle_thm thyname thmname hol4conc thy =
  1244     let
  1245         val (info,hol4conc') = disamb_term hol4conc
  1246         val i2h_conc = Thm.symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
  1247         val isaconc =
  1248             case concl_of i2h_conc of
  1249                 Const("==",_) $ lhs $ _ => lhs
  1250               | _ => error "get_isabelle_thm" "Bad rewrite rule"
  1251         val _ = (message "Original conclusion:";
  1252                  if_debug prin hol4conc';
  1253                  message "Modified conclusion:";
  1254                  if_debug prin isaconc)
  1255 
  1256         fun mk_res th = HOLThm (rens_of info, Thm.equal_elim i2h_conc th)
  1257     in
  1258         case get_hol4_mapping thyname thmname thy of
  1259             SOME (SOME thmname) =>
  1260             let
  1261                 val th1 = (SOME (Global_Theory.get_thm thy thmname)
  1262                            handle ERROR _ =>
  1263                                   (case split_name thmname of
  1264                                        SOME (listname,idx) => (SOME (List.nth(Global_Theory.get_thms thy listname,idx-1))
  1265                                                                handle _ => NONE)  (* FIXME avoid handle _ *)
  1266                                      | NONE => NONE))
  1267             in
  1268                 case th1 of
  1269                     SOME th2 =>
  1270                     (case Shuffler.set_prop thy isaconc [(thmname,th2)] of
  1271                          SOME (_,th) => (message "YES";(thy, SOME (mk_res th)))
  1272                        | NONE => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
  1273                   | NONE => (message "NO1";error "get_isabelle_thm" "Bad mapping")
  1274             end
  1275           | SOME NONE => error ("Trying to access ignored theorem " ^ thmname)
  1276           | NONE =>
  1277             let
  1278                 val _ = (message "Looking for conclusion:";
  1279                          if_debug prin isaconc)
  1280                 val cs = non_trivial_term_consts isaconc;
  1281                 val _ = (message "Looking for consts:";
  1282                          message (commas cs))
  1283                 val pot_thms = Shuffler.find_potential thy isaconc
  1284                 val _ = message (string_of_int (length pot_thms) ^ " potential theorems")
  1285             in
  1286                 case Shuffler.set_prop thy isaconc pot_thms of
  1287                     SOME (isaname,th) =>
  1288                     let
  1289                         val hth as HOLThm args = mk_res th
  1290                         val thy' =  thy |> add_hol4_theorem thyname thmname args
  1291                                         |> add_hol4_mapping thyname thmname isaname
  1292                         val _ = ImportRecorder.add_hol_theorem thyname thmname (snd args)
  1293                         val _ = ImportRecorder.add_hol_mapping thyname thmname isaname
  1294                     in
  1295                         (thy',SOME hth)
  1296                     end
  1297                   | NONE => (thy,NONE)
  1298             end
  1299     end
  1300     handle e =>
  1301       if Exn.is_interrupt e then reraise e
  1302       else
  1303         (if_debug (fn () =>
  1304             writeln ("Exception in get_isabelle_thm:\n" ^ ML_Compiler.exn_message e)) ();
  1305           (thy,NONE))
  1306 
  1307 fun get_isabelle_thm_and_warn thyname thmname hol4conc thy =
  1308   let
  1309     val (a, b) = get_isabelle_thm thyname thmname hol4conc thy
  1310     fun warn () =
  1311         let
  1312             val (info,hol4conc') = disamb_term hol4conc
  1313             val i2h_conc = Thm.symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
  1314         in
  1315             case concl_of i2h_conc of
  1316                 Const("==",_) $ lhs $ _ =>
  1317                 (warning ("Failed lookup of theorem '"^thmname^"':");
  1318                  writeln "Original conclusion:";
  1319                  prin hol4conc';
  1320                  writeln "Modified conclusion:";
  1321                  prin lhs)
  1322               | _ => ()
  1323         end
  1324   in
  1325       case b of
  1326           NONE => (warn () handle _ => (); (a,b)) (* FIXME avoid handle _ *)
  1327         | _ => (a, b)
  1328   end
  1329 
  1330 fun get_thm thyname thmname thy =
  1331     case get_hol4_thm thyname thmname thy of
  1332         SOME hth => (thy,SOME hth)
  1333       | NONE => ((case import_proof_concl thyname thmname thy of
  1334                       SOME f => get_isabelle_thm_and_warn thyname thmname (f thy) thy
  1335                     | NONE => (message "No conclusion"; (thy,NONE)))
  1336                  handle e as IO.Io _ => (message "IO exception"; (thy,NONE))
  1337                       | e as PK _ => (message "PK exception"; (thy,NONE)))
  1338 
  1339 fun rename_const thyname thy name =
  1340     case get_hol4_const_renaming thyname name thy of
  1341         SOME cname => cname
  1342       | NONE => name
  1343 
  1344 fun get_def thyname constname rhs thy =
  1345     let
  1346         val constname = rename_const thyname thy constname
  1347         val (thmname,thy') = get_defname thyname constname thy
  1348         val _ = message ("Looking for definition " ^ thyname ^ "." ^ thmname)
  1349     in
  1350         get_isabelle_thm_and_warn thyname thmname (mk_teq (thyname ^ "." ^ constname) rhs thy') thy'
  1351     end
  1352 
  1353 fun get_axiom thyname axname thy =
  1354     case get_thm thyname axname thy of
  1355         arg as (_,SOME _) => arg
  1356       | _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
  1357 
  1358 fun intern_store_thm gen_output thyname thmname hth thy =
  1359     let
  1360         val (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
  1361         val rew = rewrite_hol4_term (concl_of th) thy
  1362         val th  = Thm.equal_elim rew th
  1363         val thy' = add_hol4_pending thyname thmname args thy
  1364         val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
  1365         val th = disambiguate_frees th
  1366         val thy2 =
  1367           if gen_output
  1368           then
  1369             add_dump ("lemma " ^ (quotename thmname) ^ ": " ^
  1370                 (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^ "\n  by (import " ^
  1371                 thyname ^ " " ^ (quotename thmname) ^ ")") thy'
  1372           else thy'
  1373     in
  1374         (thy2,hth')
  1375     end
  1376 
  1377 val store_thm = intern_store_thm true
  1378 
  1379 fun mk_REFL ctm =
  1380     let
  1381         val cty = Thm.ctyp_of_term ctm
  1382     in
  1383         Drule.instantiate' [SOME cty] [SOME ctm] reflexivity_thm
  1384     end
  1385 
  1386 fun REFL tm thy =
  1387     let
  1388         val _ = message "REFL:"
  1389         val (info,tm') = disamb_term tm
  1390         val ctm = Thm.cterm_of thy tm'
  1391         val res = HOLThm(rens_of info,mk_REFL ctm)
  1392         val _ = if_debug pth res
  1393     in
  1394         (thy,res)
  1395     end
  1396 
  1397 fun ASSUME tm thy =
  1398     let
  1399         val _ = message "ASSUME:"
  1400         val (info,tm') = disamb_term tm
  1401         val ctm = Thm.cterm_of thy (HOLogic.mk_Trueprop tm')
  1402         val th = Thm.trivial ctm
  1403         val res = HOLThm(rens_of info,th)
  1404         val _ = if_debug pth res
  1405     in
  1406         (thy,res)
  1407     end
  1408 
  1409 fun INST_TYPE lambda (hth as HOLThm(rens,th)) thy =
  1410     let
  1411         val _ = message "INST_TYPE:"
  1412         val _ = if_debug pth hth
  1413         val tys_before = OldTerm.add_term_tfrees (prop_of th,[])
  1414         val th1 = Thm.varifyT_global th
  1415         val tys_after = OldTerm.add_term_tvars (prop_of th1,[])
  1416         val tyinst = map (fn (bef, iS) =>
  1417                              (case try (Lib.assoc (TFree bef)) lambda of
  1418                                   SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
  1419                                 | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
  1420                              ))
  1421                          (zip tys_before tys_after)
  1422         val res = Drule.instantiate (tyinst,[]) th1
  1423         val hth = HOLThm([],res)
  1424         val res = norm_hthm thy hth
  1425         val _ = message "RESULT:"
  1426         val _ = if_debug pth res
  1427     in
  1428         (thy,res)
  1429     end
  1430 
  1431 fun INST sigma hth thy =
  1432     let
  1433         val _ = message "INST:"
  1434         val _ = if_debug (app (fn (x,y) => (prin x; prin y))) sigma
  1435         val _ = if_debug pth hth
  1436         val (sdom,srng) = ListPair.unzip (rev sigma)
  1437         val th = hthm2thm hth
  1438         val th1 = mk_INST (map (cterm_of thy) sdom) (map (cterm_of thy) srng) th
  1439         val res = HOLThm([],th1)
  1440         val _ = message "RESULT:"
  1441         val _ = if_debug pth res
  1442     in
  1443         (thy,res)
  1444     end
  1445 
  1446 fun EQ_IMP_RULE (hth as HOLThm(rens,th)) thy =
  1447     let
  1448         val _ = message "EQ_IMP_RULE:"
  1449         val _ = if_debug pth hth
  1450         val res = HOLThm(rens,th RS eqimp_thm)
  1451         val _ = message "RESULT:"
  1452         val _ = if_debug pth res
  1453     in
  1454         (thy,res)
  1455     end
  1456 
  1457 fun mk_EQ_MP th1 th2 = [beta_eta_thm th1, beta_eta_thm th2] MRS eqmp_thm
  1458 
  1459 fun EQ_MP hth1 hth2 thy =
  1460     let
  1461         val _ = message "EQ_MP:"
  1462         val _ = if_debug pth hth1
  1463         val _ = if_debug pth hth2
  1464         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1465         val res = HOLThm(rens_of info,mk_EQ_MP th1 th2)
  1466         val _ = message "RESULT:"
  1467         val _ = if_debug pth res
  1468     in
  1469         (thy,res)
  1470     end
  1471 
  1472 fun mk_COMB th1 th2 thy =
  1473     let
  1474         val (f,g) = case concl_of th1 of
  1475                         _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
  1476                       | _ => raise ERR "mk_COMB" "First theorem not an equality"
  1477         val (x,y) = case concl_of th2 of
  1478                         _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
  1479                       | _ => raise ERR "mk_COMB" "Second theorem not an equality"
  1480         val fty = type_of f
  1481         val (fd,fr) = Term.dest_funT fty
  1482         val comb_thm' = Drule.instantiate'
  1483                             [SOME (ctyp_of thy fd),SOME (ctyp_of thy fr)]
  1484                             [SOME (cterm_of thy f),SOME (cterm_of thy g),
  1485                              SOME (cterm_of thy x),SOME (cterm_of thy y)] comb_thm
  1486     in
  1487         [th1,th2] MRS comb_thm'
  1488     end
  1489 
  1490 fun SUBST rews ctxt hth thy =
  1491     let
  1492         val _ = message "SUBST:"
  1493         val _ = if_debug (app pth) rews
  1494         val _ = if_debug prin ctxt
  1495         val _ = if_debug pth hth
  1496         val (info,th) = disamb_thm hth
  1497         val (info1,ctxt') = disamb_term_from info ctxt
  1498         val (info2,rews') = disamb_thms_from info1 rews
  1499 
  1500         val cctxt = cterm_of thy ctxt'
  1501         fun subst th [] = th
  1502           | subst th (rew::rews) = subst (mk_COMB th rew thy) rews
  1503         val res = HOLThm(rens_of info2,mk_EQ_MP (subst (mk_REFL cctxt) rews') th)
  1504         val _ = message "RESULT:"
  1505         val _ = if_debug pth res
  1506     in
  1507         (thy,res)
  1508     end
  1509 
  1510 fun DISJ_CASES hth hth1 hth2 thy =
  1511     let
  1512         val _ = message "DISJ_CASES:"
  1513         val _ = if_debug (app pth) [hth,hth1,hth2]
  1514         val (info,th) = disamb_thm hth
  1515         val (info1,th1) = disamb_thm_from info hth1
  1516         val (info2,th2) = disamb_thm_from info1 hth2
  1517         val th1 = norm_hyps th1
  1518         val th2 = norm_hyps th2
  1519         val (l,r) = case concl_of th of
  1520                         _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
  1521                       | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
  1522         val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
  1523         val th2' = rearrange thy (HOLogic.mk_Trueprop r) th2
  1524         val res1 = th RS disj_cases_thm
  1525         val res2 = uniq_compose ((nprems_of th1')-1) th1' ((nprems_of th)+1) res1
  1526         val res3 = uniq_compose ((nprems_of th2')-1) th2' (nprems_of res2) res2
  1527         val res  = HOLThm(rens_of info2,res3)
  1528         val _ = message "RESULT:"
  1529         val _ = if_debug pth res
  1530     in
  1531         (thy,res)
  1532     end
  1533 
  1534 fun DISJ1 hth tm thy =
  1535     let
  1536         val _ = message "DISJ1:"
  1537         val _ = if_debug pth hth
  1538         val _ = if_debug prin tm
  1539         val (info,th) = disamb_thm hth
  1540         val (info',tm') = disamb_term_from info tm
  1541         val ct = Thm.cterm_of thy tm'
  1542         val disj1_thm' = Drule.instantiate' [] [NONE,SOME ct] disj1_thm
  1543         val res = HOLThm(rens_of info',th RS disj1_thm')
  1544         val _ = message "RESULT:"
  1545         val _ = if_debug pth res
  1546     in
  1547         (thy,res)
  1548     end
  1549 
  1550 fun DISJ2 tm hth thy =
  1551     let
  1552         val _ = message "DISJ1:"
  1553         val _ = if_debug prin tm
  1554         val _ = if_debug pth hth
  1555         val (info,th) = disamb_thm hth
  1556         val (info',tm') = disamb_term_from info tm
  1557         val ct = Thm.cterm_of thy tm'
  1558         val disj2_thm' = Drule.instantiate' [] [NONE,SOME ct] disj2_thm
  1559         val res = HOLThm(rens_of info',th RS disj2_thm')
  1560         val _ = message "RESULT:"
  1561         val _ = if_debug pth res
  1562     in
  1563         (thy,res)
  1564     end
  1565 
  1566 fun IMP_ANTISYM hth1 hth2 thy =
  1567     let
  1568         val _ = message "IMP_ANTISYM:"
  1569         val _ = if_debug pth hth1
  1570         val _ = if_debug pth hth2
  1571         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1572         val th = [beta_eta_thm th1,beta_eta_thm th2] MRS imp_antisym_thm
  1573         val res = HOLThm(rens_of info,th)
  1574         val _ = message "RESULT:"
  1575         val _ = if_debug pth res
  1576     in
  1577         (thy,res)
  1578     end
  1579 
  1580 fun SYM (hth as HOLThm(rens,th)) thy =
  1581     let
  1582         val _ = message "SYM:"
  1583         val _ = if_debug pth hth
  1584         val th = th RS symmetry_thm
  1585         val res = HOLThm(rens,th)
  1586         val _ = message "RESULT:"
  1587         val _ = if_debug pth res
  1588     in
  1589         (thy,res)
  1590     end
  1591 
  1592 fun MP hth1 hth2 thy =
  1593     let
  1594         val _ = message "MP:"
  1595         val _ = if_debug pth hth1
  1596         val _ = if_debug pth hth2
  1597         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1598         val th = [beta_eta_thm th1,beta_eta_thm th2] MRS mp_thm
  1599         val res = HOLThm(rens_of info,th)
  1600         val _ = message "RESULT:"
  1601         val _ = if_debug pth res
  1602     in
  1603         (thy,res)
  1604     end
  1605 
  1606 fun CONJ hth1 hth2 thy =
  1607     let
  1608         val _ = message "CONJ:"
  1609         val _ = if_debug pth hth1
  1610         val _ = if_debug pth hth2
  1611         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1612         val th = [th1,th2] MRS conj_thm
  1613         val res = HOLThm(rens_of info,th)
  1614         val _ = message "RESULT:"
  1615         val _ = if_debug pth res
  1616     in
  1617         (thy,res)
  1618     end
  1619 
  1620 fun CONJUNCT1 (hth as HOLThm(rens,th)) thy =
  1621     let
  1622         val _ = message "CONJUNCT1:"
  1623         val _ = if_debug pth hth
  1624         val res = HOLThm(rens,th RS conjunct1_thm)
  1625         val _ = message "RESULT:"
  1626         val _ = if_debug pth res
  1627     in
  1628         (thy,res)
  1629     end
  1630 
  1631 fun CONJUNCT2 (hth as HOLThm(rens,th)) thy =
  1632     let
  1633         val _ = message "CONJUNCT1:"
  1634         val _ = if_debug pth hth
  1635         val res = HOLThm(rens,th RS conjunct2_thm)
  1636         val _ = message "RESULT:"
  1637         val _ = if_debug pth res
  1638     in
  1639         (thy,res)
  1640     end
  1641 
  1642 fun EXISTS ex wit hth thy =
  1643     let
  1644         val _ = message "EXISTS:"
  1645         val _ = if_debug prin ex
  1646         val _ = if_debug prin wit
  1647         val _ = if_debug pth hth
  1648         val (info,th) = disamb_thm hth
  1649         val (info',[ex',wit']) = disamb_terms_from info [ex,wit]
  1650         val cwit = cterm_of thy wit'
  1651         val cty = ctyp_of_term cwit
  1652         val a = case ex' of
  1653                     (Const(@{const_name Ex},_) $ a) => a
  1654                   | _ => raise ERR "EXISTS" "Argument not existential"
  1655         val ca = cterm_of thy a
  1656         val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm)
  1657         val th1 = beta_eta_thm th
  1658         val th2 = implies_elim_all th1
  1659         val th3 = th2 COMP exists_thm'
  1660         val th  = implies_intr_hyps th3
  1661         val res = HOLThm(rens_of info',th)
  1662         val _   = message "RESULT:"
  1663         val _   = if_debug pth res
  1664     in
  1665         (thy,res)
  1666     end
  1667 
  1668 fun CHOOSE v hth1 hth2 thy =
  1669     let
  1670         val _ = message "CHOOSE:"
  1671         val _ = if_debug prin v
  1672         val _ = if_debug pth hth1
  1673         val _ = if_debug pth hth2
  1674         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1675         val (info',v') = disamb_term_from info v
  1676         fun strip 0 _ th = th
  1677           | strip n (p::ps) th =
  1678             strip (n-1) ps (Thm.implies_elim th (Thm.assume p))
  1679           | strip _ _ _ = raise ERR "CHOOSE" "strip error"
  1680         val cv = cterm_of thy v'
  1681         val th2 = norm_hyps th2
  1682         val cvty = ctyp_of_term cv
  1683         val c = HOLogic.dest_Trueprop (concl_of th2)
  1684         val cc = cterm_of thy c
  1685         val a = case concl_of th1 of
  1686                     _ $ (Const(@{const_name Ex},_) $ a) => a
  1687                   | _ => raise ERR "CHOOSE" "Conclusion not existential"
  1688         val ca = cterm_of (theory_of_thm th1) a
  1689         val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
  1690         val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
  1691         val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
  1692         val th23 = beta_eta_thm (Thm.forall_intr cv th22)
  1693         val th11 = implies_elim_all (beta_eta_thm th1)
  1694         val th' = th23 COMP (th11 COMP choose_thm')
  1695         val th = implies_intr_hyps th'
  1696         val res = HOLThm(rens_of info',th)
  1697         val _   = message "RESULT:"
  1698         val _   = if_debug pth res
  1699     in
  1700         (thy,res)
  1701     end
  1702 
  1703 fun GEN v hth thy =
  1704     let
  1705       val _ = message "GEN:"
  1706         val _ = if_debug prin v
  1707         val _ = if_debug pth hth
  1708         val (info,th) = disamb_thm hth
  1709         val (info',v') = disamb_term_from info v
  1710         val res = HOLThm(rens_of info',mk_GEN v' th thy)
  1711         val _ = message "RESULT:"
  1712         val _ = if_debug pth res
  1713     in
  1714         (thy,res)
  1715     end
  1716 
  1717 fun SPEC tm hth thy =
  1718     let
  1719         val _ = message "SPEC:"
  1720         val _ = if_debug prin tm
  1721         val _ = if_debug pth hth
  1722         val (info,th) = disamb_thm hth
  1723         val (info',tm') = disamb_term_from info tm
  1724         val ctm = Thm.cterm_of thy tm'
  1725         val cty = Thm.ctyp_of_term ctm
  1726         val spec' = Drule.instantiate' [SOME cty] [NONE,SOME ctm] spec_thm
  1727         val th = th RS spec'
  1728         val res = HOLThm(rens_of info',th)
  1729         val _ = message "RESULT:"
  1730         val _ = if_debug pth res
  1731     in
  1732         (thy,res)
  1733     end
  1734 
  1735 fun COMB hth1 hth2 thy =
  1736     let
  1737         val _ = message "COMB:"
  1738         val _ = if_debug pth hth1
  1739         val _ = if_debug pth hth2
  1740         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1741         val res = HOLThm(rens_of info,mk_COMB th1 th2 thy)
  1742         val _ = message "RESULT:"
  1743         val _ = if_debug pth res
  1744     in
  1745         (thy,res)
  1746     end
  1747 
  1748 fun TRANS hth1 hth2 thy =
  1749     let
  1750         val _ = message "TRANS:"
  1751         val _ = if_debug pth hth1
  1752         val _ = if_debug pth hth2
  1753         val (info,[th1,th2]) = disamb_thms [hth1,hth2]
  1754         val th = [th1,th2] MRS trans_thm
  1755         val res = HOLThm(rens_of info,th)
  1756         val _ = message "RESULT:"
  1757         val _ = if_debug pth res
  1758     in
  1759         (thy,res)
  1760     end
  1761 
  1762 
  1763 fun CCONTR tm hth thy =
  1764     let
  1765         val _ = message "SPEC:"
  1766         val _ = if_debug prin tm
  1767         val _ = if_debug pth hth
  1768         val (info,th) = disamb_thm hth
  1769         val (info',tm') = disamb_term_from info tm
  1770         val th = norm_hyps th
  1771         val ct = cterm_of thy tm'
  1772         val th1 = rearrange thy (HOLogic.mk_Trueprop (Const(@{const_name Not},HOLogic.boolT-->HOLogic.boolT) $ tm')) th
  1773         val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm
  1774         val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
  1775         val res = HOLThm(rens_of info',res1)
  1776         val _ = message "RESULT:"
  1777         val _ = if_debug pth res
  1778     in
  1779         (thy,res)
  1780     end
  1781 
  1782 fun mk_ABS v th thy =
  1783     let
  1784         val cv = cterm_of thy v
  1785         val th1 = implies_elim_all (beta_eta_thm th)
  1786         val (f,g) = case concl_of th1 of
  1787                         _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
  1788                       | _ => raise ERR "mk_ABS" "Bad conclusion"
  1789         val (fd,fr) = Term.dest_funT (type_of f)
  1790         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
  1791         val th2 = Thm.forall_intr cv th1
  1792         val th3 = th2 COMP abs_thm'
  1793         val res = implies_intr_hyps th3
  1794     in
  1795         res
  1796     end
  1797 
  1798 fun ABS v hth thy =
  1799     let
  1800         val _ = message "ABS:"
  1801         val _ = if_debug prin v
  1802         val _ = if_debug pth hth
  1803         val (info,th) = disamb_thm hth
  1804         val (info',v') = disamb_term_from info v
  1805         val res = HOLThm(rens_of info',mk_ABS v' th thy)
  1806         val _ = message "RESULT:"
  1807         val _ = if_debug pth res
  1808     in
  1809         (thy,res)
  1810     end
  1811 
  1812 fun GEN_ABS copt vlist hth thy =
  1813     let
  1814         val _ = message "GEN_ABS:"
  1815         val _ = case copt of
  1816                     SOME c => if_debug prin c
  1817                   | NONE => ()
  1818         val _ = if_debug (app prin) vlist
  1819         val _ = if_debug pth hth
  1820         val (info,th) = disamb_thm hth
  1821         val (info',vlist') = disamb_terms_from info vlist
  1822         val th1 =
  1823             case copt of
  1824                 SOME (c as Const(cname,cty)) =>
  1825                 let
  1826                     fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!"
  1827                       | inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty
  1828                                                             then ty2
  1829                                                             else ty
  1830                       | inst_type ty1 ty2 (ty as Type(name,tys)) =
  1831                         Type(name,map (inst_type ty1 ty2) tys)
  1832                 in
  1833                     fold_rev (fn v => fn th =>
  1834                               let
  1835                                   val cdom = fst (Term.dest_funT (fst (Term.dest_funT cty)))
  1836                                   val vty  = type_of v
  1837                                   val newcty = inst_type cdom vty cty
  1838                                   val cc = cterm_of thy (Const(cname,newcty))
  1839                               in
  1840                                   mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy
  1841                               end) vlist' th
  1842                 end
  1843               | SOME _ => raise ERR "GEN_ABS" "Bad constant"
  1844               | NONE =>
  1845                 fold_rev (fn v => fn th => mk_ABS v th thy) vlist' th
  1846         val res = HOLThm(rens_of info',th1)
  1847         val _ = message "RESULT:"
  1848         val _ = if_debug pth res
  1849     in
  1850         (thy,res)
  1851     end
  1852 
  1853 fun NOT_INTRO (hth as HOLThm(rens,th)) thy =
  1854     let
  1855         val _ = message "NOT_INTRO:"
  1856         val _ = if_debug pth hth
  1857         val th1 = implies_elim_all (beta_eta_thm th)
  1858         val a = case concl_of th1 of
  1859                     _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
  1860                   | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
  1861         val ca = cterm_of thy a
  1862         val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
  1863         val res = HOLThm(rens,implies_intr_hyps th2)
  1864         val _ = message "RESULT:"
  1865         val _ = if_debug pth res
  1866     in
  1867         (thy,res)
  1868     end
  1869 
  1870 fun NOT_ELIM (hth as HOLThm(rens,th)) thy =
  1871     let
  1872         val _ = message "NOT_INTRO:"
  1873         val _ = if_debug pth hth
  1874         val th1 = implies_elim_all (beta_eta_thm th)
  1875         val a = case concl_of th1 of
  1876                     _ $ (Const(@{const_name Not},_) $ a) => a
  1877                   | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
  1878         val ca = cterm_of thy a
  1879         val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
  1880         val res = HOLThm(rens,implies_intr_hyps th2)
  1881         val _ = message "RESULT:"
  1882         val _ = if_debug pth res
  1883     in
  1884         (thy,res)
  1885     end
  1886 
  1887 fun DISCH tm hth thy =
  1888     let
  1889         val _ = message "DISCH:"
  1890         val _ = if_debug prin tm
  1891         val _ = if_debug pth hth
  1892         val (info,th) = disamb_thm hth
  1893         val (info',tm') = disamb_term_from info tm
  1894         val prems = prems_of th
  1895         val th1 = beta_eta_thm th
  1896         val th2 = implies_elim_all th1
  1897         val th3 = Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop tm')) th2
  1898         val th4 = th3 COMP disch_thm
  1899         val res = HOLThm (rens_of info', implies_intr_hyps th4)
  1900         val _ = message "RESULT:"
  1901         val _ = if_debug pth res
  1902     in
  1903         (thy,res)
  1904     end
  1905 
  1906 val spaces = space_implode " "
  1907 
  1908 fun new_definition thyname constname rhs thy =
  1909     let
  1910         val constname = rename_const thyname thy constname
  1911         val redeclared = is_some (Sign.const_type thy (Sign.intern_const thy constname));
  1912         val _ = warning ("Introducing constant " ^ constname)
  1913         val (thmname,thy) = get_defname thyname constname thy
  1914         val (info,rhs') = disamb_term rhs
  1915         val ctype = type_of rhs'
  1916         val csyn = mk_syn thy constname
  1917         val thy1 = case HOL4DefThy.get thy of
  1918                        Replaying _ => thy
  1919                      | _ => (ImportRecorder.add_consts [(constname, ctype, csyn)];
  1920                               Sign.add_consts_i [(Binding.name constname,ctype,csyn)] thy)
  1921         val eq = mk_defeq constname rhs' thy1
  1922         val (thms, thy2) = Global_Theory.add_defs false [((Binding.name thmname,eq),[])] thy1
  1923         val _ = ImportRecorder.add_defs thmname eq
  1924         val def_thm = hd thms
  1925         val thm' = def_thm RS meta_eq_to_obj_eq_thm
  1926         val (thy',th) = (thy2, thm')
  1927         val fullcname = Sign.intern_const thy' constname
  1928         val thy'' = add_hol4_const_mapping thyname constname true fullcname thy'
  1929         val _ = ImportRecorder.add_hol_const_mapping thyname constname fullcname
  1930         val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
  1931         val rew = rewrite_hol4_term eq thy''
  1932         val crhs = cterm_of thy'' (#2 (Logic.dest_equals (prop_of rew)))
  1933         val thy22 =
  1934           if Thm.def_name constname = thmname andalso not redeclared andalso csyn = NoSyn
  1935           then
  1936               let
  1937                   val ctxt = Syntax.init_pretty_global thy''
  1938                   val p1 = quotename constname
  1939                   val p2 = Syntax.string_of_typ ctxt ctype
  1940                   val p3 = string_of_mixfix csyn
  1941                   val p4 = smart_string_of_cterm ctxt crhs
  1942               in
  1943                 add_dump ("definition\n  " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ " where\n  " ^ p4) thy''
  1944               end
  1945           else
  1946               let val ctxt = Syntax.init_pretty_global thy'' in
  1947                 add_dump ("consts\n  " ^ quotename constname ^ " :: \"" ^
  1948                   Syntax.string_of_typ ctxt ctype ^
  1949                   "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n  " ^
  1950                   quotename thmname ^ ": " ^ smart_string_of_cterm ctxt crhs) thy''
  1951               end
  1952         val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
  1953                       SOME (_,res) => HOLThm(rens_of linfo,res)
  1954                     | NONE => raise ERR "new_definition" "Bad conclusion"
  1955         val fullname = Sign.full_bname thy22 thmname
  1956         val thy22' = case opt_get_output_thy thy22 of
  1957                          "" => (ImportRecorder.add_hol_mapping thyname thmname fullname;
  1958                                 add_hol4_mapping thyname thmname fullname thy22)
  1959                        | output_thy =>
  1960                          let
  1961                              val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
  1962                              val _ = ImportRecorder.add_hol_move fullname moved_thmname
  1963                              val _ = ImportRecorder.add_hol_mapping thyname thmname moved_thmname
  1964                          in
  1965                              thy22 |> add_hol4_move fullname moved_thmname
  1966                                    |> add_hol4_mapping thyname thmname moved_thmname
  1967                          end
  1968         val _ = message "new_definition:"
  1969         val _ = if_debug pth hth
  1970     in
  1971         (thy22',hth)
  1972     end
  1973 
  1974 fun new_specification thyname thmname names hth thy =
  1975     case HOL4DefThy.get thy of
  1976         Replaying _ => (thy,hth)
  1977       | _ =>
  1978         let
  1979             val _ = message "NEW_SPEC:"
  1980             val _ = if_debug pth hth
  1981             val names = map (rename_const thyname thy) names
  1982             val _ = warning ("Introducing constants " ^ commas names)
  1983             val (HOLThm(rens,th)) = norm_hthm thy hth
  1984             val thy1 = case HOL4DefThy.get thy of
  1985                            Replaying _ => thy
  1986                          | _ =>
  1987                            let
  1988                                fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
  1989                                  | dest_eta_abs body =
  1990                                    let
  1991                                        val (dT,rT) = Term.dest_funT (type_of body)
  1992                                    in
  1993                                        ("x",dT,body $ Bound 0)
  1994                                    end
  1995                                    handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
  1996                                fun dest_exists (Const(@{const_name Ex},_) $ abody) =
  1997                                    dest_eta_abs abody
  1998                                  | dest_exists tm =
  1999                                    raise ERR "new_specification" "Bad existential formula"
  2000 
  2001                                val (consts,_) = Library.foldl (fn ((cs,ex),cname) =>
  2002                                                           let
  2003                                                               val (_,cT,p) = dest_exists ex
  2004                                                           in
  2005                                                               ((cname,cT,mk_syn thy cname)::cs,p)
  2006                                                           end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
  2007                                val str = Library.foldl (fn (acc, (c, T, csyn)) =>
  2008                                    acc ^ "\n  " ^ quotename c ^ " :: \"" ^
  2009                                    Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
  2010                                val thy' = add_dump str thy
  2011                                val _ = ImportRecorder.add_consts consts
  2012                            in
  2013                                Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy'
  2014                            end
  2015 
  2016             val thy1 = fold_rev (fn name => fn thy =>
  2017                                 snd (get_defname thyname name thy)) names thy1
  2018             fun new_name name = fst (get_defname thyname name thy1)
  2019             val names' = map (fn name => (new_name name,name,false)) names
  2020             val (thy',res) = Choice_Specification.add_specification NONE
  2021                                  names'
  2022                                  (thy1,th)
  2023             val _ = ImportRecorder.add_specification names' th
  2024             val res' = Thm.unvarify_global res
  2025             val hth = HOLThm(rens,res')
  2026             val rew = rewrite_hol4_term (concl_of res') thy'
  2027             val th  = Thm.equal_elim rew res'
  2028             fun handle_const (name,thy) =
  2029                 let
  2030                     val defname = Thm.def_name name
  2031                     val (newname,thy') = get_defname thyname name thy
  2032                 in
  2033                     (if defname = newname
  2034                      then quotename name
  2035                      else (quotename newname) ^ ": " ^ (quotename name),thy')
  2036                 end
  2037             val (new_names,thy') = fold_rev (fn name => fn (names, thy) =>
  2038                                             let
  2039                                                 val (name',thy') = handle_const (name,thy)
  2040                                             in
  2041                                                 (name'::names,thy')
  2042                                             end) names ([], thy')
  2043             val thy'' =
  2044               add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^
  2045                 (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^
  2046                 "\n  by (import " ^ thyname ^ " " ^ thmname ^ ")")
  2047                 thy'
  2048             val _ = message "RESULT:"
  2049             val _ = if_debug pth hth
  2050         in
  2051             intern_store_thm false thyname thmname hth thy''
  2052         end
  2053 
  2054 fun new_axiom name tm thy = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
  2055 
  2056 fun to_isa_thm (hth as HOLThm(_,th)) =
  2057     let
  2058         val (HOLThm args) = norm_hthm (theory_of_thm th) hth
  2059     in
  2060         apsnd Thm.strip_shyps args
  2061     end
  2062 
  2063 fun to_isa_term tm = tm
  2064 
  2065 local
  2066     val light_nonempty = @{thm light_ex_imp_nonempty}
  2067     val ex_imp_nonempty = @{thm ex_imp_nonempty}
  2068     val typedef_hol2hol4 = @{thm typedef_hol2hol4}
  2069     val typedef_hol2hollight = @{thm typedef_hol2hollight}
  2070 in
  2071 fun new_type_definition thyname thmname tycname hth thy =
  2072     case HOL4DefThy.get thy of
  2073         Replaying _ => (thy,hth)
  2074       | _ =>
  2075         let
  2076             val _ = message "TYPE_DEF:"
  2077             val _ = if_debug pth hth
  2078             val _ = warning ("Introducing type " ^ tycname)
  2079             val (HOLThm(rens,td_th)) = norm_hthm thy hth
  2080             val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
  2081             val c = case concl_of th2 of
  2082                         _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
  2083                       | _ => raise ERR "new_type_definition" "Bad type definition theorem"
  2084             val tfrees = OldTerm.term_tfrees c
  2085             val tnames = map fst tfrees
  2086             val tsyn = mk_syn thy tycname
  2087             val typ = (tycname,tnames,tsyn)
  2088             val ((_, typedef_info), thy') =
  2089               Typedef.add_typedef_global false (SOME (Binding.name thmname))
  2090                 (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c NONE (rtac th2 1) thy
  2091             val _ = ImportRecorder.add_typedef (SOME thmname) typ c NONE th2
  2092 
  2093             val th3 = (#type_definition (#2 typedef_info)) RS typedef_hol2hol4
  2094 
  2095             val fulltyname = Sign.intern_type thy' tycname
  2096             val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
  2097             val _ = ImportRecorder.add_hol_type_mapping thyname tycname fulltyname
  2098 
  2099             val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th3))
  2100             val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
  2101                     else ()
  2102             val thy4 = add_hol4_pending thyname thmname args thy''
  2103             val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
  2104 
  2105             val rew = rewrite_hol4_term (concl_of td_th) thy4
  2106             val th  = Thm.equal_elim rew (Thm.transfer thy4 td_th)
  2107             val c   = case HOLogic.dest_Trueprop (prop_of th) of
  2108                           Const(@{const_name Ex},exT) $ P =>
  2109                           let
  2110                               val PT = domain_type exT
  2111                           in
  2112                               Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P
  2113                           end
  2114                         | _ => error "Internal error in ProofKernel.new_typedefinition"
  2115             val tnames_string = if null tnames
  2116                                 then ""
  2117                                 else "(" ^ commas tnames ^ ") "
  2118             val proc_prop =
  2119               smart_string_of_cterm
  2120                 (Syntax.init_pretty_global thy4
  2121                   |> not (null tnames) ? Config.put show_all_types true)
  2122             val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " "
  2123                                  ^ (string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
  2124 
  2125             val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hol4 [OF type_definition_" ^ tycname ^ "]") thy5
  2126 
  2127             val _ = message "RESULT:"
  2128             val _ = if_debug pth hth'
  2129         in
  2130             (thy6,hth')
  2131         end
  2132 
  2133 fun add_dump_syntax thy name =
  2134     let
  2135       val n = quotename name
  2136       val syn = string_of_mixfix (mk_syn thy name)
  2137     in
  2138       add_dump ("syntax\n  "^n^" :: _ "^syn) thy
  2139     end
  2140 
  2141 fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
  2142     case HOL4DefThy.get thy of
  2143         Replaying _ => (thy,
  2144           HOLThm([], Global_Theory.get_thm thy (thmname^"_@intern")) handle ERROR _ => hth)
  2145       | _ =>
  2146         let
  2147             val _ = message "TYPE_INTRO:"
  2148             val _ = if_debug pth hth
  2149             val _ = warning ("Introducing type " ^ tycname ^ " (with morphisms " ^ abs_name ^ " and " ^ rep_name ^ ")")
  2150             val (HOLThm(rens,td_th)) = norm_hthm thy hth
  2151             val tT = type_of t
  2152             val light_nonempty' =
  2153                 Drule.instantiate' [SOME (ctyp_of thy tT)]
  2154                                    [SOME (cterm_of thy P),
  2155                                     SOME (cterm_of thy t)] light_nonempty
  2156             val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
  2157             val c = case concl_of th2 of
  2158                         _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
  2159                       | _ => raise ERR "type_introduction" "Bad type definition theorem"
  2160             val tfrees = OldTerm.term_tfrees c
  2161             val tnames = sort_strings (map fst tfrees)
  2162             val tsyn = mk_syn thy tycname
  2163             val typ = (tycname,tnames,tsyn)
  2164             val ((_, typedef_info), thy') =
  2165               Typedef.add_typedef_global false NONE
  2166                 (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c
  2167                 (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
  2168             val _ = ImportRecorder.add_typedef NONE typ c (SOME(rep_name,abs_name)) th2
  2169             val fulltyname = Sign.intern_type thy' tycname
  2170             val aty = Type (fulltyname, map mk_vartype tnames)
  2171             val abs_ty = tT --> aty
  2172             val rep_ty = aty --> tT
  2173             val typedef_hol2hollight' =
  2174                 Drule.instantiate'
  2175                     [SOME (ctyp_of thy' aty), SOME (ctyp_of thy' tT)]
  2176                     [NONE, NONE, NONE, SOME (cterm_of thy' (Free ("a", aty))), SOME (cterm_of thy' (Free ("r", tT)))]
  2177                     typedef_hol2hollight
  2178             val th4 = (#type_definition (#2 typedef_info)) RS typedef_hol2hollight'
  2179             val _ = null (Thm.fold_terms Term.add_tvars th4 []) orelse
  2180               raise ERR "type_introduction" "no type variables expected any more"
  2181             val _ = null (Thm.fold_terms Term.add_vars th4 []) orelse
  2182               raise ERR "type_introduction" "no term variables expected any more"
  2183             val _ = message ("step 3: thyname="^thyname^", tycname="^tycname^", fulltyname="^fulltyname)
  2184             val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
  2185             val _ = ImportRecorder.add_hol_type_mapping thyname tycname fulltyname
  2186             val _ = message "step 4"
  2187             val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th4))
  2188             val thy4 = add_hol4_pending thyname thmname args thy''
  2189             val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
  2190 
  2191             val P' = P (* why !? #2 (Logic.dest_equals (concl_of (rewrite_hol4_term P thy4))) *)
  2192             val c   =
  2193                 let
  2194                     val PT = type_of P'
  2195                 in
  2196                     Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P'
  2197                 end
  2198 
  2199             val tnames_string = if null tnames
  2200                                 then ""
  2201                                 else "(" ^ commas tnames ^ ") "
  2202             val proc_prop =
  2203               smart_string_of_cterm
  2204                 (Syntax.init_pretty_global thy4
  2205                   |> not (null tnames) ? Config.put show_all_types true)
  2206             val thy = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^
  2207               " = " ^ (proc_prop (cterm_of thy4 c)) ^ " " ^
  2208               (string_of_mixfix tsyn) ^ " morphisms "^
  2209               (quote rep_name)^" "^(quote abs_name)^"\n"^
  2210               ("  apply (rule light_ex_imp_nonempty[where t="^
  2211               (proc_prop (cterm_of thy4 t))^"])\n"^
  2212               ("  by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
  2213             val str_aty = Syntax.string_of_typ_global thy aty
  2214             val thy = add_dump_syntax thy rep_name
  2215             val thy = add_dump_syntax thy abs_name
  2216             val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
  2217               " = typedef_hol2hollight \n"^
  2218               "  [where a=\"a :: "^str_aty^"\" and r=r" ^
  2219               " ,\n   OF "^(quotename ("type_definition_" ^ tycname)) ^ "]") thy
  2220             val _ = message "RESULT:"
  2221             val _ = if_debug pth hth'
  2222         in
  2223             (thy,hth')
  2224         end
  2225 end
  2226 
  2227 val prin = prin
  2228 
  2229 end