src/Pure/Isar/locale.ML
author ballarin
Tue, 04 Jul 2006 14:47:01 +0200
changeset 19984 29bb4659f80a
parent 19942 dc92e3aebc71
child 19991 0c9650664d47
permissions -rw-r--r--
Method intro_locales replaced by intro_locales and unfold_locales.

(*  Title:      Pure/Isar/locale.ML
    ID:         $Id$
    Author:     Clemens Ballarin, TU Muenchen; Markus Wenzel, LMU/TU Muenchen

Locales -- Isar proof contexts as meta-level predicates, with local
syntax and implicit structures.

Draws basic ideas from Florian Kammueller's original version of
locales, but uses the richer infrastructure of Isar instead of the raw
meta-logic.  Furthermore, we provide structured import of contexts
(with merge and rename operations), as well as type-inference of the
signature parts, and predicate definitions of the specification text.

Interpretation enables the reuse of theorems of locales in other
contexts, namely those defined by theories, structured proofs and
locales themselves.

See also:

[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
    In Stefano Berardi et al., Types for Proofs and Programs: International
    Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
    Dependencies between Locales. Technical Report TUM-I0607, Technische
    Universitaet Muenchen, 2006.
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
    Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
    pages 31-43, 2006.
*)

(* TODO:
- beta-eta normalisation of interpretation parameters
- dangling type frees in locales
- test subsumption of interpretations when merging theories
- var vs. fixes in locale to be interpreted (interpretation in locale)
  (implicit locale expressions generated by multiple registrations)
*)

signature LOCALE =
sig
  datatype expr =
    Locale of string |
    Rename of expr * (string * mixfix option) option list |
    Merge of expr list
  val empty: expr
  datatype 'a element = Elem of 'a | Expr of expr

  val intern: theory -> xstring -> string
  val extern: theory -> string -> xstring
  val init: string -> theory -> cterm list * Proof.context

  (* The specification of a locale *)

  val parameters_of: theory -> string ->
    ((string * typ) * mixfix) list
  val parameters_of_expr: theory -> expr ->
    ((string * typ) * mixfix) list
  val local_asms_of: theory -> string ->
    ((string * Attrib.src list) * term list) list
  val global_asms_of: theory -> string ->
    ((string * Attrib.src list) * term list) list

  (* Processing of locale statements *)
  val read_context_statement: xstring option -> Element.context element list ->
    (string * string list) list list -> Proof.context ->
    string option * (cterm list * Proof.context) * (cterm list * Proof.context) *
      (term * term list) list list
  val cert_context_statement: string option -> Element.context_i element list ->
    (term * term list) list list -> Proof.context ->
    string option * (cterm list * Proof.context) * (cterm list * Proof.context) *
      (term * term list) list list
  val read_expr: expr -> Element.context list -> Proof.context ->
    Element.context_i list * Proof.context
  val cert_expr: expr -> Element.context_i list -> Proof.context ->
    Element.context_i list * Proof.context

  (* Diagnostic functions *)
  val print_locales: theory -> unit
  val print_locale: theory -> bool -> expr -> Element.context list -> unit
  val print_global_registrations: bool -> string -> theory -> unit
  val print_local_registrations': bool -> string -> Proof.context -> unit
  val print_local_registrations: bool -> string -> Proof.context -> unit

  val add_locale: bool -> bstring -> expr -> Element.context list -> theory
    -> (string * Proof.context (*body context!*)) * theory
  val add_locale_i: bool -> bstring -> expr -> Element.context_i list -> theory
    -> (string * Proof.context (*body context!*)) * theory

  (* Storing results *)
  val note_thmss: string -> xstring ->
    ((string * Attrib.src list) * (thmref * Attrib.src list) list) list ->
    theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory)
  val note_thmss_i: string -> string ->
    ((string * Attrib.src list) * (thm list * Attrib.src list) list) list ->
    theory -> ((string * thm list) list * (string * thm list) list) * (Proof.context * theory)
  val add_thmss: string -> string ->
    ((string * Attrib.src list) * (thm list * Attrib.src list) list) list ->
    cterm list * Proof.context ->
    ((string * thm list) list * (string * thm list) list) * Proof.context
  val add_term_syntax: string -> (Proof.context -> Proof.context) ->
    cterm list * Proof.context -> Proof.context

  val theorem: string -> Method.text option -> (thm list list -> theory -> theory) ->
    string * Attrib.src list -> Element.context element list -> Element.statement ->
    theory -> Proof.state
  val theorem_i: string -> Method.text option -> (thm list list -> theory -> theory) ->
    string * Attrib.src list -> Element.context_i element list -> Element.statement_i ->
    theory -> Proof.state
  val theorem_in_locale: string -> Method.text option ->
    (thm list list -> thm list list -> theory -> theory) ->
    xstring -> string * Attrib.src list -> Element.context element list -> Element.statement ->
    theory -> Proof.state
  val theorem_in_locale_i: string -> Method.text option ->
    (thm list list -> thm list list -> theory -> theory) ->
    string -> string * Attrib.src list -> Element.context_i element list ->
    Element.statement_i -> theory -> Proof.state
  val smart_theorem: string -> xstring option ->
    string * Attrib.src list -> Element.context element list -> Element.statement ->
    theory -> Proof.state
  val interpretation: string * Attrib.src list -> expr -> string option list ->
    theory -> Proof.state
  val interpretation_in_locale: xstring * expr -> theory -> Proof.state
  val interpret: string * Attrib.src list -> expr -> string option list ->
    bool -> Proof.state -> Proof.state
end;

structure Locale: LOCALE =
struct

fun legacy_unvarifyT thm =
  let
    val cT = Thm.ctyp_of (Thm.theory_of_thm thm);
    val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) (Drule.tvars_of thm);
  in Drule.instantiate' tfrees [] thm end;

fun legacy_unvarify raw_thm =
  let
    val thm = legacy_unvarifyT raw_thm;
    val ct = Thm.cterm_of (Thm.theory_of_thm thm);
    val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) (Drule.vars_of thm);
  in Drule.instantiate' [] frees thm end;

(** locale elements and expressions **)

datatype ctxt = datatype Element.ctxt;

datatype expr =
  Locale of string |
  Rename of expr * (string * mixfix option) option list |
  Merge of expr list;

val empty = Merge [];

datatype 'a element =
  Elem of 'a | Expr of expr;

fun map_elem f (Elem e) = Elem (f e)
  | map_elem _ (Expr e) = Expr e;

type locale =
 {axiom: Element.witness list,
    (* For locales that define predicates this is [A [A]], where A is the locale
       specification.  Otherwise []. *)
  import: expr,                                                     (*dynamic import*)
  elems: (Element.context_i * stamp) list,
    (* Static content, neither Fixes nor Constrains elements *)
  params: ((string * typ) * mixfix) list,                           (*all params*)
  lparams: string list,                                             (*local params*)
  term_syntax: ((Proof.context -> Proof.context) * stamp) list, (* FIXME depend on morphism *)
  regs: ((string * string list) * Element.witness list) list,
    (* Registrations: indentifiers and witnesses of locales interpreted in the locale. *)
  intros: thm list * thm list}
    (* Introduction rules: of delta predicate and locale predicate. *)

(* CB: an internal (Int) locale element was either imported or included,
   an external (Ext) element appears directly in the locale text. *)

datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;



(** management of registrations in theories and proof contexts **)

structure Registrations :
  sig
    type T
    val empty: T
    val join: T * T -> T
    val dest: T -> (term list * ((string * Attrib.src list) * Element.witness list)) list
    val lookup: theory -> T * term list ->
      ((string * Attrib.src list) * Element.witness list) option
    val insert: theory -> term list * (string * Attrib.src list) -> T ->
      T * (term list * ((string * Attrib.src list) * Element.witness list)) list
    val add_witness: term list -> Element.witness -> T -> T
  end =
