| 1460 |      1 | (*  Title: 	Pure/section-utils
 | 
| 579 |      2 |     ID:         $Id$
 | 
| 1460 |      3 |     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 579 |      4 |     Copyright   1994  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Utilities for writing new theory sections
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | fun ap t u = t$u;
 | 
|  |     10 | fun app t (u1,u2) = t $ u1 $ u2;
 | 
|  |     11 | 
 | 
|  |     12 | (*Make distinct individual variables a1, a2, a3, ..., an. *)
 | 
|  |     13 | fun mk_frees a [] = []
 | 
|  |     14 |   | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (bump_string a) Ts;
 | 
|  |     15 | 
 | 
|  |     16 | (*Make a definition lhs==rhs*)
 | 
|  |     17 | fun mk_defpair (lhs, rhs) = 
 | 
|  |     18 |   let val Const(name, _) = head_of lhs
 | 
| 3934 |     19 |   in (Sign.base_name name ^ "_def", Logic.mk_equals (lhs, rhs)) end;
 | 
| 579 |     20 | 
 | 
|  |     21 | fun get_def thy s = get_axiom thy (s^"_def");
 | 
|  |     22 | 
 | 
|  |     23 | 
 | 
|  |     24 | (*Read an assumption in the given theory*)
 | 
|  |     25 | fun assume_read thy a = assume (read_cterm (sign_of thy) (a,propT));
 | 
|  |     26 | 
 | 
| 1390 |     27 | (*Read a term from string "b", with expected type T*)
 | 
|  |     28 | fun readtm sign T b = 
 | 
|  |     29 |     read_cterm sign (b,T) |> term_of
 | 
| 3536 |     30 |     handle ERROR => error ("The error(s) above occurred for " ^ b);
 | 
| 579 |     31 | 
 | 
|  |     32 | (*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)
 | 
|  |     33 | fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
 | 
|  |     34 |   | tryres (th, []) = raise THM("tryres", 0, [th]);
 | 
|  |     35 | 
 | 
|  |     36 | fun gen_make_elim elim_rls rl = 
 | 
|  |     37 |       standard (tryres (rl, elim_rls @ [revcut_rl]));
 | 
|  |     38 | 
 | 
|  |     39 | (** String manipulation **)
 | 
|  |     40 | 
 | 
|  |     41 | (*Skipping initial blanks, find the first identifier*)
 | 
|  |     42 | fun scan_to_id s = 
 | 
|  |     43 |     s |> explode |> take_prefix is_blank |> #2 |> Scanner.scan_id |> #1
 | 
|  |     44 |     handle LEXICAL_ERROR => error ("Expected to find an identifier in " ^ s);
 | 
|  |     45 | 
 | 
|  |     46 | fun is_backslash c = c = "\\";
 | 
|  |     47 | 
 | 
|  |     48 | (*Apply string escapes to a quoted string; see Def of Standard ML, page 3
 | 
|  |     49 |   Does not handle the \ddd form;  no error checking*)
 | 
|  |     50 | fun escape [] = []
 | 
|  |     51 |   | escape cs = (case take_prefix (not o is_backslash) cs of
 | 
| 1460 |     52 | 	 (front, []) => front
 | 
| 579 |     53 |        | (front, _::"n"::rest) => front @ ("\n" :: escape rest)
 | 
|  |     54 |        | (front, _::"t"::rest) => front @ ("\t" :: escape rest)
 | 
|  |     55 |        | (front, _::"^"::c::rest) => front @ (chr(ord(c)-64) :: escape rest)
 | 
|  |     56 |        | (front, _::"\""::rest) => front @ ("\"" :: escape rest)
 | 
|  |     57 |        | (front, _::"\\"::rest) => front @ ("\\" :: escape rest)
 | 
|  |     58 |        | (front, b::c::rest) => 
 | 
| 1460 |     59 | 	   if is_blank c   (*remove any further blanks and the following \ *)
 | 
|  |     60 | 	   then front @ escape (tl (snd (take_prefix is_blank rest)))
 | 
|  |     61 | 	   else error ("Unrecognized string escape: " ^ implode(b::c::rest)));
 | 
| 579 |     62 | 
 | 
|  |     63 | (*Remove the first and last charaters -- presumed to be quotes*)
 | 
|  |     64 | val trim = implode o escape o rev o tl o rev o tl o explode;
 | 
| 654 |     65 | 
 | 
|  |     66 | 
 | 
|  |     67 | (*Check for some named theory*)
 | 
|  |     68 | fun require_thy thy name sect =
 | 
| 3974 |     69 |   if exists (equal name) (Sign.stamp_names_of (sign_of thy)) then ()
 | 
| 2863 |     70 |   else error ("Need theory " ^ quote name ^ " as an ancestor for " ^ sect);
 |