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