Nicer names in FindTheorems.
* Patch NameSpace.get_accesses, contributed by Timothy Bourke:
  NameSpace.get_accesses has been patched to fix the following
  bug:
      theory OverHOL imports Main begin
        lemma conjI: "a & b --> b"
          by blast
        ML {* val ns = PureThy.facts_of @{theory} |> Facts.space_of;
  	    val x1 = NameSpace.get_accesses ns "HOL.conjI";
  	    val x2 = NameSpace.get_accesses ns "OverHOL.conjI"; *}
      end
  where x1 = ["conjI", "HOL.conjI"] and x2 = ["conjI", "OverHOL.conjI"],
  but x1 should be just ["HOL.conjI"].
  NameSpace.get_accesses is only used within the NameSpace
  structure itself. The two uses have been modified to retain
  their original behaviour.
  Note that NameSpace.valid_accesses gives different results:
      get_accesses ns "HOL.eq_class.eq"
      gives ["eq", "eq_class.eq", "HOL.eq_class.eq"]
      but,
      valid_accesses ns "HOL.eq_class.eq"
      gives ["HOL.eq_class.eq", "eq_class.eq", "HOL.eq", "eq"]
* Patch FindTheorems:
  Prefer names that are shorter to type in the current context.
* Re-export space_of.
(*  Title:      Pure/Thy/term_style.ML
    Author:     Florian Haftmann, TU Muenchen
Styles for terms, to use with the "term_style" and "thm_style"
antiquotations.
*)
signature TERM_STYLE =
sig
  val the_style: theory -> string -> (Proof.context -> term -> term)
  val add_style: string -> (Proof.context -> term -> term) -> theory -> theory
  val print_styles: theory -> unit
end;
structure TermStyle: TERM_STYLE =
struct
(* style data *)
fun err_dup_style name =
  error ("Duplicate declaration of antiquote style: " ^ quote name);
structure StyleData = TheoryDataFun
(
  type T = ((Proof.context -> term -> term) * stamp) Symtab.table;
  val empty = Symtab.empty;
  val copy = I;
  val extend = I;
  fun merge _ tabs : T = Symtab.merge (eq_snd (op =)) tabs
    handle Symtab.DUP dup => err_dup_style dup;
);
fun print_styles thy =
  Pretty.writeln (Pretty.strs ("antiquote styles:" :: Symtab.keys (StyleData.get thy)));
(* accessors *)
fun the_style thy name =
  (case Symtab.lookup (StyleData.get thy) name of
    NONE => error ("Unknown antiquote style: " ^ quote name)
  | SOME (style, _) => style);
fun add_style name style thy =
  StyleData.map (Symtab.update_new (name, (style, stamp ()))) thy
    handle Symtab.DUP _ => err_dup_style name;
(* predefined styles *)
fun style_binargs ctxt t =
  let
    val concl = ObjectLogic.drop_judgment (ProofContext.theory_of ctxt)
      (Logic.strip_imp_concl t)
  in
    case concl of (_ $ l $ r) => (l, r)
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
  end;
fun style_parm_premise i ctxt t =
  let val prems = Logic.strip_imp_prems t in
    if i <= length prems then nth prems (i - 1)
    else error ("Not enough premises for prem" ^ string_of_int i ^
      " in propositon: " ^ Syntax.string_of_term ctxt t)
  end;
val _ = Context.>> (Context.map_theory
 (add_style "lhs" (fst oo style_binargs) #>
  add_style "rhs" (snd oo style_binargs) #>
  add_style "prem1" (style_parm_premise 1) #>
  add_style "prem2" (style_parm_premise 2) #>
  add_style "prem3" (style_parm_premise 3) #>
  add_style "prem4" (style_parm_premise 4) #>
  add_style "prem5" (style_parm_premise 5) #>
  add_style "prem6" (style_parm_premise 6) #>
  add_style "prem7" (style_parm_premise 7) #>
  add_style "prem8" (style_parm_premise 8) #>
  add_style "prem9" (style_parm_premise 9) #>
  add_style "prem10" (style_parm_premise 10) #>
  add_style "prem11" (style_parm_premise 11) #>
  add_style "prem12" (style_parm_premise 12) #>
  add_style "prem13" (style_parm_premise 13) #>
  add_style "prem14" (style_parm_premise 14) #>
  add_style "prem15" (style_parm_premise 15) #>
  add_style "prem16" (style_parm_premise 16) #>
  add_style "prem17" (style_parm_premise 17) #>
  add_style "prem18" (style_parm_premise 18) #>
  add_style "prem19" (style_parm_premise 19) #>
  add_style "concl" (K Logic.strip_imp_concl)));
end;