src/Pure/section_utils.ML
author wenzelm
Wed, 17 Mar 1999 16:33:00 +0100
changeset 6390 5d58c100ca3f
parent 6039 01f67f5f8dd0
child 7744 f27d54c1ef39
permissions -rw-r--r--
qualify Theory.sign_of etc.;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4848
99c8d95c51d6 moved mk_defpair to logic.ML;
wenzelm
parents: 4695
diff changeset
     1
(*  Title: 	Pure/section_utils.ML
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     5
4848
99c8d95c51d6 moved mk_defpair to logic.ML;
wenzelm
parents: 4695
diff changeset
     6
Utilities for writing new theory sections.
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     7
*)
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     8
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
     9
fun ap t u = t$u;
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    10
fun app t (u1,u2) = t $ u1 $ u2;
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    11
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    12
(*Make distinct individual variables a1, a2, a3, ..., an. *)
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    13
fun mk_frees a [] = []
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    14
  | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (bump_string a) Ts;
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    15
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    16
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    17
(*Read an assumption in the given theory*)
6390
5d58c100ca3f qualify Theory.sign_of etc.;
wenzelm
parents: 6039
diff changeset
    18
fun assume_read thy a = Thm.assume (read_cterm (Theory.sign_of thy) (a,propT));
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    19
1390
bf523422a3df Commented and renamed vars in readtm
paulson
parents: 654
diff changeset
    20
(*Read a term from string "b", with expected type T*)
bf523422a3df Commented and renamed vars in readtm
paulson
parents: 654
diff changeset
    21
fun readtm sign T b = 
bf523422a3df Commented and renamed vars in readtm
paulson
parents: 654
diff changeset
    22
    read_cterm sign (b,T) |> term_of
3536
8fb4150e2ad3 tuned error propagation msg;
wenzelm
parents: 2863
diff changeset
    23
    handle ERROR => error ("The error(s) above occurred for " ^ b);
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    24
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    25
(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    26
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    27
  | tryres (th, []) = raise THM("tryres", 0, [th]);
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    28
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    29
fun gen_make_elim elim_rls rl = 
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    30
      standard (tryres (rl, elim_rls @ [revcut_rl]));
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    31
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    32
(** String manipulation **)
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    33
4695
6aa25ee18fc4 adapted to new scanners and baroque chars;
wenzelm
parents: 3974
diff changeset
    34
(*Skipping initial blanks, find the first identifier*)	(* FIXME *)
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    35
fun scan_to_id s = 
4695
6aa25ee18fc4 adapted to new scanners and baroque chars;
wenzelm
parents: 3974
diff changeset
    36
    s |> Symbol.explode
4938
c8bbbf3c59fa Symbol.stopper;
wenzelm
parents: 4848
diff changeset
    37
    |> Scan.error (Scan.finite Symbol.stopper
4695
6aa25ee18fc4 adapted to new scanners and baroque chars;
wenzelm
parents: 3974
diff changeset
    38
      (!! (fn _ => "Expected to find an identifier in " ^ s)
6aa25ee18fc4 adapted to new scanners and baroque chars;
wenzelm
parents: 3974
diff changeset
    39
        (Scan.any Symbol.is_blank |-- Syntax.scan_id)))
6aa25ee18fc4 adapted to new scanners and baroque chars;
wenzelm
parents: 3974
diff changeset
    40
    |> #1;
579
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    41
08f465e23dc5 new file of useful things for writing theory sections
lcp
parents:
diff changeset
    42
fun is_backslash c = c = "\\";