struct
  (* a registration consists of theorems instantiating locale assumptions
     and prefix and attributes, indexed by parameter instantiation *)
  type T = ((string * Attrib.src list) * Element.witness list) Termtab.table;

  val empty = Termtab.empty;

  (* term list represented as single term, for simultaneous matching *)
  fun termify ts =
    Term.list_comb (Const ("", map fastype_of ts ---> propT), ts);
  fun untermify t =
    let fun ut (Const _) ts = ts
          | ut (s $ t) ts = ut s (t::ts)
    in ut t [] end;

  (* joining of registrations: prefix and attributes of left theory,
     thms are equal, no attempt to subsumption testing *)
  fun join (r1, r2) = Termtab.join (fn _ => fn (reg, _) => reg) (r1, r2);

  fun dest regs = map (apfst untermify) (Termtab.dest regs);

  (* registrations that subsume t *)
  fun subsumers thy t regs =
    filter (fn (t', _) => Pattern.matches thy (t', t)) (Termtab.dest regs);

  (* look up registration, pick one that subsumes the query *)
  fun lookup thy (regs, ts) =
    let
      val t = termify ts;
      val subs = subsumers thy t regs;
    in (case subs of
        [] => NONE
      | ((t', (attn, thms)) :: _) => let
            val (tinst, inst) = Pattern.match thy (t', t) (Vartab.empty, Vartab.empty);
            (* thms contain Frees, not Vars *)
            val tinst' = tinst |> Vartab.dest
                 |> map (fn ((x, 0), (_, T)) => (x, Logic.legacy_unvarifyT T))
                 |> Symtab.make;
            val inst' = inst |> Vartab.dest
                 |> map (fn ((x, 0), (_, t)) => (x, Logic.legacy_unvarify t))
                 |> Symtab.make;
          in
            SOME (attn, map (Element.inst_witness thy (tinst', inst')) thms)
          end)
    end;

  (* add registration if not subsumed by ones already present,
     additionally returns registrations that are strictly subsumed *)
  fun insert sign (ts, attn) regs =
    let
      val t = termify ts;
      val subs = subsumers sign t regs ;
    in (case subs of
        [] => let
                val sups =
                  filter (fn (t', _) => Pattern.matches sign (t, t')) (Termtab.dest regs);
                val sups' = map (apfst untermify) sups
              in (Termtab.update (t, (attn, [])) regs, sups') end
      | _ => (regs, []))
    end;

  (* add witness theorem to registration,
     only if instantiation is exact, otherwise exception Option raised *)
  fun add_witness ts thm regs =
    let
      val t = termify ts;
      val (x, thms) = the (Termtab.lookup regs t);
    in
      Termtab.update (t, (x, thm::thms)) regs
    end;
end;


(** theory data **)

structure GlobalLocalesData = TheoryDataFun
(struct
  val name = "Isar/locales";
  type T = NameSpace.T * locale Symtab.table * Registrations.T Symtab.table;
    (* 1st entry: locale namespace,
       2nd entry: locales of the theory,
       3rd entry: registrations, indexed by locale name *)

  val empty = (NameSpace.empty, Symtab.empty, Symtab.empty);
  val copy = I;
  val extend = I;

  fun join_locs _ ({axiom, import, elems, params, lparams, term_syntax, regs, intros}: locale,
      {elems = elems', term_syntax = term_syntax', regs = regs', ...}: locale) =
     {axiom = axiom,
      import = import,
      elems = gen_merge_lists (eq_snd (op =)) elems elems',
      params = params,
      lparams = lparams,
      term_syntax = Library.merge (eq_snd (op =)) (term_syntax, term_syntax'),
      regs = merge_alists regs regs',
      intros = intros};
  fun merge _ ((space1, locs1, regs1), (space2, locs2, regs2)) =
    (NameSpace.merge (space1, space2), Symtab.join join_locs (locs1, locs2),
     Symtab.join (K Registrations.join) (regs1, regs2));

  fun print _ (space, locs, _) =
    Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (space, locs)))
    |> Pretty.writeln;
end);

val _ = Context.add_setup GlobalLocalesData.init;



(** context data **)

structure LocalLocalesData = ProofDataFun
(struct
  val name = "Isar/locales";
  type T = Registrations.T Symtab.table;
    (* registrations, indexed by locale name *)
  fun init _ = Symtab.empty;
  fun print _ _ = ();
end);

val _ = Context.add_setup LocalLocalesData.init;


(* access locales *)

val print_locales = GlobalLocalesData.print;

val intern = NameSpace.intern o #1 o GlobalLocalesData.get;
val extern = NameSpace.extern o #1 o GlobalLocalesData.get;

fun declare_locale name thy =
  thy |> GlobalLocalesData.map (fn (space, locs, regs) =>
    (Sign.declare_name thy name space, locs, regs));

fun put_locale name loc =
  GlobalLocalesData.map (fn (space, locs, regs) =>
    (space, Symtab.update (name, loc) locs, regs));

fun get_locale thy name = Symtab.lookup (#2 (GlobalLocalesData.get thy)) name;

fun the_locale thy name =
  (case get_locale thy name of
    SOME loc => loc
  | NONE => error ("Unknown locale " ^ quote name));

fun change_locale name f thy =
  let
    val {axiom, import, elems, params, lparams, term_syntax, regs, intros} =
        the_locale thy name;
    val (axiom', import', elems', params', lparams', term_syntax', regs', intros') =
      f (axiom, import, elems, params, lparams, term_syntax, regs, intros);
  in
    put_locale name {axiom = axiom', import = import', elems = elems',
      params = params', lparams = lparams', term_syntax = term_syntax', regs = regs',
      intros = intros'} thy
  end;


(* access registrations *)

(* Ids of global registrations are varified,
   Ids of local registrations are not.
   Thms of registrations are never varified. *)

(* retrieve registration from theory or context *)

fun gen_get_registrations get thy_ctxt name =
  case Symtab.lookup (get thy_ctxt) name of
      NONE => []
    | SOME reg => Registrations.dest reg;

val get_global_registrations =
     gen_get_registrations (#3 o GlobalLocalesData.get);
val get_local_registrations =
     gen_get_registrations LocalLocalesData.get;

fun gen_get_registration get thy_of thy_ctxt (name, ps) =
  case Symtab.lookup (get thy_ctxt) name of
      NONE => NONE
    | SOME reg => Registrations.lookup (thy_of thy_ctxt) (reg, ps);

val get_global_registration =
     gen_get_registration (#3 o GlobalLocalesData.get) I;
val get_local_registration =
     gen_get_registration LocalLocalesData.get ProofContext.theory_of;

val test_global_registration = is_some oo get_global_registration;
val test_local_registration = is_some oo get_local_registration;
fun smart_test_registration ctxt id =
  let
    val thy = ProofContext.theory_of ctxt;
  in
    test_global_registration thy id orelse test_local_registration ctxt id
  end;


(* add registration to theory or context, ignored if subsumed *)

fun gen_put_registration map_data thy_of (name, ps) attn thy_ctxt =
  map_data (fn regs =>
    let
      val thy = thy_of thy_ctxt;
      val reg = the_default Registrations.empty (Symtab.lookup regs name);
      val (reg', sups) = Registrations.insert thy (ps, attn) reg;
      val _ = if not (null sups) then warning
                ("Subsumed interpretation(s) of locale " ^
                 quote (extern thy name) ^
                 "\nby interpretation(s) with the following prefix(es):\n" ^
                  commas_quote (map (fn (_, ((s, _), _)) => s) sups))
              else ();
    in Symtab.update (name, reg') regs end) thy_ctxt;

val put_global_registration =
     gen_put_registration (fn f =>
       GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, f regs))) I;
val put_local_registration =
     gen_put_registration LocalLocalesData.map ProofContext.theory_of;

fun put_registration_in_locale name id =
  change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
    (axiom, import, elems, params, lparams, term_syntax, regs @ [(id, [])], intros));


(* add witness theorem to registration in theory or context,
   ignored if registration not present *)

fun add_witness (name, ps) thm =
  Symtab.map_entry name (Registrations.add_witness ps thm);

fun add_global_witness id thm =
  GlobalLocalesData.map (fn (space, locs, regs) => (space, locs, add_witness id thm regs));

val add_local_witness = LocalLocalesData.map oo add_witness;

fun add_witness_in_locale name id thm =
  change_locale name (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
    let
      fun add (id', thms) =
        if id = id' then (id', thm :: thms) else (id', thms);
    in (axiom, import, elems, params, lparams, term_syntax, map add regs, intros) end);


(* printing of registrations *)

fun gen_print_registrations get_regs mk_ctxt msg show_wits loc thy_ctxt =
  let
    val ctxt = mk_ctxt thy_ctxt;
    val thy = ProofContext.theory_of ctxt;

    val prt_term = Pretty.quote o ProofContext.pretty_term ctxt;
    fun prt_inst ts =
        Pretty.enclose "(" ")" (Pretty.breaks (map prt_term ts));
    fun prt_attn (prfx, atts) =
        Pretty.breaks (Pretty.str prfx :: Args.pretty_attribs ctxt atts);
    fun prt_witns witns = Pretty.enclose "[" "]"
      (Pretty.breaks (map (prt_term o Element.witness_prop) witns));
    fun prt_reg (ts, (("", []), witns)) =
        if show_wits
          then Pretty.block [prt_inst ts, Pretty.fbrk, prt_witns witns]
          else prt_inst ts
      | prt_reg (ts, (attn, witns)) =
        if show_wits
          then Pretty.block ((prt_attn attn @
            [Pretty.str ":", Pretty.brk 1, prt_inst ts, Pretty.fbrk,
              prt_witns witns]))
          else Pretty.block ((prt_attn attn @
            [Pretty.str ":", Pretty.brk 1, prt_inst ts]));

    val loc_int = intern thy loc;
    val regs = get_regs thy_ctxt;
    val loc_regs = Symtab.lookup regs loc_int;
  in
    (case loc_regs of
        NONE => Pretty.str ("no interpretations in " ^ msg)
      | SOME r => let
            val r' = Registrations.dest r;
            val r'' = Library.sort_wrt (fn (_, ((prfx, _), _)) => prfx) r';
          in Pretty.big_list ("interpretations in " ^ msg ^ ":")
            (map prt_reg r'')
          end)
    |> Pretty.writeln
  end;

val print_global_registrations =
     gen_print_registrations (#3 o GlobalLocalesData.get)
       ProofContext.init "theory";
val print_local_registrations' =
     gen_print_registrations LocalLocalesData.get
       I "context";
fun print_local_registrations show_wits loc ctxt =
  (print_global_registrations show_wits loc (ProofContext.theory_of ctxt);
   print_local_registrations' show_wits loc ctxt);


(* diagnostics *)

fun err_in_locale ctxt msg ids =
  let
    val thy = ProofContext.theory_of ctxt;
    fun prt_id (name, parms) =
      [Pretty.block (Pretty.breaks (map Pretty.str (extern thy name :: parms)))];
    val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));
    val err_msg =
      if forall (equal "" o #1) ids then msg
      else msg ^ "\n" ^ Pretty.string_of (Pretty.block
        (Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));
  in error err_msg end;

fun err_in_locale' ctxt msg ids' = err_in_locale ctxt msg (map fst ids');


fun pretty_ren NONE = Pretty.str "_"
  | pretty_ren (SOME (x, NONE)) = Pretty.str x
  | pretty_ren (SOME (x, SOME syn)) =
      Pretty.block [Pretty.str x, Pretty.brk 1, Syntax.pretty_mixfix syn];

fun pretty_expr thy (Locale name) = Pretty.str (extern thy name)
  | pretty_expr thy (Rename (expr, xs)) =
      Pretty.block [pretty_expr thy expr, Pretty.brk 1, Pretty.block (map pretty_ren xs |> Pretty.breaks)]
  | pretty_expr thy (Merge es) =
      Pretty.separate "+" (map (pretty_expr thy) es) |> Pretty.block;

fun err_in_expr _ msg (Merge []) = error msg
  | err_in_expr ctxt msg expr =
    error (msg ^ "\n" ^ Pretty.string_of (Pretty.block
      [Pretty.str "The error(s) above occured in locale expression:", Pretty.brk 1,
       pretty_expr (ProofContext.theory_of ctxt) expr]));


(** structured contexts: rename + merge + implicit type instantiation **)

(* parameter types *)

fun frozen_tvars ctxt Ts =
  #1 (Variable.importT_inst (map Logic.mk_type Ts) ctxt)
  |> map (fn ((xi, S), T) => (xi, (S, T)));

fun unify_frozen ctxt maxidx Ts Us =
  let
    fun paramify NONE i = (NONE, i)
      | paramify (SOME T) i = apfst SOME (TypeInfer.paramify_dummies T i);

    val (Ts', maxidx') = fold_map paramify Ts maxidx;
    val (Us', maxidx'') = fold_map paramify Us maxidx';
    val thy = ProofContext.theory_of ctxt;

    fun unify (SOME T, SOME U) env = (Sign.typ_unify thy (U, T) env
          handle Type.TUNIFY => raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
      | unify _ env = env;
    val (unifier, _) = fold unify (Ts' ~~ Us') (Vartab.empty, maxidx'');
    val Vs = map (Option.map (Envir.norm_type unifier)) Us';
    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map_filter I Vs));
  in map (Option.map (Envir.norm_type unifier')) Vs end;

fun params_of elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst) elemss);
fun params_of' elemss = distinct (eq_fst (op = : string * string -> bool)) (maps (snd o fst o fst) elemss);
fun params_syn_of syn elemss =
  distinct (eq_fst (op =)) (maps (snd o fst) elemss) |>
    map (apfst (fn x => (x, the (Symtab.lookup syn x))));


(* CB: param_types has the following type:
  ('a * 'b option) list -> ('a * 'b) list *)
fun param_types ps = map_filter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;


fun merge_syntax ctxt ids ss = Symtab.merge (op = : mixfix * mixfix -> bool) ss
  handle Symtab.DUPS xs => err_in_locale ctxt
    ("Conflicting syntax for parameter(s): " ^ commas_quote xs) (map fst ids);


(* Distinction of assumed vs. derived identifiers.
   The former may have axioms relating assumptions of the context to
   assumptions of the specification fragment (for locales with
   predicates).  The latter have witnesses relating assumptions of the
   specification fragment to assumptions of other (assumed) specification
   fragments. *)

datatype 'a mode = Assumed of 'a | Derived of 'a;

fun map_mode f (Assumed x) = Assumed (f x)
  | map_mode f (Derived x) = Derived (f x);


(* flatten expressions *)

local

fun unify_parms ctxt fixed_parms raw_parmss =
  let
    val thy = ProofContext.theory_of ctxt;
    val maxidx = length raw_parmss;
    val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;

    fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));
    fun varify_parms (i, ps) = map (apsnd (varify i)) (param_types ps);
    val parms = fixed_parms @ maps varify_parms idx_parmss;

    fun unify T U envir = Sign.typ_unify thy (U, T) envir
      handle Type.TUNIFY =>
        let val prt = Sign.string_of_typ thy in
          raise TYPE ("unify_parms: failed to unify types " ^
            prt U ^ " and " ^ prt T, [U, T], [])
        end;
    fun unify_list (T :: Us) = fold (unify T) Us
      | unify_list [] = I;
    val (unifier, _) = fold unify_list (map #2 (Symtab.dest (Symtab.make_list parms)))
      (Vartab.empty, maxidx);

    val parms' = map (apsnd (Envir.norm_type unifier)) (distinct (eq_fst (op =)) parms);
    val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms'));

    fun inst_parms (i, ps) =
      foldr Term.add_typ_tfrees [] (map_filter snd ps)
      |> map_filter (fn (a, S) =>
          let val T = Envir.norm_type unifier' (TVar ((a, i), S))
          in if T = TFree (a, S) then NONE else SOME (a, T) end)
      |> Symtab.make;
  in map inst_parms idx_parmss end;

in

fun unify_elemss _ _ [] = []
  | unify_elemss _ [] [elems] = [elems]
  | unify_elemss ctxt fixed_parms elemss =
      let
        val thy = ProofContext.theory_of ctxt;
        val envs = unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss);
        fun inst ((((name, ps), mode), elems), env) =
          (((name, map (apsnd (Option.map (Element.instT_type env))) ps),
              map_mode (map (Element.instT_witness thy env)) mode),
            map (Element.instT_ctxt thy env) elems);
      in map inst (elemss ~~ envs) end;

(* like unify_elemss, but does not touch mode, additional
   parameter c_parms for enforcing further constraints (eg. syntax) *)
(* FIXME avoid code duplication *)

fun unify_elemss' _ _ [] [] = []
  | unify_elemss' _ [] [elems] [] = [elems]
  | unify_elemss' ctxt fixed_parms elemss c_parms =
      let
        val thy = ProofContext.theory_of ctxt;
        val envs =
          unify_parms ctxt fixed_parms (map (snd o fst o fst) elemss @ map single c_parms);
        fun inst ((((name, ps), (ps', mode)), elems), env) =
          (((name, map (apsnd (Option.map (Element.instT_type env))) ps), (ps', mode)),
           map (Element.instT_ctxt thy env) elems);
      in map inst (elemss ~~ Library.take (length elemss, envs)) end;


fun renaming xs parms = zip_options parms xs
  handle Library.UnequalLengths =>
    error ("Too many arguments in renaming: " ^
      commas (map (fn NONE => "_" | SOME x => quote (fst x)) xs));


(* params_of_expr:
   Compute parameters (with types and syntax) of locale expression.
*)

fun params_of_expr ctxt fixed_params expr (prev_parms, prev_types, prev_syn) =
  let
    val thy = ProofContext.theory_of ctxt;

    fun merge_tenvs fixed tenv1 tenv2 =
        let
          val [env1, env2] = unify_parms ctxt fixed
                [tenv1 |> Symtab.dest |> map (apsnd SOME),
                 tenv2 |> Symtab.dest |> map (apsnd SOME)]
        in
          Symtab.merge (op =) (Symtab.map (Element.instT_type env1) tenv1,
            Symtab.map (Element.instT_type env2) tenv2)
        end;

    fun merge_syn expr syn1 syn2 =
        Symtab.merge (op = : mixfix * mixfix -> bool) (syn1, syn2)
        handle Symtab.DUPS xs => err_in_expr ctxt
          ("Conflicting syntax for parameter(s): " ^ commas_quote xs) expr;
            
    fun params_of (expr as Locale name) =
          let
            val {import, params, ...} = the_locale thy name;
            val parms = map (fst o fst) params;
            val (parms', types', syn') = params_of import;
            val all_parms = merge_lists parms' parms;
            val all_types = merge_tenvs [] types' (params |> map fst |> Symtab.make);
            val all_syn = merge_syn expr syn' (params |> map (apfst fst) |> Symtab.make);
          in (all_parms, all_types, all_syn) end
      | params_of (expr as Rename (e, xs)) =
          let
            val (parms', types', syn') = params_of e;
            val ren = renaming xs parms';
            (* renaming may reduce number of parameters *)
            val new_parms = map (Element.rename ren) parms' |> distinct (op =);
            val ren_syn = syn' |> Symtab.dest |> map (Element.rename_var ren);
            val new_syn = fold (Symtab.insert (op =)) ren_syn Symtab.empty
                handle Symtab.DUP x =>
                  err_in_expr ctxt ("Conflicting syntax for parameter: " ^ quote x) expr;
            val syn_types = map (apsnd (fn mx => SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (TypeInfer.mixfixT mx) 0))))) (Symtab.dest new_syn);
            val ren_types = types' |> Symtab.dest |> map (apfst (Element.rename ren));
            val (env :: _) = unify_parms ctxt [] 
                ((ren_types |> map (apsnd SOME)) :: map single syn_types);
            val new_types = fold (Symtab.insert (op =))
                (map (apsnd (Element.instT_type env)) ren_types) Symtab.empty;
          in (new_parms, new_types, new_syn) end
      | params_of (Merge es) =
          fold (fn e => fn (parms, types, syn) =>
                   let
                     val (parms', types', syn') = params_of e
                   in
                     (merge_lists parms parms', merge_tenvs [] types types',
                      merge_syn e syn syn')
                   end)
            es ([], Symtab.empty, Symtab.empty)

      val (parms, types, syn) = params_of expr;
    in
      (merge_lists prev_parms parms, merge_tenvs fixed_params prev_types types,
       merge_syn expr prev_syn syn)
    end;

fun make_params_ids params = [(("", params), ([], Assumed []))];
fun make_raw_params_elemss (params, tenv, syn) =
    [((("", map (fn p => (p, Symtab.lookup tenv p)) params), Assumed []),
      Int [Fixes (map (fn p =>
        (p, Symtab.lookup tenv p, Symtab.lookup syn p |> the)) params)])];


(* flatten_expr:
   Extend list of identifiers by those new in locale expression expr.
   Compute corresponding list of lists of locale elements (one entry per
   identifier).

   Identifiers represent locale fragments and are in an extended form:
     ((name, ps), (ax_ps, axs))
   (name, ps) is the locale name with all its parameters.
   (ax_ps, axs) is the locale axioms with its parameters;
     axs are always taken from the top level of the locale hierarchy,
     hence axioms may contain additional parameters from later fragments:
     ps subset of ax_ps.  axs is either singleton or empty.

   Elements are enriched by identifier-like information:
     (((name, ax_ps), axs), elems)
   The parameters in ax_ps are the axiom parameters, but enriched by type
   info: now each entry is a pair of string and typ option.  Axioms are
   type-instantiated.

*)

fun flatten_expr ctxt ((prev_idents, prev_syntax), expr) =
  let
    val thy = ProofContext.theory_of ctxt;

    fun rename_parms top ren ((name, ps), (parms, mode)) =
        ((name, map (Element.rename ren) ps),
         if top
         then (map (Element.rename ren) parms,
               map_mode (map (Element.rename_witness ren)) mode)
         else (parms, mode));

    (* add registrations of (name, ps), recursively; adjust hyps of witnesses *)

    fun add_regs (name, ps) (ths, ids) =
        let
          val {params, regs, ...} = the_locale thy name;
          val ps' = map #1 params;
          val ren = map #1 ps' ~~ map (fn (x, _) => (x, NONE)) ps;
            (* dummy syntax, since required by rename *)
          val ps'' = map (fn ((p, _), (_, T)) => (p, T)) (ps ~~ ps');
          val [env] = unify_parms ctxt ps [map (apsnd SOME) ps''];
            (* propagate parameter types, to keep them consistent *)
          val regs' = map (fn ((name, ps), ths) =>
              ((name, map (Element.rename ren) ps), ths)) regs;
          val new_regs = gen_rems (eq_fst (op =)) (regs', ids);
          val new_ids = map fst new_regs;
          val new_idTs = map (apsnd (map (fn p => (p, (the o AList.lookup (op =) ps) p)))) new_ids;

          val new_ths = new_regs |> map (#2 #> map
            (Element.instT_witness thy env #>
              Element.rename_witness ren #>
              Element.satisfy_witness ths));
          val new_ids' = map (fn (id, ths) =>
              (id, ([], Derived ths))) (new_ids ~~ new_ths);
        in
          fold add_regs new_idTs (ths @ flat new_ths, ids @ new_ids')
        end;

    (* distribute top-level axioms over assumed ids *)

    fun axiomify all_ps ((name, parms), (_, Assumed _)) axioms =
        let
          val {elems, ...} = the_locale thy name;
          val ts = maps
            (fn (Assumes asms, _) => maps (map #1 o #2) asms
              | _ => [])
            elems;
          val (axs1, axs2) = chop (length ts) axioms;
        in (((name, parms), (all_ps, Assumed axs1)), axs2) end
      | axiomify all_ps (id, (_, Derived ths)) axioms =
          ((id, (all_ps, Derived ths)), axioms);

    (* identifiers of an expression *)

    fun identify top (Locale name) =
    (* CB: ids_ax is a list of tuples of the form ((name, ps), axs),
       where name is a locale name, ps a list of parameter names and axs
       a list of axioms relating to the identifier, axs is empty unless
       identify at top level (top = true);
       parms is accumulated list of parameters *)
          let
            val {axiom, import, params, ...} =
              the_locale thy name;
            val ps = map (#1 o #1) params;
            val (ids', parms', _) = identify false import;
                (* acyclic import dependencies *)

            val ids'' = ids' @ [((name, ps), ([], Assumed []))];
            val (_, ids''') = add_regs (name, map #1 params) ([], ids'');
            val (_, ids4) = chop (length ids' + 1) ids''';
            val ids5 = ids' @ ids4 @ [((name, ps), ([], Assumed []))];
            val ids_ax = if top then fst (fold_map (axiomify ps) ids5 axiom) else ids5;
            val syn = Symtab.make (map (apfst fst) params);
            in (ids_ax, merge_lists parms' ps, syn) end
      | identify top (Rename (e, xs)) =
          let
            val (ids', parms', syn') = identify top e;
            val ren = renaming xs parms'
              handle ERROR msg => err_in_locale' ctxt msg ids';

            val ids'' = distinct (eq_fst (op =)) (map (rename_parms top ren) ids');
            val parms'' = distinct (op =) (maps (#2 o #1) ids'');
            val syn'' = fold (Symtab.insert (op =))
                (map (Element.rename_var ren) (Symtab.dest syn')) Symtab.empty;
          in (ids'', parms'', syn'') end
      | identify top (Merge es) =
          fold (fn e => fn (ids, parms, syn) =>
                   let
                     val (ids', parms', syn') = identify top e
                   in
                     (merge_alists ids ids',
                      merge_lists parms parms',
                      merge_syntax ctxt ids' (syn, syn'))
                   end)
            es ([], [], Symtab.empty);


    (* CB: enrich identifiers by parameter types and
       the corresponding elements (with renamed parameters),
       also takes care of parameter syntax *)

    fun eval syn ((name, xs), axs) =
      let
        val {params = ps, lparams = qs, elems, ...} = the_locale thy name;
        val ps' = map (apsnd SOME o #1) ps;
        fun lookup_syn x = (case Symtab.lookup syn x of SOME Structure => NONE | opt => opt);
        val ren = map #1 ps' ~~ map (fn x => (x, lookup_syn x)) xs;
        val (params', elems') =
          if null ren then (ps', map #1 elems)
          else (map (apfst (Element.rename ren)) ps',
            map (Element.rename_ctxt ren o #1) elems);
        val elems'' = elems' |> map (Element.map_ctxt
          {var = I, typ = I, term = I, fact = I, attrib = I,
            name = NameSpace.qualified (space_implode "_" xs)});
      in (((name, params'), axs), elems'') end;

    (* type constraint for renamed parameter with syntax *)
    fun mixfix_type mx =
      SOME (Type.freeze_type (#1 (TypeInfer.paramify_dummies (TypeInfer.mixfixT mx) 0)));

    (* compute identifiers and syntax, merge with previous ones *)
    val (ids, _, syn) = identify true expr;
    val idents = gen_rems (eq_fst (op =)) (ids, prev_idents);
    val syntax = merge_syntax ctxt ids (syn, prev_syntax);
    (* add types to params and unify them *)
    val raw_elemss = map (eval syntax) idents;
    val elemss = unify_elemss' ctxt [] raw_elemss (map (apsnd mixfix_type) (Symtab.dest syntax));
    (* replace params in ids by params from axioms,
       adjust types in mode *)
    val all_params' = params_of' elemss;
    val all_params = param_types all_params';
    val elemss' = map (fn (((name, _), (ps, mode)), elems) =>
         (((name, map (fn p => (p, AList.lookup (op =) all_params p)) ps), mode), elems))
         elemss;
    fun inst_th (t, th) = let
         val {hyps, prop, ...} = Thm.rep_thm th;
         val ps = map (apsnd SOME) (fold Term.add_frees (prop :: hyps) []);
         val [env] = unify_parms ctxt all_params [ps];
         val t' = Element.instT_term env t;
         val th' = Element.instT_thm thy env th;
       in (t', th') end;
    val final_elemss = map (fn ((id, mode), elems) =>
         ((id, map_mode (map (Element.map_witness inst_th)) mode), elems)) elemss';

  in ((prev_idents @ idents, syntax), final_elemss) end;

end;


(* activate elements *)

local

fun axioms_export axs _ hyps =
  Element.satisfy_thm axs
  #> Drule.implies_intr_list (Library.drop (length axs, hyps))
  #> Seq.single;


(* NB: derived ids contain only facts at this stage *)

fun activate_elem _ _ ((ctxt, mode), Fixes fixes) =
      ((ctxt |> ProofContext.add_fixes_i fixes |> snd, mode), [])
  | activate_elem _ _ ((ctxt, mode), Constrains _) =
      ((ctxt, mode), [])
  | activate_elem ax_in_ctxt _ ((ctxt, Assumed axs), Assumes asms) =
      let
        val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) asms;
        val ts = maps (map #1 o #2) asms';
        val (ps, qs) = chop (length ts) axs;
        val (_, ctxt') =
          ctxt |> fold ProofContext.fix_frees ts
          |> ProofContext.add_assms_i (axioms_export (if ax_in_ctxt then ps else [])) asms';
      in ((ctxt', Assumed qs), []) end
  | activate_elem _ _ ((ctxt, Derived ths), Assumes asms) =
      ((ctxt, Derived ths), [])
  | activate_elem _ _ ((ctxt, Assumed axs), Defines defs) =
      let
        val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs;
        val asms = defs' |> map (fn ((name, atts), (t, ps)) =>
            let val ((c, _), t') = LocalDefs.cert_def ctxt t
            in (t', ((if name = "" then Thm.def_name c else name, atts), [(t', ps)])) end);
        val (_, ctxt') =
          ctxt |> fold (ProofContext.fix_frees o #1) asms
          |> ProofContext.add_assms_i LocalDefs.def_export (map #2 asms);
      in ((ctxt', Assumed axs), []) end
  | activate_elem _ _ ((ctxt, Derived ths), Defines defs) =
      ((ctxt, Derived ths), [])
  | activate_elem _ is_ext ((ctxt, mode), Notes facts) =
      let
        val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts;
        val (res, ctxt') = ctxt |> ProofContext.note_thmss_i facts';
      in ((ctxt', mode), if is_ext then res else []) end;

fun activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt =
  let
    val thy = ProofContext.theory_of ctxt;
    val ((ctxt', _), res) =
        foldl_map (activate_elem ax_in_ctxt (name = "")) ((ProofContext.qualified_names ctxt, mode), elems)
      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]
    val ctxt'' = if name = "" then ctxt'
          else let
              val ps' = map (fn (n, SOME T) => Free (n, T)) ps;
              val ctxt'' = put_local_registration (name, ps') ("", []) ctxt'
            in case mode of
                Assumed axs =>
                  fold (add_local_witness (name, ps') o
                    Element.assume_witness thy o Element.witness_prop) axs ctxt''
              | Derived ths => fold (add_local_witness (name, ps')) ths ctxt''
            end
  in (ProofContext.restore_naming ctxt ctxt'', res) end;

fun activate_elemss ax_in_ctxt prep_facts =
    fold_map (fn (((name, ps), mode), raw_elems) => fn ctxt =>
      let
        val elems = map (prep_facts ctxt) raw_elems;
        val (ctxt', res) = apsnd flat
            (activate_elems ax_in_ctxt (((name, ps), mode), elems) ctxt);
        val elems' = elems |> map (Element.map_ctxt
          {name = I, var = I, typ = I, term = I, fact = I, attrib = Args.closure});
      in (((((name, ps), mode), elems'), res), ctxt') end);

in

(* CB: activate_facts prep_facts (ctxt, elemss),
   where elemss is a list of pairs consisting of identifiers and
   context elements, extends ctxt by the context elements yielding
   ctxt' and returns (ctxt', (elemss', facts)).
   Identifiers in the argument are of the form ((name, ps), axs) and
   assumptions use the axioms in the identifiers to set up exporters
   in ctxt'.  elemss' does not contain identifiers and is obtained
   from elemss and the intermediate context with prep_facts.
   If read_facts or cert_facts is used for prep_facts, these also remove
   the internal/external markers from elemss. *)

fun activate_facts ax_in_ctxt prep_facts (ctxt, args) =
  let val ((elemss, factss), ctxt') = activate_elemss ax_in_ctxt prep_facts args ctxt |>> split_list
  in (ctxt', (elemss, flat factss)) end;

end;



(** prepare locale elements **)

(* expressions *)

fun intern_expr thy (Locale xname) = Locale (intern thy xname)
  | intern_expr thy (Merge exprs) = Merge (map (intern_expr thy) exprs)
  | intern_expr thy (Rename (expr, xs)) = Rename (intern_expr thy expr, xs);


(* propositions and bindings *)

(* flatten (ctxt, prep_expr) ((ids, syn), expr)
   normalises expr (which is either a locale
   expression or a single context element) wrt.
   to the list ids of already accumulated identifiers.
   It returns ((ids', syn'), elemss) where ids' is an extension of ids
   with identifiers generated for expr, and elemss is the list of
   context elements generated from expr.
   syn and syn' are symtabs mapping parameter names to their syntax.  syn'
   is an extension of syn.
   For details, see flatten_expr.

   Additionally, for a locale expression, the elems are grouped into a single
   Int; individual context elements are marked Ext.  In this case, the
   identifier-like information of the element is as follows:
   - for Fixes: (("", ps), []) where the ps have type info NONE
   - for other elements: (("", []), []).
   The implementation of activate_facts relies on identifier names being
   empty strings for external elements.
*)

fun flatten (ctxt, _) ((ids, syn), Elem (Fixes fixes)) = let
        val ids' = ids @ [(("", map #1 fixes), ([], Assumed []))]
      in
        ((ids',
         merge_syntax ctxt ids'
           (syn, Symtab.make (map (fn fx => (#1 fx, #3 fx)) fixes))
           handle Symtab.DUPS xs => err_in_locale ctxt
             ("Conflicting syntax for parameters: " ^ commas_quote xs)
             (map #1 ids')),
         [((("", map (rpair NONE o #1) fixes), Assumed []), Ext (Fixes fixes))])
      end
  | flatten _ ((ids, syn), Elem elem) =
      ((ids @ [(("", []), ([], Assumed []))], syn), [((("", []), Assumed []), Ext elem)])
  | flatten (ctxt, prep_expr) ((ids, syn), Expr expr) =
      apsnd (map (apsnd Int)) (flatten_expr ctxt ((ids, syn), prep_expr expr));

local

local

fun declare_int_elem (ctxt, Fixes fixes) =
      (ctxt |> ProofContext.add_fixes_i (map (fn (x, T, mx) =>
        (x, Option.map (Term.map_type_tfree (TypeInfer.param 0)) T, mx)) fixes) |> snd, [])
  | declare_int_elem (ctxt, _) = (ctxt, []);

fun declare_ext_elem prep_vars (ctxt, Fixes fixes) =
      let val (vars, _) = prep_vars fixes ctxt
      in (ctxt |> ProofContext.add_fixes_i vars |> snd, []) end
  | declare_ext_elem prep_vars (ctxt, Constrains csts) =
      let val (_, ctxt') = prep_vars (map (fn (x, T) => (x, SOME T, NoSyn)) csts) ctxt
      in (ctxt', []) end
  | declare_ext_elem _ (ctxt, Assumes asms) = (ctxt, map #2 asms)
  | declare_ext_elem _ (ctxt, Defines defs) = (ctxt, map (fn (_, (t, ps)) => [(t, ps)]) defs)
  | declare_ext_elem _ (ctxt, Notes facts) = (ctxt, []);

fun declare_elems prep_vars (ctxt, (((name, ps), Assumed _), elems)) =
    let val (ctxt', propps) =
      (case elems of
        Int es => foldl_map declare_int_elem (ctxt, es)
      | Ext e => foldl_map (declare_ext_elem prep_vars) (ctxt, [e]))
      handle ERROR msg => err_in_locale ctxt msg [(name, map fst ps)]
    in (ctxt', propps) end
  | declare_elems _ (ctxt, ((_, Derived _), elems)) = (ctxt, []);

in

fun declare_elemss prep_vars fixed_params raw_elemss ctxt =
  let
    (* CB: fix of type bug of goal in target with context elements.
       Parameters new in context elements must receive types that are
       distinct from types of parameters in target (fixed_params).  *)
    val ctxt_with_fixed =
      fold Variable.declare_term (map Free fixed_params) ctxt;
    val int_elemss =
      raw_elemss
      |> map_filter (fn (id, Int es) => SOME (id, es) | _ => NONE)
      |> unify_elemss ctxt_with_fixed fixed_params;
    val (_, raw_elemss') =
      foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x)
        (int_elemss, raw_elemss);
  in foldl_map (declare_elems prep_vars) (ctxt, raw_elemss') end;

end;

local

val norm_term = Envir.beta_norm oo Term.subst_atomic;

fun abstract_thm thy eq =
  Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;

fun bind_def ctxt (name, ps) eq (xs, env, ths) =
  let
    val ((y, T), b) = LocalDefs.abs_def eq;
    val b' = norm_term env b;
    val th = abstract_thm (ProofContext.theory_of ctxt) eq;
    fun err msg = err_in_locale ctxt (msg ^ ": " ^ quote y) [(name, map fst ps)];
  in
    conditional (exists (equal y o #1) xs) (fn () =>
      err "Attempt to define previously specified variable");
    conditional (exists (fn (Free (y', _), _) => y = y' | _ => false) env) (fn () =>
      err "Attempt to redefine variable");
    (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
  end;


(* CB: for finish_elems (Int and Ext),
   extracts specification, only of assumed elements *)

fun eval_text _ _ _ (Fixes _) text = text
  | eval_text _ _ _ (Constrains _) text = text
  | eval_text _ (_, Assumed _) is_ext (Assumes asms)
        (((exts, exts'), (ints, ints')), (xs, env, defs)) =
      let
        val ts = maps (map #1 o #2) asms;
        val ts' = map (norm_term env) ts;
        val spec' =
          if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
          else ((exts, exts'), (ints @ ts, ints' @ ts'));
      in (spec', (fold Term.add_frees ts' xs, env, defs)) end
  | eval_text _ (_, Derived _) _ (Assumes _) text = text
  | eval_text ctxt (id, Assumed _) _ (Defines defs) (spec, binds) =
      (spec, fold (bind_def ctxt id o #1 o #2) defs binds)
  | eval_text _ (_, Derived _) _ (Defines _) text = text
  | eval_text _ _ _ (Notes _) text = text;


(* for finish_elems (Int),
   remove redundant elements of derived identifiers,
   turn assumptions and definitions into facts,
   adjust hypotheses of facts using witnesses *)

fun finish_derived _ _ (Assumed _) (Fixes fixes) = SOME (Fixes fixes)
  | finish_derived _ _ (Assumed _) (Constrains csts) = SOME (Constrains csts)
  | finish_derived _ _ (Assumed _) (Assumes asms) = SOME (Assumes asms)
  | finish_derived _ _ (Assumed _) (Defines defs) = SOME (Defines defs)

  | finish_derived _ _ (Derived _) (Fixes _) = NONE
  | finish_derived _ _ (Derived _) (Constrains _) = NONE
  | finish_derived sign wits (Derived _) (Assumes asms) = asms
      |> map (apsnd (map (fn (a, _) => ([Thm.assume (cterm_of sign a)], []))))
      |> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME
  | finish_derived sign wits (Derived _) (Defines defs) = defs
      |> map (apsnd (fn (d, _) => [([Thm.assume (cterm_of sign d)], [])]))
      |> Notes |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME

  | finish_derived _ wits _ (Notes facts) = (Notes facts)
      |> Element.map_ctxt_values I I (Element.satisfy_thm wits) |> SOME;

(* CB: for finish_elems (Ext) *)

fun closeup _ false elem = elem
  | closeup ctxt true elem =
      let
        fun close_frees t =
          let val frees = rev (filter_out (Variable.is_fixed ctxt o #1) (Term.add_frees t []))
          in Term.list_all_free (frees, t) end;

        fun no_binds [] = []
          | no_binds _ = error "Illegal term bindings in locale element";
      in
        (case elem of
          Assumes asms => Assumes (asms |> map (fn (a, propps) =>
            (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
        | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
            (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
        | e => e)
      end;


fun finish_ext_elem parms _ (Fixes fixes, _) = Fixes (map (fn (x, _, mx) =>
      (x, AList.lookup (op =) parms x, mx)) fixes)
  | finish_ext_elem parms _ (Constrains _, _) = Constrains []
  | finish_ext_elem _ close (Assumes asms, propp) =
      close (Assumes (map #1 asms ~~ propp))
  | finish_ext_elem _ close (Defines defs, propp) =
      close (Defines (map #1 defs ~~ map (fn [(t, ps)] => (t, ps)) propp))
  | finish_ext_elem _ _ (Notes facts, _) = Notes facts;


(* CB: finish_parms introduces type info from parms to identifiers *)
(* CB: only needed for types that have been NONE so far???
   If so, which are these??? *)

fun finish_parms parms (((name, ps), mode), elems) =
  (((name, map (fn (x, _) => (x, AList.lookup (op = : string * string -> bool) parms x)) ps), mode), elems);

fun finish_elems ctxt parms _ ((text, wits), ((id, Int e), _)) =
      let
        val [(id' as (_, mode), es)] = unify_elemss ctxt parms [(id, e)];
        val wits' = case mode of Assumed _ => wits | Derived ths => wits @ ths;
        val text' = fold (eval_text ctxt id' false) es text;
        val es' = map_filter
              (finish_derived (ProofContext.theory_of ctxt) wits' mode) es;
      in ((text', wits'), (id', map Int es')) end
  | finish_elems ctxt parms do_close ((text, wits), ((id, Ext e), [propp])) =
      let
        val e' = finish_ext_elem parms (closeup ctxt do_close) (e, propp);
        val text' = eval_text ctxt id true e' text;
      in ((text', wits), (id, [Ext e'])) end

in

(* CB: only called by prep_elemss *)

fun finish_elemss ctxt parms do_close =
  foldl_map (apsnd (finish_parms parms) o finish_elems ctxt parms do_close);

end;


(* Remove duplicate Defines elements: temporary workaround to fix Afp/Category. *)

fun defs_ord (defs1, defs2) =
    list_ord (fn ((_, (d1, _)), (_, (d2, _))) =>
      Term.fast_term_ord (d1, d2)) (defs1, defs2);
structure Defstab =
    TableFun(type key = ((string * Attrib.src list) * (term * term list)) list
        val ord = defs_ord);

fun rem_dup_defs es ds =
    fold_map (fn e as (Defines defs) => (fn ds =>
                 if Defstab.defined ds defs
                 then (Defines [], ds)
                 else (e, Defstab.update (defs, ()) ds))
               | e => (fn ds => (e, ds))) es ds;
fun rem_dup_elemss (Int es) ds = apfst Int (rem_dup_defs es ds)
  | rem_dup_elemss (Ext e) ds = (Ext e, ds);
fun rem_dup_defines raw_elemss =
    fold_map (fn (id as (_, (Assumed _)), es) => (fn ds =>
                     apfst (pair id) (rem_dup_elemss es ds))
               | (id as (_, (Derived _)), es) => (fn ds =>
                     ((id, es), ds))) raw_elemss Defstab.empty |> #1;

(* CB: type inference and consistency checks for locales.

   Works by building a context (through declare_elemss), extracting the
   required information and adjusting the context elements (finish_elemss).
   Can also universally close free vars in assms and defs.  This is only
   needed for Ext elements and controlled by parameter do_close.

   Only elements of assumed identifiers are considered.
*)

fun prep_elemss prep_vars prepp do_close context fixed_params raw_elemss raw_concl =
  let
    (* CB: contexts computed in the course of this function are discarded.
       They are used for type inference and consistency checks only. *)
    (* CB: fixed_params are the parameters (with types) of the target locale,
       empty list if there is no target. *)
    (* CB: raw_elemss are list of pairs consisting of identifiers and
       context elements, the latter marked as internal or external. *)
    val raw_elemss = rem_dup_defines raw_elemss;
    val (raw_ctxt, raw_proppss) = declare_elemss prep_vars fixed_params raw_elemss context;
    (* CB: raw_ctxt is context with additional fixed variables derived from
       the fixes elements in raw_elemss,
       raw_proppss contains assumptions and definitions from the
       external elements in raw_elemss. *)
    fun prep_prop raw_propp (raw_ctxt, raw_concl)  =
      let
        (* CB: add type information from fixed_params to context (declare_term) *)
        (* CB: process patterns (conclusion and external elements only) *)
        val (ctxt, all_propp) =
          prepp (fold Variable.declare_term (map Free fixed_params) raw_ctxt, raw_concl @ raw_propp);
        (* CB: add type information from conclusion and external elements to context *)
        val ctxt = fold Variable.declare_term (maps (map fst) all_propp) ctxt;
        (* CB: resolve schematic variables (patterns) in conclusion and external elements. *)
        val all_propp' = map2 (curry (op ~~))
          (#1 (#2 (ProofContext.bind_propp_schematic_i (ctxt, all_propp)))) (map (map snd) all_propp);
        val (concl, propp) = chop (length raw_concl) all_propp';
      in (propp, (ctxt, concl)) end

    val (proppss, (ctxt, concl)) =
      (fold_burrow o fold_burrow) prep_prop raw_proppss (raw_ctxt, raw_concl);

    (* CB: obtain all parameters from identifier part of raw_elemss *)
    val xs = map #1 (params_of' raw_elemss);
    val typing = unify_frozen ctxt 0
      (map (Variable.default_type raw_ctxt) xs)
      (map (Variable.default_type ctxt) xs);
    val parms = param_types (xs ~~ typing);
    (* CB: parms are the parameters from raw_elemss, with correct typing. *)

    (* CB: extract information from assumes and defines elements
       (fixes, constrains and notes in raw_elemss don't have an effect on
       text and elemss), compute final form of context elements. *)
    val ((text, _), elemss) = finish_elemss ctxt parms do_close
      ((((([], []), ([], [])), ([], [], [])), []), raw_elemss ~~ proppss);
    (* CB: text has the following structure:
           (((exts, exts'), (ints, ints')), (xs, env, defs))
       where
         exts: external assumptions (terms in external assumes elements)
         exts': dito, normalised wrt. env
         ints: internal assumptions (terms in internal assumes elements)
         ints': dito, normalised wrt. env
         xs: the free variables in exts' and ints' and rhss of definitions,
           this includes parameters except defined parameters
         env: list of term pairs encoding substitutions, where the first term
           is a free variable; substitutions represent defines elements and
           the rhs is normalised wrt. the previous env
         defs: theorems representing the substitutions from defines elements
           (thms are normalised wrt. env).
       elemss is an updated version of raw_elemss:
         - type info added to Fixes and modified in Constrains
         - axiom and definition statement replaced by corresponding one
           from proppss in Assumes and Defines
         - Facts unchanged
       *)
  in ((parms, elemss, concl), text) end;

in

fun read_elemss x = prep_elemss ProofContext.read_vars ProofContext.read_propp_schematic x;
fun cert_elemss x = prep_elemss ProofContext.cert_vars ProofContext.cert_propp_schematic x;

end;


(* facts and attributes *)

local

fun prep_name name =
  if NameSpace.is_qualified name then error ("Illegal qualified name: " ^ quote name)
  else name;

fun prep_facts _ _ ctxt (Int elem) =
      Element.map_ctxt_values I I (Thm.transfer (ProofContext.theory_of ctxt)) elem
  | prep_facts get intern ctxt (Ext elem) = elem |> Element.map_ctxt
     {var = I, typ = I, term = I,
      name = prep_name,
      fact = get ctxt,
      attrib = Args.assignable o intern (ProofContext.theory_of ctxt)};

in

fun read_facts x = prep_facts ProofContext.get_thms Attrib.intern_src x;
fun cert_facts x = prep_facts (K I) (K I) x;

end;


(* Get the specification of a locale *)

(*The global specification is made from the parameters and global
  assumptions, the local specification from the parameters and the
  local assumptions.*)

local

fun gen_asms_of get thy name =
  let
    val ctxt = ProofContext.init thy;
    val (_, raw_elemss) = flatten (ctxt, I) (([], Symtab.empty), Expr (Locale name));
    val ((_, elemss, _), _) = read_elemss false ctxt [] raw_elemss [];
  in
    elemss |> get
      |> maps (fn (_, es) => map (fn Int e => e) es)
      |> maps (fn Assumes asms => asms | _ => [])
      |> map (apsnd (map fst))
  end;

in

fun parameters_of thy name =
  the_locale thy name |> #params;

fun parameters_of_expr thy expr =
  let
    val ctxt = ProofContext.init thy;
    val pts = params_of_expr ctxt [] (intern_expr thy expr)
        ([], Symtab.empty, Symtab.empty);
    val raw_params_elemss = make_raw_params_elemss pts;
    val ((_, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
        (([], Symtab.empty), Expr expr);
    val ((parms, _, _), _) =
        read_elemss false ctxt [] (raw_params_elemss @ raw_elemss) [];
  in map (fn p as (n, _) => (p, Symtab.lookup syn n |> the)) parms end;

fun local_asms_of thy name =
  gen_asms_of (single o Library.last_elem) thy name;

fun global_asms_of thy name =
  gen_asms_of I thy name;

end;


(* full context statements: import + elements + conclusion *)

local

fun prep_context_statement prep_expr prep_elemss prep_facts
    do_close fixed_params import elements raw_concl context =
  let
    val thy = ProofContext.theory_of context;

    val (import_params, import_tenv, import_syn) =
      params_of_expr context fixed_params (prep_expr thy import)
        ([], Symtab.empty, Symtab.empty);
    val includes = map_filter (fn Expr e => SOME e | Elem _ => NONE) elements;
    val (incl_params, incl_tenv, incl_syn) = fold (params_of_expr context fixed_params)
      (map (prep_expr thy) includes) (import_params, import_tenv, import_syn);

    val ((import_ids, _), raw_import_elemss) =
      flatten (context, prep_expr thy) (([], Symtab.empty), Expr import);
    (* CB: normalise "includes" among elements *)
    val ((ids, syn), raw_elemsss) = foldl_map (flatten (context, prep_expr thy))
      ((import_ids, incl_syn), elements);

    val raw_elemss = flat raw_elemsss;
    (* CB: raw_import_elemss @ raw_elemss is the normalised list of
       context elements obtained from import and elements. *)
    (* Now additional elements for parameters are inserted. *)
    val import_params_ids = make_params_ids import_params;
    val incl_params_ids =
        make_params_ids (incl_params \\ import_params);
    val raw_import_params_elemss =
        make_raw_params_elemss (import_params, incl_tenv, incl_syn);
    val raw_incl_params_elemss =
        make_raw_params_elemss (incl_params \\ import_params, incl_tenv, incl_syn);
    val ((parms, all_elemss, concl), (spec, (_, _, defs))) = prep_elemss do_close
      context fixed_params
      (raw_import_params_elemss @ raw_import_elemss @ raw_incl_params_elemss @ raw_elemss) raw_concl;

    (* replace extended ids (for axioms) by ids *)
    val (import_ids', incl_ids) = chop (length import_ids) ids;
    val add_ids = import_params_ids @ import_ids' @ incl_params_ids @ incl_ids;
    val all_elemss' = map (fn (((_, ps), _), (((n, ps'), mode), elems)) =>
        (((n, map (fn p => (p, (the o AList.lookup (op =) ps') p)) ps), mode), elems))
      (add_ids ~~ all_elemss);
    (* CB: all_elemss and parms contain the correct parameter types *)

    val (ps, qs) = chop (length raw_import_params_elemss + length raw_import_elemss) all_elemss';
    val (import_ctxt, (import_elemss, _)) =
      activate_facts false prep_facts (context, ps);

    val (ctxt, (elemss, _)) =
      activate_facts false prep_facts (import_ctxt, qs);
  in
    ((((import_ctxt, import_elemss), (ctxt, elemss, syn)),
      (parms, spec, defs)), ([], concl))
(* FIXME: change signature, remove [] *)
  end;

fun prep_statement prep_locale prep_ctxt raw_locale elems concl ctxt =
  let
    val thy = ProofContext.theory_of ctxt;
    val locale = Option.map (prep_locale thy) raw_locale;
    val (fixed_params, import) =
      (case locale of
        NONE => ([], empty)
      | SOME name =>
          let val {params = ps, ...} = the_locale thy name
          in (map fst ps, Locale name) end);
    val ((((locale_ctxt, _), (elems_ctxt, _, _)), _), ([], concl')) =
      prep_ctxt false fixed_params import elems concl ctxt;
  in (locale, ([], locale_ctxt), ([], elems_ctxt), concl') end;
  (* FIXME: change signatures, remove empty statement lists *)

fun prep_expr prep import body ctxt =
  let
    val (((_, import_elemss), (ctxt', elemss, _)), _) = prep import body ctxt;
    val all_elems = maps snd (import_elemss @ elemss);
  in (all_elems, ctxt') end;

in

val read_ctxt = prep_context_statement intern_expr read_elemss read_facts;
val cert_ctxt = prep_context_statement (K I) cert_elemss cert_facts;

fun read_context import body ctxt = #1 (read_ctxt true [] import (map Elem body) [] ctxt);
fun cert_context import body ctxt = #1 (cert_ctxt true [] import (map Elem body) [] ctxt);

val read_expr = prep_expr read_context;
val cert_expr = prep_expr cert_context;

fun read_context_statement raw_locale = prep_statement intern read_ctxt raw_locale ;
fun cert_context_statement raw_locale = prep_statement (K I) cert_ctxt raw_locale ;

end;


(* print locale *)

fun print_locale thy show_facts import body =
  let
    val (all_elems, ctxt) = read_expr import body (ProofContext.init thy);
    val prt_elem = Element.pretty_ctxt ctxt;
  in
    Pretty.big_list "locale elements:" (all_elems
      |> (if show_facts then I else filter (fn Notes _ => false | _ => true))
      |> map (Pretty.chunks o prt_elem))
    |> Pretty.writeln
  end;



(** store results **)

(* naming of interpreted theorems *)

fun global_note_prefix_i kind prfx args thy =
  thy
  |> Theory.qualified_names
  |> Theory.sticky_prefix prfx
  |> PureThy.note_thmss_i kind args
  ||> Theory.restore_naming thy;

fun local_note_prefix_i prfx args ctxt =
  ctxt
  |> ProofContext.qualified_names
  |> ProofContext.sticky_prefix prfx
  |> ProofContext.note_thmss_i args
  ||> ProofContext.restore_naming ctxt;


(* collect witnesses for global registration;
   requires parameters and flattened list of (assumed!) identifiers
   instead of recomputing it from the target *)

fun collect_global_witnesses thy parms ids vts = let
    val ts = map Logic.unvarify vts;
    val (parms, parmTs) = split_list parms;
    val parmvTs = map Logic.varifyT parmTs;
    val vtinst = fold (Sign.typ_match thy) (parmvTs ~~ map Term.fastype_of ts) Vartab.empty;
    val tinst = Vartab.dest vtinst |> map (fn ((x, 0), (_, T)) => (x, T))
        |> Symtab.make;
    (* replace parameter names in ids by instantiations *)
    val vinst = Symtab.make (parms ~~ vts);
    fun vinst_names ps = map (the o Symtab.lookup vinst) ps;
    val inst = Symtab.make (parms ~~ ts);
    val ids' = map (apsnd vinst_names) ids;
    val wits = maps (snd o the o get_global_registration thy) ids';
  in ((tinst, inst), wits) end;


(* store instantiations of args for all registered interpretations
   of the theory *)

fun note_thmss_registrations kind target args thy =
  let
    val parms = the_locale thy target |> #params |> map fst;
    val ids = flatten (ProofContext.init thy, intern_expr thy)
      (([], Symtab.empty), Expr (Locale target)) |> fst |> fst
      |> map_filter (fn (id, (_, Assumed _)) => SOME id | _ => NONE)

    val regs = get_global_registrations thy target;

    (* add args to thy for all registrations *)

    fun activate (vts, ((prfx, atts2), _)) thy =
      let
        val (insts, prems) = collect_global_witnesses thy parms ids vts;
        val inst_atts =
          Args.map_values I (Element.instT_type (#1 insts))
            (Element.inst_term insts) (Element.inst_thm thy insts);
        val args' = args |> map (fn ((n, atts), [(ths, [])]) =>
            ((n, map (Attrib.attribute_i thy) (map inst_atts atts @ atts2)),
             [(map (Drule.standard o Element.satisfy_thm prems o
            Element.inst_thm thy insts) ths, [])]));
      in global_note_prefix_i kind prfx args' thy |> snd end;
  in fold activate regs thy end;


(* term syntax *)

fun add_term_syntax loc syn =
  snd #> syn #> ProofContext.map_theory (change_locale loc
    (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
      (axiom, import, elems, params, lparams, (syn, stamp ()) :: term_syntax, regs, intros)));

fun init_term_syntax loc ctxt =
  fold_rev (fn (f, _) => fn ctxt' => f ctxt')
    (#term_syntax (the_locale (ProofContext.theory_of ctxt) loc)) ctxt;

fun init loc =
  ProofContext.init
  #> init_term_syntax loc
  #> (#2 o cert_context_statement (SOME loc) [] []);


(* theory/locale results *)

fun theory_results kind prefix results (view, ctxt) thy =
  let
    val thy_ctxt = ProofContext.init thy;
    val export = ProofContext.export_view view ctxt thy_ctxt;
    val facts = map (fn (name, ths) => ((name, []), [(map export ths, [])])) results;
  in thy |> PureThy.note_thmss_qualified kind prefix facts end;

local

fun gen_thmss prep_facts global_results kind loc args (view, ctxt) thy =
  let
    val (ctxt', ([(_, [Notes args'])], facts)) =
      activate_facts true prep_facts (ctxt, [((("", []), Assumed []), [Ext (Notes args)])]);
    val (facts', thy') =
      thy
      |> change_locale loc (fn (axiom, import, elems, params, lparams, term_syntax, regs, intros) =>
        (axiom, import, elems @ [(Notes args', stamp ())], params, lparams, term_syntax, regs, intros))
      |> note_thmss_registrations kind loc args'
      |> global_results (map (#1 o #1) args' ~~ map #2 facts) (view, ctxt);
  in ((facts, facts'), (ProofContext.transfer thy' ctxt', thy')) end;

fun gen_note prep_locale prep_facts kind raw_loc args thy =
  let
    val loc = prep_locale thy raw_loc;
    val prefix = Sign.base_name loc;
  in gen_thmss prep_facts (theory_results kind prefix) kind loc args (init loc thy) thy end;

in

val note_thmss = gen_note intern read_facts;
val note_thmss_i = gen_note (K I) cert_facts;

fun add_thmss kind loc args (view, ctxt) =
  gen_thmss cert_facts (theory_results kind "")
    kind loc args (view, ctxt) (ProofContext.theory_of ctxt)
  ||> #1;

fun locale_results kind loc args (ctxt, thy) =
  thy |> gen_thmss cert_facts (K (K (pair [])))
    kind loc (map (apsnd Thm.simple_fact) args) ([], ctxt)
  |>> #1;

end;



(** define locales **)

(* predicate text *)
(* CB: generate locale predicates and delta predicates *)

local

(* introN: name of theorems for introduction rules of locale and
     delta predicates;
   axiomsN: name of theorem set with destruct rules for locale predicates,
     also name suffix of delta predicates. *)

val introN = "intro";
val axiomsN = "axioms";

fun atomize_spec thy ts =
  let
    val t = Logic.mk_conjunction_list ts;
    val body = ObjectLogic.atomize_term thy t;
    val bodyT = Term.fastype_of body;
  in
    if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
    else (body, bodyT, ObjectLogic.atomize_cterm (Thm.cterm_of thy t))
  end;

fun aprop_tr' n c = (c, fn args =>
  if length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.free c, args)
  else raise Match);

(* CB: define one predicate including its intro rule and axioms
   - bname: predicate name
   - parms: locale parameters
   - defs: thms representing substitutions from defines elements
   - ts: terms representing locale assumptions (not normalised wrt. defs)
   - norm_ts: terms representing locale assumptions (normalised wrt. defs)
   - thy: the theory
*)

fun def_pred bname parms defs ts norm_ts thy =
  let
    val name = Sign.full_name thy bname;

    val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
    val env = Term.add_term_free_names (body, []);
    val xs = filter (fn (x, _) => x mem_string env) parms;
    val Ts = map #2 xs;
    val extraTs = (Term.term_tfrees body \\ foldr Term.add_typ_tfrees [] Ts)
      |> Library.sort_wrt #1 |> map TFree;
    val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;

    val args = map Logic.mk_type extraTs @ map Free xs;
    val head = Term.list_comb (Const (name, predT), args);
    val statement = ObjectLogic.ensure_propT thy head;

    val ([pred_def], defs_thy) =
      thy
      |> (if bodyT <> propT then I else
        Theory.add_trfuns ([], [], map (aprop_tr' (length args)) (NameSpace.accesses' name), []))
      |> Theory.add_consts_i [(bname, predT, NoSyn)]
      |> PureThy.add_defs_i false [((Thm.def_name bname, Logic.mk_equals (head, body)), [])];

    val cert = Thm.cterm_of defs_thy;

    val intro = Drule.standard (Goal.prove defs_thy [] norm_ts statement (fn _ =>
      Tactic.rewrite_goals_tac [pred_def] THEN
      Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
      Tactic.compose_tac (false, Conjunction.intr_list (map (Thm.assume o cert) norm_ts), 0) 1));

    val conjuncts =
      (Drule.equal_elim_rule2 OF [body_eq,
        Tactic.rewrite_rule [pred_def] (Thm.assume (cert statement))])
      |> Conjunction.elim_precise [length ts] |> hd;
    val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
      Element.prove_witness defs_thy t
       (Tactic.rewrite_goals_tac defs THEN
        Tactic.compose_tac (false, ax, 0) 1));
  in ((statement, intro, axioms), defs_thy) end;

fun assumes_to_notes (Assumes asms) axms =
    fold_map (fn (a, spec) => fn axs =>
       let val (ps, qs) = chop (length spec) axs
       in ((a, [(ps, [])]), qs) end) asms axms
    |> apfst Notes
  | assumes_to_notes e axms = (e, axms);

(* CB: the following two change only "new" elems, these have identifier ("", _). *)

(* turn Assumes into Notes elements *)

fun change_assumes_elemss axioms elemss =
  let
    fun change (id as ("", _), es) =
          fold_map assumes_to_notes
            (map (Element.map_ctxt_values I I (Element.satisfy_thm axioms)) es)
          #-> (fn es' => pair (id, es'))
      | change e = pair e;
  in
    fst (fold_map change elemss (map Element.conclude_witness axioms))
  end;

(* adjust hyps of Notes elements *)

fun change_elemss_hyps axioms elemss =
  let
    fun change (id as ("", _), es) =
        (id, map (fn e as (Notes _) => Element.map_ctxt_values I I (Element.satisfy_thm axioms) e
                   | e => e) es)
      | change e = e;
  in map change elemss end;

in

(* CB: main predicate definition function *)

fun define_preds bname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =
  let
    val ((elemss', more_ts), a_elem, a_intro, thy') =
      if null exts then ((elemss, []), [], [], thy)
      else
        let
          val aname = if null ints then bname else bname ^ "_" ^ axiomsN;
          val ((statement, intro, axioms), def_thy) =
            thy |> def_pred aname parms defs exts exts';
          val elemss' = change_assumes_elemss axioms elemss;
          val def_thy' = def_thy
            |> PureThy.note_thmss_qualified "" aname [((introN, []), [([intro], [])])]
            |> snd
          val a_elem = [(("", []), [Assumes [((bname ^ "_" ^ axiomsN, []), [(statement, [])])]])];
        in ((elemss', [statement]), a_elem, [intro], def_thy') end;
    val (predicate, stmt', elemss'', b_intro, thy'') =
      if null ints then (([], []), more_ts, elemss' @ a_elem, [], thy')
      else
        let
          val ((statement, intro, axioms), def_thy) =
            thy' |> def_pred bname parms defs (ints @ more_ts) (ints' @ more_ts);
          val cstatement = Thm.cterm_of def_thy statement;
          val elemss'' = change_elemss_hyps axioms elemss';
          val def_thy' =
          def_thy
          |> PureThy.note_thmss_qualified "" bname
               [((introN, []), [([intro], [])]),
                ((axiomsN, []), [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
          |> snd;
          val b_elem = [(("", []),
               [Assumes [((bname ^ "_" ^ axiomsN, []), [(statement, [])])]])];
        in (([cstatement], axioms), [statement], elemss'' @ b_elem, [intro], def_thy') end;
  in (((elemss'', predicate, stmt'), (a_intro, b_intro)), thy'') end;

end;


(* add_locale(_i) *)

local

(* turn Defines into Notes elements, accumulate definition terms *)

fun defines_to_notes is_ext thy (Defines defs) defns =
    let
      val defs' = map (fn (_, (def, _)) => (("", []), (def, []))) defs
      val notes = map (fn (a, (def, _)) =>
        (a, [([assume (cterm_of thy def)], [])])) defs
    in (if is_ext then SOME (Notes notes) else NONE, defns @ [Defines defs']) end
  | defines_to_notes _ _ e defns = (SOME e, defns);

fun change_defines_elemss thy elemss defns =
  let
    fun change (id as (n, _), es) defns =
        let
          val (es', defns') = fold_map (defines_to_notes (n="") thy) es defns
        in ((id, map_filter I es'), defns') end
  in fold_map change elemss defns end;

fun gen_add_locale prep_ctxt prep_expr
    do_predicate bname raw_import raw_body thy =
  let
    val name = Sign.full_name thy bname;
    val _ = conditional (is_some (get_locale thy name)) (fn () =>
      error ("Duplicate definition of locale " ^ quote name));

    val thy_ctxt = ProofContext.init thy;
    val (((import_ctxt, import_elemss), (body_ctxt, body_elemss, syn)),
      text as (parms, ((_, exts'), _), defs)) =
      prep_ctxt raw_import raw_body thy_ctxt;
    val elemss = import_elemss @ body_elemss |>
        map_filter (fn ((id, Assumed axs), elems) => SOME (id, elems) | _ => NONE);
    val import = prep_expr thy raw_import;

    val extraTs = foldr Term.add_term_tfrees [] exts' \\
      foldr Term.add_typ_tfrees [] (map snd parms);
    val _ = if null extraTs then ()
      else warning ("Additional type variable(s) in locale specification " ^ quote bname);

    val ((((elemss', predicate as (predicate_statement, predicate_axioms), stmt'), intros),
          pred_thy), (import, regs)) =
      if do_predicate then
        let
          val (elemss', defns) = change_defines_elemss thy elemss [];
          val elemss'' = elemss' @
(*
              [(("", []), [Defines (map (fn t => (("", []), (t, []))) defs)])];
*)
              [(("", []), defns)];
          val (((elemss''', predicate as (_, axioms), stmt'), intros), thy') =
            define_preds bname text elemss'' thy;
          fun mk_regs elemss wits =
            fold_map (fn (id, elems) => fn wts => let
                val ts = List.concat (List.mapPartial (fn (Assumes asms) =>
                  SOME (List.concat (map (map #1 o #2) asms)) | _ => NONE) elems);
                val (wts1, wts2) = chop (length ts) wts
              in ((apsnd (map fst) id, wts1), wts2) end) elemss wits |> fst;
          val regs = mk_regs elemss''' axioms |>
                map_filter (fn (("", _), _) => NONE | e => SOME e);
        in ((((elemss''', predicate, stmt'), intros), thy'), (empty, regs)) end
      else ((((elemss, ([], []), []), ([], [])), thy), (import, []));

    fun axiomify axioms elemss =
      (axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let
                   val ts = flat (map_filter (fn (Assumes asms) =>
                     SOME (maps (map #1 o #2) asms) | _ => NONE) elems);
                   val (axs1, axs2) = chop (length ts) axs;
                 in (axs2, ((id, Assumed axs1), elems)) end)
        |> snd;
    val pred_ctxt = ProofContext.init pred_thy;
    val (ctxt, (_, facts)) = activate_facts true (K I)
      (pred_ctxt, axiomify predicate_axioms elemss');
    val export = ProofContext.export_view predicate_statement ctxt thy_ctxt;
    val facts' = facts |> map (fn (a, ths) => ((a, []), [(map export ths, [])]));
    val elems' = maps #2 (filter (equal "" o #1 o #1) elemss');
    val elems'' = map_filter (fn (Fixes _) => NONE | e => SOME e) elems';
    val axs' = map (Element.assume_witness pred_thy) stmt';
    val thy' = pred_thy
      |> PureThy.note_thmss_qualified "" bname facts' |> snd
      |> declare_locale name
      |> put_locale name
       {axiom = axs',
        import = import,
        elems = map (fn e => (e, stamp ())) elems'',
        params = params_of elemss' |> map (fn (x, SOME T) => ((x, T), the (Symtab.lookup syn x))),
        lparams = map #1 (params_of' body_elemss),
        term_syntax = [],
        regs = regs,
        intros = intros};
  in ((name, ProofContext.transfer thy' body_ctxt), thy') end;

in

val add_locale = gen_add_locale read_context intern_expr;
val add_locale_i = gen_add_locale cert_context (K I);

end;

val _ = Context.add_setup
 (add_locale_i true "var" empty [Fixes [(Term.internal "x", NONE, NoSyn)]] #> snd #>
  add_locale_i true "struct" empty [Fixes [(Term.internal "S", NONE, Structure)]] #> snd);



(** locale goals **)

local

fun intern_attrib thy = map_elem (Element.map_ctxt
  {name = I, var = I, typ = I, term = I, fact = I, attrib = Attrib.intern_src thy});

val global_goal = Proof.global_goal ProofDisplay.present_results
  Attrib.attribute_i ProofContext.bind_propp_schematic_i;

fun conclusion prep_att (Element.Shows concl) =
      (([], concl), fn stmt => fn ctxt => ((Attrib.map_specs prep_att stmt, []), ctxt))
  | conclusion _ (Element.Obtains cases) = apfst (apfst (map Elem)) (Obtain.statement cases);

fun gen_theorem prep_src prep_elem prep_stmt
    kind before_qed after_qed (name, raw_atts) raw_elems raw_concl thy =
  let
    val atts = map (prep_src thy) raw_atts;
    val ((concl_elems, concl), mk_stmt) = conclusion (prep_src thy) raw_concl;
    val thy_ctxt = ProofContext.init thy;
    val elems = map (prep_elem thy) (raw_elems @ concl_elems);
    val (_, _, (view, ctxt), propp) = prep_stmt NONE elems (map snd concl) thy_ctxt;
    val ((stmt, facts), goal_ctxt) = ctxt
      |> ProofContext.add_view thy_ctxt view
      |> mk_stmt (map fst concl ~~ propp);
  in
    global_goal kind before_qed after_qed (SOME "") (name, atts) stmt goal_ctxt
    |> Proof.refine_insert facts
  end;

fun gen_theorem_in_locale prep_locale prep_src prep_elem prep_stmt no_target
    kind before_qed after_qed raw_loc (name, atts) raw_elems raw_concl thy =
  let
    val ((concl_elems, concl), mk_stmt) = conclusion (prep_src thy) raw_concl;
    val loc = prep_locale thy raw_loc;
    val loc_atts = map (prep_src thy) atts;
    val loc_attss = map (map (prep_src thy) o snd o fst) concl;
    val target = if no_target then NONE else SOME (extern thy loc);
    val elems = map (prep_elem thy) (raw_elems @ concl_elems);
    val names = map (fst o fst) concl;

    val thy_ctxt = init_term_syntax loc (ProofContext.init thy);
    val (_, (loc_view, loc_ctxt), (elems_view, elems_ctxt), propp) =
      prep_stmt (SOME raw_loc) elems (map snd concl) thy_ctxt;
    val elems_ctxt' = elems_ctxt
      |> ProofContext.add_view loc_ctxt elems_view
      |> ProofContext.add_view thy_ctxt loc_view;
    val loc_ctxt' = loc_ctxt
      |> ProofContext.add_view thy_ctxt loc_view;

    val ((stmt, facts), goal_ctxt) =
      elems_ctxt' |> mk_stmt (map (apsnd (K []) o fst) concl ~~ propp);

    fun after_qed' results =
      let val loc_results = results |> map
          (ProofContext.export_standard goal_ctxt loc_ctxt') in
        curry (locale_results kind loc ((names ~~ loc_attss) ~~ loc_results)) loc_ctxt
        #-> (fn res =>
          if name = "" andalso null loc_atts then I
          else #2 o locale_results kind loc [((name, loc_atts), maps #2 res)])
        #> #2
        #> after_qed loc_results results
      end;
  in
    global_goal kind before_qed after_qed' target (name, []) stmt goal_ctxt
    |> Proof.refine_insert facts
  end;

in

val theorem = gen_theorem Attrib.intern_src intern_attrib read_context_statement;
val theorem_i = gen_theorem (K I) (K I) cert_context_statement;

val theorem_in_locale = gen_theorem_in_locale intern Attrib.intern_src intern_attrib
  read_context_statement false;

val theorem_in_locale_i = gen_theorem_in_locale (K I) (K I) (K I)
  cert_context_statement false;

val theorem_in_locale_no_target = gen_theorem_in_locale (K I) (K I) (K I)
  cert_context_statement true;

fun smart_theorem kind NONE a [] (Element.Shows concl) =
      Proof.theorem kind NONE (K I) (SOME "") a concl o ProofContext.init
  | smart_theorem kind NONE a elems concl =
      theorem kind NONE (K I) a elems concl
  | smart_theorem kind (SOME loc) a elems concl =
      theorem_in_locale kind NONE (K (K I)) loc a elems concl;

end;


(** Normalisation of locale statements ---
    discharges goals implied by interpretations **)

local

fun locale_assm_intros thy =
  Symtab.fold (fn (_, {intros = (a, _), ...}) => fn intros => (a @ intros))
    (#2 (GlobalLocalesData.get thy)) [];
fun locale_base_intros thy =
  Symtab.fold (fn (_, {intros = (_, b), ...}) => fn intros => (b @ intros))
    (#2 (GlobalLocalesData.get thy)) [];

fun all_witnesses ctxt =
  let
    val thy = ProofContext.theory_of ctxt;
    fun get registrations = Symtab.fold (fn (_, regs) => fn thms =>
        (Registrations.dest regs |> map (fn (_, (_, wits)) =>
          map Element.conclude_witness wits) |> flat) @ thms)
      registrations [];
    val globals = get (#3 (GlobalLocalesData.get thy));
    val locals = get (LocalLocalesData.get ctxt);
  in globals @ locals end;
(* FIXME: proper varification *)

in

fun intro_locales_tac eager ctxt facts st =
  let
    val wits = all_witnesses ctxt |> map Thm.varifyT;
    val thy = ProofContext.theory_of ctxt;
    val intros = locale_base_intros thy @ (if eager then locale_assm_intros thy else []);
  in
    (ALLGOALS (Method.insert_tac facts THEN'
        REPEAT_ALL_NEW (resolve_tac (wits @ intros)))
      THEN Tactic.distinct_subgoals_tac) st
  end;

val _ = Context.add_setup (Method.add_methods
  [("intro_locales",
    Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt)),
    "back-chain introduction rules of locales without unfolding predicates"),
   ("unfold_locales",
    Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt)),
    "back-chain all introduction rules of locales")]);

end;


(** Interpretation commands **)

local

(* extract proof obligations (assms and defs) from elements *)

fun extract_asms_elems ((id, Assumed _), elems) = (id, maps Element.prems_of elems)
  | extract_asms_elems ((id, Derived _), _) = (id, []);


(* activate instantiated facts in theory or context *)

fun gen_activate_facts_elemss get_reg note attrib std put_reg add_wit
        attn all_elemss new_elemss propss thmss thy_ctxt =
    let
      fun activate_elem disch (prfx, atts) (Notes facts) thy_ctxt =
          let
            val facts' = facts
              (* discharge hyps in attributes *)
              |> Attrib.map_facts (attrib thy_ctxt o Args.map_values I I I disch)
              (* insert interpretation attributes *)
              |> map (apfst (apsnd (fn a => a @ map (attrib thy_ctxt) atts)))
              (* discharge hyps *)
              |> map (apsnd (map (apfst (map disch))));
          in snd (note prfx facts' thy_ctxt) end
        | activate_elem _ _ _ thy_ctxt = thy_ctxt;

      fun activate_elems disch ((id, _), elems) thy_ctxt =
          let
            val ((prfx, atts2), _) = the (get_reg thy_ctxt id)
                handle Option => sys_error ("Unknown registration of " ^
                  quote (fst id) ^ " while activating facts.");
          in
            fold (activate_elem disch (prfx, atts2)) elems thy_ctxt
          end;

      val thy_ctxt' = thy_ctxt
        (* add registrations *)
        |> fold (fn ((id, _), _) => put_reg id attn) new_elemss
        (* add witnesses of Assumed elements *)
        |> fold (fn (id, thms) => fold (add_wit id) thms) (map fst propss ~~ thmss);

      val prems = flat (map_filter
            (fn ((id, Assumed _), _) => Option.map snd (get_reg thy_ctxt' id)
              | ((_, Derived _), _) => NONE) all_elemss);
      val thy_ctxt'' = thy_ctxt'
        (* add witnesses of Derived elements *)
        |> fold (fn (id, thms) => fold (add_wit id o Element.satisfy_witness prems) thms)
           (map_filter (fn ((_, Assumed _), _) => NONE
              | ((id, Derived thms), _) => SOME (id, thms)) new_elemss)

      val disch' = std o Element.satisfy_thm prems;  (* FIXME *)
    in
      thy_ctxt''
        (* add facts to theory or context *)
        |> fold (activate_elems disch') new_elemss
    end;

fun global_activate_facts_elemss x = gen_activate_facts_elemss
      (fn thy => fn (name, ps) =>
        get_global_registration thy (name, map Logic.varify ps))
      (global_note_prefix_i "")
      Attrib.attribute_i Drule.standard
      (fn (name, ps) => put_global_registration (name, map Logic.varify ps))
      (fn (n, ps) => add_global_witness (n, map Logic.varify ps) o
        Element.map_witness (fn (t, th) => (Logic.legacy_unvarify t, legacy_unvarify th))
        (* FIXME *)) x;

fun local_activate_facts_elemss x = gen_activate_facts_elemss
      get_local_registration
      local_note_prefix_i
      (Attrib.attribute_i o ProofContext.theory_of) I
      put_local_registration
      add_local_witness x;

fun gen_prep_registration mk_ctxt is_local read_terms test_reg activate
    attn expr insts thy_ctxt =
  let
    val ctxt = mk_ctxt thy_ctxt;
    val thy = ProofContext.theory_of ctxt;

    val ctxt' = ctxt |> ProofContext.theory_of |> ProofContext.init;
    val pts = params_of_expr ctxt' [] (intern_expr thy expr)
          ([], Symtab.empty, Symtab.empty);
    val params_ids = make_params_ids (#1 pts);
    val raw_params_elemss = make_raw_params_elemss pts;
    val ((ids, _), raw_elemss) = flatten (ctxt', intern_expr thy)
          (([], Symtab.empty), Expr expr);
    val ((parms, all_elemss, _), (_, (_, defs, _))) =
          read_elemss false ctxt' [] (raw_params_elemss @ raw_elemss) [];

    (** compute instantiation **)

    (* user input *)
    val insts = if length parms < length insts
         then error "More arguments than parameters in instantiation."
         else insts @ replicate (length parms - length insts) NONE;
    val (ps, pTs) = split_list parms;
    val pvTs = map Logic.legacy_varifyT pTs;

    (* instantiations given by user *)
    val given = map_filter (fn (_, (NONE, _)) => NONE
         | (x, (SOME inst, T)) => SOME (x, (inst, T))) (ps ~~ (insts ~~ pvTs));
    val (given_ps, given_insts) = split_list given;
    val tvars = foldr Term.add_typ_tvars [] pvTs;
    val used = foldr Term.add_typ_varnames [] pvTs;
    fun sorts (a, i) = AList.lookup (op =) tvars (a, i);
    val (vs, vinst) = read_terms thy_ctxt sorts used given_insts;
    val vars = foldl Term.add_term_tvar_ixns [] vs \\ map fst tvars;
    val vars' = fold Term.add_term_varnames vs vars;
    val _ = if null vars' then ()
         else error ("Illegal schematic variable(s) in instantiation: " ^
           commas_quote (map Syntax.string_of_vname vars'));
    (* replace new types (which are TFrees) by ones with new names *)
    val new_Tnames = foldr Term.add_term_tfree_names [] vs;
    val new_Tnames' = Term.invent_names used "'a" (length new_Tnames);
    val renameT =
          if is_local then I
          else Logic.legacy_unvarifyT o Term.map_type_tfree (fn (a, s) =>
            TFree ((the o AList.lookup (op =) (new_Tnames ~~ new_Tnames')) a, s));
    val rename =
          if is_local then I
          else Term.map_term_types renameT;

    val tinst = Symtab.make (map
                (fn ((x, 0), T) => (x, T |> renameT)
                  | ((_, n), _) => sys_error "Internal error var in prep_registration") vinst);
    val inst = Symtab.make (given_ps ~~ map rename vs);

    (* defined params without user input *)
    val not_given = map_filter (fn (x, (NONE, T)) => SOME (x, T)
         | (_, (SOME _, _)) => NONE) (ps ~~ (insts ~~ pTs));
    fun add_def (p, pT) inst =
      let
        val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of
               NONE => error ("Instance missing for parameter " ^ quote p)
             | SOME (Free (_, T), t) => (t, T);
        val d = Element.inst_term (tinst, inst) t;
      in Symtab.update_new (p, d) inst end;
    val inst = fold add_def not_given inst;
    val insts = (tinst, inst);
    (* Note: insts contain no vars. *)


    (** compute proof obligations **)

    (* restore "small" ids *)
    val ids' = map (fn ((n, ps), (_, mode)) =>
          ((n, map (fn p => Free (p, (the o AList.lookup (op =) parms) p)) ps), mode))
        ids;
    val (_, all_elemss') = chop (length raw_params_elemss) all_elemss
    (* instantiate ids and elements *)
    val inst_elemss = (ids' ~~ all_elemss') |> map (fn (((n, ps), _), ((_, mode), elems)) =>
      ((n, map (Element.inst_term insts) ps),
        map (fn Int e => Element.inst_ctxt thy insts e) elems)
      |> apfst (fn id => (id, map_mode (map (Element.inst_witness thy insts)) mode)));

    (* remove fragments already registered with theory or context *)
    val new_inst_elemss = filter (fn ((id, _), _) =>
          not (test_reg thy_ctxt id)) inst_elemss;
    val new_ids = map #1 new_inst_elemss;

    val propss = map extract_asms_elems new_inst_elemss;

    val bind_attrib = Attrib.crude_closure ctxt o Args.assignable;
    val attn' = apsnd (map (bind_attrib o Attrib.intern_src thy)) attn;

  in (propss, activate attn' inst_elemss new_inst_elemss propss) end;

val prep_global_registration = gen_prep_registration
     ProofContext.init false
     (fn thy => fn sorts => fn used =>
       Sign.read_def_terms (thy, K NONE, sorts) used true)
     (fn thy => fn (name, ps) =>
       test_global_registration thy (name, map Logic.varify ps))
     global_activate_facts_elemss;

val prep_local_registration = gen_prep_registration
     I true
     (fn ctxt => ProofContext.read_termTs ctxt (K false) (K NONE))
     smart_test_registration
     local_activate_facts_elemss;

fun prep_registration_in_locale target expr thy =
  (* target already in internal form *)
  let
    val ctxt = ProofContext.init thy;
    val ((raw_target_ids, target_syn), _) = flatten (ctxt, I)
        (([], Symtab.empty), Expr (Locale target));
    val fixed = the_locale thy target |> #params |> map #1;
    val ((all_ids, syn), raw_elemss) = flatten (ctxt, intern_expr thy)
        ((raw_target_ids, target_syn), Expr expr);
    val (target_ids, ids) = chop (length raw_target_ids) all_ids;
    val ((parms, elemss, _), _) = read_elemss false ctxt fixed raw_elemss [];

    (** compute proof obligations **)

    (* restore "small" ids, with mode *)
    val ids' = map (apsnd snd) ids;
    (* remove Int markers *)
    val elemss' = map (fn (_, es) =>
        map (fn Int e => e) es) elemss
    (* extract assumptions and defs *)
    val ids_elemss = ids' ~~ elemss';
    val propss = map extract_asms_elems ids_elemss;

    (** activation function:
        - add registrations to the target locale
        - add induced registrations for all global registrations of
          the target, unless already present
        - add facts of induced registrations to theory **)

    val t_ids = map_filter
        (fn (id, (_, Assumed _)) => SOME id | _ => NONE) target_ids;

    fun activate thmss thy = let
        val satisfy = Element.satisfy_thm (flat thmss);
        val ids_elemss_thmss = ids_elemss ~~ thmss;
        val regs = get_global_registrations thy target;

        fun activate_id (((id, Assumed _), _), thms) thy =
            thy |> put_registration_in_locale target id
                |> fold (add_witness_in_locale target id) thms
          | activate_id _ thy = thy;

        fun activate_reg (vts, ((prfx, atts2), _)) thy = let
            val (insts, wits) = collect_global_witnesses thy fixed t_ids vts;
            fun inst_parms ps = map
                  (the o AList.lookup (op =) (map #1 fixed ~~ vts)) ps;
            val disch = Element.satisfy_thm wits;
            val new_elemss = filter (fn (((name, ps), _), _) =>
                not (test_global_registration thy (name, inst_parms ps))) (ids_elemss);
            fun activate_assumed_id (((_, Derived _), _), _) thy = thy
              | activate_assumed_id ((((name, ps), Assumed _), _), thms) thy = let
                val ps' = inst_parms ps;
              in
                if test_global_registration thy (name, ps')
                then thy
                else thy
                  |> put_global_registration (name, ps') (prfx, atts2)
                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
                     (Element.inst_witness thy insts witn) thy) thms
              end;

            fun activate_derived_id ((_, Assumed _), _) thy = thy
              | activate_derived_id (((name, ps), Derived ths), _) thy = let
                val ps' = inst_parms ps;
              in
                if test_global_registration thy (name, ps')
                then thy
                else thy
                  |> put_global_registration (name, ps') (prfx, atts2)
                  |> fold (fn witn => fn thy => add_global_witness (name, ps')
                       (witn |> Element.map_witness (fn (t, th) =>  (* FIXME *)
                       (Element.inst_term insts t,
                        disch (Element.inst_thm thy insts (satisfy th))))) thy) ths
              end;

            fun activate_elem (Notes facts) thy =
                let
                  val facts' = facts
                      |> Attrib.map_facts (Attrib.attribute_i thy o
                         Args.map_values I (Element.instT_type (#1 insts))
                           (Element.inst_term insts)
                           (disch o Element.inst_thm thy insts o satisfy))
                      |> map (apfst (apsnd (fn a => a @ map (Attrib.attribute thy) atts2)))
                      |> map (apsnd (map (apfst (map (disch o Element.inst_thm thy insts o satisfy)))))
                in
                  thy
                  |> global_note_prefix_i "" prfx facts'
                  |> snd
                end
              | activate_elem _ thy = thy;

            fun activate_elems (_, elems) thy = fold activate_elem elems thy;

          in thy |> fold activate_assumed_id ids_elemss_thmss
                 |> fold activate_derived_id ids_elemss
                 |> fold activate_elems new_elemss end;
      in
        thy |> fold activate_id ids_elemss_thmss
            |> fold activate_reg regs
      end;

  in (propss, activate) end;

fun prep_propp propss = propss |> map (fn ((name, _), props) =>
  (("", []), map (rpair [] o Element.mark_witness) props));

fun prep_result propps thmss =
  ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);

fun goal_name thy kind target propss =
    kind ^ Proof.goal_names (Option.map (extern thy) target) ""
      (propss |> map (fn ((name, _), _) => extern thy name));

in

fun interpretation (prfx, atts) expr insts thy =
  let
    val (propss, activate) = prep_global_registration (prfx, atts) expr insts thy;
    val kind = goal_name thy "interpretation" NONE propss;
    val after_qed = activate o prep_result propss;
  in
    ProofContext.init thy
    |> Proof.theorem_i kind NONE after_qed NONE ("", []) (prep_propp propss)
    |> Element.refine_witness |> Seq.hd
  end;

fun interpretation_in_locale (raw_target, expr) thy =
  let
    val target = intern thy raw_target;
    val (propss, activate) = prep_registration_in_locale target expr thy;
    val kind = goal_name thy "interpretation" (SOME target) propss;
    val after_qed = K (activate o prep_result propss);
  in
    thy
    |> theorem_in_locale_no_target kind NONE after_qed target ("", []) []
      (Element.Shows (prep_propp propss))
    |> Element.refine_witness |> Seq.hd
  end;

fun interpret (prfx, atts) expr insts int state =
  let
    val _ = Proof.assert_forward_or_chain state;
    val ctxt = Proof.context_of state;
    val (propss, activate) = prep_local_registration (prfx, atts) expr insts ctxt;
    val kind = goal_name (Proof.theory_of state) "interpret" NONE propss;
    fun after_qed results =
      Proof.map_context (K (ctxt |> activate (prep_result propss results)))
      #> Proof.put_facts NONE
      #> Seq.single;
  in
    state
    |> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i
      kind NONE after_qed (prep_propp propss)
    |> Element.refine_witness |> Seq.hd
  end;

end;

end;