--- a/src/HOL/Import/proof_kernel.ML Fri Aug 28 18:23:24 2009 +0200
+++ b/src/HOL/Import/proof_kernel.ML Fri Aug 28 21:04:03 2009 +0200
@@ -199,12 +199,12 @@
val ct = (cterm_of thy (HOLogic.dest_Trueprop t)
handle TERM _ => ct)
in
- quote(
+ quote (
PrintMode.setmp [] (
Library.setmp show_brackets false (
Library.setmp show_all_types true (
Library.setmp Syntax.ambiguity_is_error false (
- Library.setmp show_sorts true Display.string_of_cterm))))
+ Library.setmp show_sorts true (Syntax.string_of_term_global thy o Thm.term_of)))))
ct)
end
@@ -226,7 +226,8 @@
| G _ = raise SMART_STRING
fun F n =
let
- val str = Library.setmp show_brackets false (G n Display.string_of_cterm) ct
+ val str =
+ Library.setmp show_brackets false (G n (Syntax.string_of_term ctxt o term_of)) ct
val u = Syntax.parse_term ctxt str
|> TypeInfer.constrain T |> Syntax.check_term ctxt
in
@@ -234,8 +235,9 @@
then quote str
else F (n+1)
end
- handle ERROR mesg => F (n+1)
- | SMART_STRING => error ("smart_string failed for: "^(G 0 Display.string_of_cterm ct))
+ handle ERROR mesg => F (n + 1)
+ | SMART_STRING =>
+ error ("smart_string failed for: "^ G 0 (Syntax.string_of_term ctxt o term_of) ct)
in
PrintMode.setmp [] (Library.setmp Syntax.ambiguity_is_error true F) 0
end
@@ -243,8 +245,7 @@
val smart_string_of_thm = smart_string_of_cterm o cprop_of
-fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th)
-fun prc ct = writeln (PrintMode.setmp [] Display.string_of_cterm ct)
+fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th);
fun prin t = writeln (PrintMode.setmp []
(fn () => Syntax.string_of_term (ML_Context.the_local_context ()) t) ());
fun pth (HOLThm(ren,thm)) =
@@ -1939,16 +1940,17 @@
then
let
val p1 = quotename constname
- val p2 = Display.string_of_ctyp (ctyp_of thy'' ctype)
+ val p2 = Syntax.string_of_typ_global thy'' ctype
val p3 = string_of_mixfix csyn
val p4 = smart_string_of_cterm crhs
in
- add_dump ("constdefs\n " ^p1^ " :: \"" ^p2^ "\" "^p3^ "\n " ^p4) thy''
+ add_dump ("constdefs\n " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ "\n " ^ p4) thy''
end
else
- (add_dump ("consts\n " ^ (quotename constname) ^ " :: \"" ^ Display.string_of_ctyp (ctyp_of thy'' ctype) ^
- "\" " ^ (string_of_mixfix csyn) ^ "\n\ndefs\n " ^ (quotename thmname) ^ ": " ^ (smart_string_of_cterm crhs))
- thy'')
+ add_dump ("consts\n " ^ quotename constname ^ " :: \"" ^
+ Syntax.string_of_typ_global thy'' ctype ^
+ "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n " ^
+ quotename thmname ^ ": " ^ smart_string_of_cterm crhs) thy''
val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
SOME (_,res) => HOLThm(rens_of linfo,res)
| NONE => raise ERR "new_definition" "Bad conclusion"
@@ -2008,8 +2010,9 @@
in
((cname,cT,mk_syn thy cname)::cs,p)
end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
- val str = Library.foldl (fn (acc,(c,T,csyn)) =>
- acc ^ "\n " ^ (quotename c) ^ " :: \"" ^ Display.string_of_ctyp (ctyp_of thy T) ^ "\" " ^ (string_of_mixfix csyn)) ("consts",consts)
+ val str = Library.foldl (fn (acc, (c, T, csyn)) =>
+ acc ^ "\n " ^ quotename c ^ " :: \"" ^
+ Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
val thy' = add_dump str thy
val _ = ImportRecorder.add_consts consts
in
@@ -2137,7 +2140,7 @@
fun add_dump_constdefs thy defname constname rhs ty =
let
val n = quotename constname
- val t = Display.string_of_ctyp (ctyp_of thy ty)
+ val t = Syntax.string_of_typ_global thy ty
val syn = string_of_mixfix (mk_syn thy constname)
(*val eq = smart_string_of_cterm (cterm_of thy (Const(rhs, ty)))*)
val eq = quote (constname ^ " == "^rhs)
@@ -2224,7 +2227,7 @@
(" apply (rule light_ex_imp_nonempty[where t="^
(proc_prop (cterm_of thy4 t))^"])\n"^
(" by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
- val str_aty = Display.string_of_ctyp (ctyp_of thy aty)
+ val str_aty = Syntax.string_of_typ_global thy aty
val thy = add_dump_syntax thy rep_name
val thy = add_dump_syntax thy abs_name
val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^