src/Pure/Isar/expression.ML
author haftmann
Fri, 05 Dec 2008 18:43:42 +0100
changeset 28999 abe0f11cfa4e
parent 28994 49f602ae24e5
child 29020 3e95d28114a1
permissions -rw-r--r--
Name.name_of -> Binding.base_name
     1 (*  Title:      Pure/Isar/expression.ML
     2     Author:     Clemens Ballarin, TU Muenchen
     3 
     4 New locale development --- experimental.
     5 *)
     6 
     7 signature EXPRESSION =
     8 sig
     9   datatype 'term map = Positional of 'term option list | Named of (string * 'term) list;
    10   type 'term expr = (string * (string * 'term map)) list;
    11   type expression = string expr * (Binding.T * string option * mixfix) list;
    12   type expression_i = term expr * (Binding.T * typ option * mixfix) list;
    13 
    14   (* Processing of context statements *)
    15   val read_statement: Element.context list -> (string * string list) list list ->
    16     Proof.context ->  (term * term list) list list * Proof.context;
    17   val cert_statement: Element.context_i list -> (term * term list) list list ->
    18     Proof.context -> (term * term list) list list * Proof.context;
    19 
    20   (* Declaring locales *)
    21   val add_locale_cmd: string -> bstring -> expression -> Element.context list -> theory ->
    22     string * Proof.context
    23   val add_locale: string -> bstring -> expression_i -> Element.context_i list -> theory ->
    24     string * Proof.context
    25 
    26   (* Interpretation *)
    27   val sublocale_cmd: string -> expression -> theory -> Proof.state;
    28   val sublocale: string -> expression_i -> theory -> Proof.state;
    29   val interpretation_cmd: expression -> theory -> Proof.state;
    30   val interpretation: expression_i -> theory -> Proof.state;
    31 (*
    32   val interpret_cmd: Morphism.morphism -> expression -> bool -> Proof.state -> Proof.state;
    33   val interpret: Morphism.morphism -> expression_i -> bool -> Proof.state -> Proof.state;
    34 *)
    35 
    36   (* Debugging and development *)
    37   val parse_expression: OuterParse.token list -> expression * OuterParse.token list
    38     (* FIXME to spec_parse.ML *)
    39 end;
    40 
    41 
    42 structure Expression : EXPRESSION =
    43 struct
    44 
    45 datatype ctxt = datatype Element.ctxt;
    46 
    47 
    48 (*** Expressions ***)
    49 
    50 datatype 'term map =
    51   Positional of 'term option list |
    52   Named of (string * 'term) list;
    53 
    54 type 'term expr = (string * (string * 'term map)) list;
    55 
    56 type expression = string expr * (Binding.T * string option * mixfix) list;
    57 type expression_i = term expr * (Binding.T * typ option * mixfix) list;
    58 
    59 
    60 (** Parsing and printing **)
    61 
    62 local
    63 
    64 structure P = OuterParse;
    65 
    66 val loc_keyword = P.$$$ "fixes" || P.$$$ "constrains" || P.$$$ "assumes" ||
    67    P.$$$ "defines" || P.$$$ "notes";
    68 fun plus1_unless test scan =
    69   scan ::: Scan.repeat (P.$$$ "+" |-- Scan.unless test (P.!!! scan));
    70 
    71 val prefix = P.name --| P.$$$ ":";
    72 val named = P.name -- (P.$$$ "=" |-- P.term);
    73 val position = P.maybe P.term;
    74 val instance = P.$$$ "where" |-- P.and_list1 named >> Named ||
    75   Scan.repeat1 position >> Positional;
    76 
    77 in
    78 
    79 val parse_expression =
    80   let
    81     fun expr2 x = P.xname x;
    82     fun expr1 x = (Scan.optional prefix "" -- expr2 --
    83       Scan.optional instance (Named []) >> (fn ((p, l), i) => (l, (p, i)))) x;
    84     fun expr0 x = (plus1_unless loc_keyword expr1) x;
    85   in expr0 -- P.for_fixes end;
    86 
    87 end;
    88 
    89 fun pretty_expr thy expr =
    90   let
    91     fun pretty_pos NONE = Pretty.str "_"
    92       | pretty_pos (SOME x) = Pretty.str x;
    93     fun pretty_named (x, y) = [Pretty.str x, Pretty.brk 1, Pretty.str "=",
    94           Pretty.brk 1, Pretty.str y] |> Pretty.block;
    95     fun pretty_ren (Positional ps) = take_suffix is_none ps |> snd |>
    96           map pretty_pos |> Pretty.breaks
    97       | pretty_ren (Named []) = []
    98       | pretty_ren (Named ps) = Pretty.str "where" :: Pretty.brk 1 ::
    99           (ps |> map pretty_named |> Pretty.separate "and");
   100     fun pretty_rename (loc, ("", ren)) =
   101           Pretty.block (Pretty.str (NewLocale.extern thy loc) :: Pretty.brk 1 :: pretty_ren ren) 
   102       | pretty_rename (loc, (prfx, ren)) =
   103           Pretty.block (Pretty.str prfx :: Pretty.brk 1 :: Pretty.str (NewLocale.extern thy loc) ::
   104             Pretty.brk 1 :: pretty_ren ren);
   105   in Pretty.separate "+" (map pretty_rename expr) |> Pretty.block end;
   106 
   107 fun err_in_expr thy msg expr =
   108   let
   109     val err_msg =
   110       if null expr then msg
   111       else msg ^ "\n" ^ Pretty.string_of (Pretty.block
   112         [Pretty.str "The above error(s) occurred in expression:", Pretty.brk 1,
   113           pretty_expr thy expr])
   114   in error err_msg end;
   115 
   116 
   117 (** Internalise locale names in expr **)
   118 
   119 fun intern thy instances =  map (apfst (NewLocale.intern thy)) instances;
   120 
   121 
   122 (** Parameters of expression.
   123 
   124    Sanity check of instantiations and extraction of implicit parameters.
   125    The latter only occurs iff strict = false.
   126    Positional instantiations are extended to match full length of parameter list
   127    of instantiated locale. **)
   128 
   129 fun parameters_of thy strict (expr, fixed) =
   130   let
   131     fun reject_dups message xs =
   132       let val dups = duplicates (op =) xs
   133       in
   134         if null dups then () else error (message ^ commas dups)
   135       end;
   136 
   137     fun match_bind (n, b) = (n = Binding.base_name b);
   138     fun bind_eq (b1, b2) = (Binding.base_name b1 = Binding.base_name b2);
   139       (* FIXME: cannot compare bindings for equality. *)
   140 
   141     fun params_loc loc =
   142           (NewLocale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
   143     fun params_inst (expr as (loc, (prfx, Positional insts))) =
   144           let
   145             val (ps, loc') = params_loc loc;
   146 	    val d = length ps - length insts;
   147 	    val insts' =
   148 	      if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
   149                 quote (NewLocale.extern thy loc))
   150 	      else insts @ replicate d NONE;
   151             val ps' = (ps ~~ insts') |>
   152               map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
   153           in (ps', (loc', (prfx, Positional insts'))) end
   154       | params_inst (expr as (loc, (prfx, Named insts))) =
   155           let
   156             val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
   157               (map fst insts);
   158 
   159             val (ps, loc') = params_loc loc;
   160             val ps' = fold (fn (p, _) => fn ps =>
   161               if AList.defined match_bind ps p then AList.delete match_bind p ps
   162               else error (quote p ^" not a parameter of instantiated expression.")) insts ps;
   163           in (ps', (loc', (prfx, Named insts))) end;
   164     fun params_expr is =
   165           let
   166             val (is', ps') = fold_map (fn i => fn ps =>
   167               let
   168                 val (ps', i') = params_inst i;
   169                 val ps'' = AList.join bind_eq (fn p => fn (mx1, mx2) =>
   170                   (* FIXME: should check for bindings being the same.
   171                      Instead we check for equal name and syntax. *)
   172                   if mx1 = mx2 then mx1
   173                   else error ("Conflicting syntax for parameter" ^ quote (Binding.display p) ^
   174                     " in expression.")) (ps, ps')
   175               in (i', ps'') end) is []
   176           in (ps', is') end;
   177 
   178     val (implicit, expr') = params_expr expr;
   179 
   180     val implicit' = map (#1 #> Binding.base_name) implicit;
   181     val fixed' = map (#1 #> Binding.base_name) fixed;
   182     val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
   183     val implicit'' = if strict then []
   184       else let val _ = reject_dups
   185           "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed')
   186         in map (fn (b, mx) => (b, NONE, mx)) implicit end;
   187 
   188   in (expr', implicit'' @ fixed) end;
   189 
   190 
   191 (** Read instantiation **)
   192 
   193 (* Parse positional or named instantiation *)
   194 
   195 local
   196 
   197 fun prep_inst parse_term parms (Positional insts) ctxt =
   198       (insts ~~ parms) |> map (fn
   199         (NONE, p) => Syntax.parse_term ctxt p |
   200         (SOME t, _) => parse_term ctxt t)
   201   | prep_inst parse_term parms (Named insts) ctxt =
   202       parms |> map (fn p => case AList.lookup (op =) insts p of
   203         SOME t => parse_term ctxt t |
   204         NONE => Syntax.parse_term ctxt p);
   205 
   206 in
   207 
   208 fun parse_inst x = prep_inst Syntax.parse_term x;
   209 fun make_inst x = prep_inst (K I) x;
   210 
   211 end;
   212 
   213 
   214 (* Instantiation morphism *)
   215 
   216 fun inst_morph (parm_names, parm_types) (prfx, insts') ctxt =
   217   let
   218     (* parameters *)
   219     val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
   220 
   221     (* type inference and contexts *)
   222     val parm_types' = map (TypeInfer.paramify_vars o Logic.varifyT) parm_types;
   223     val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar);
   224     val arg = type_parms @ map2 TypeInfer.constrain parm_types' insts';
   225     val res = Syntax.check_terms ctxt arg;
   226     val ctxt' = ctxt |> fold Variable.auto_fixes res;
   227     
   228     (* instantiation *)
   229     val (type_parms'', res') = chop (length type_parms) res;
   230     val insts'' = (parm_names ~~ res') |> map_filter
   231       (fn (inst as (x, Free (y, _))) => if x = y then NONE else SOME inst |
   232         inst => SOME inst);
   233     val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
   234     val inst = Symtab.make insts'';
   235   in
   236     (Element.inst_morphism (ProofContext.theory_of ctxt) (instT, inst) $>
   237       Morphism.binding_morphism (Binding.qualify prfx), ctxt')
   238   end;
   239 
   240 
   241 (*** Locale processing ***)
   242 
   243 (** Parsing **)
   244 
   245 fun parse_elem prep_typ prep_term ctxt elem =
   246   Element.map_ctxt {binding = I, var = I, typ = prep_typ ctxt,
   247     term = prep_term ctxt, fact = I, attrib = I} elem;
   248 
   249 fun parse_concl prep_term ctxt concl =
   250   (map o map) (fn (t, ps) =>
   251     (prep_term ctxt, map (prep_term ctxt) ps)) concl;
   252 
   253 
   254 (** Simultaneous type inference: instantiations + elements + conclusion **)
   255 
   256 local
   257 
   258 fun mk_type T = (Logic.mk_type T, []);
   259 fun mk_term t = (t, []);
   260 fun mk_propp (p, pats) = (Syntax.type_constraint propT p, pats);
   261 
   262 fun dest_type (T, []) = Logic.dest_type T;
   263 fun dest_term (t, []) = t;
   264 fun dest_propp (p, pats) = (p, pats);
   265 
   266 fun extract_inst (_, (_, ts)) = map mk_term ts;
   267 fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs));
   268 
   269 fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes
   270   | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts
   271   | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms
   272   | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs
   273   | extract_elem (Notes _) = [];
   274 
   275 fun restore_elem (Fixes fixes, css) =
   276       (fixes ~~ css) |> map (fn ((x, _, mx), cs) =>
   277         (x, cs |> map dest_type |> try hd, mx)) |> Fixes
   278   | restore_elem (Constrains csts, css) =
   279       (csts ~~ css) |> map (fn ((x, _), cs) =>
   280         (x, cs |> map dest_type |> hd)) |> Constrains
   281   | restore_elem (Assumes asms, css) =
   282       (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes
   283   | restore_elem (Defines defs, css) =
   284       (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines
   285   | restore_elem (Notes notes, _) = Notes notes;
   286 
   287 fun check cs context =
   288   let
   289     fun prep (_, pats) (ctxt, t :: ts) =
   290       let val ctxt' = Variable.auto_fixes t ctxt
   291       in
   292         ((t, Syntax.check_props (ProofContext.set_mode ProofContext.mode_pattern ctxt') pats),
   293           (ctxt', ts))
   294       end
   295     val (cs', (context', _)) = fold_map prep cs
   296       (context, Syntax.check_terms
   297         (ProofContext.set_mode ProofContext.mode_schematic context) (map fst cs));
   298   in (cs', context') end;
   299 
   300 in
   301 
   302 fun check_autofix insts elems concl ctxt =
   303   let
   304     val inst_cs = map extract_inst insts;
   305     val elem_css = map extract_elem elems;
   306     val concl_cs = (map o map) mk_propp concl;
   307     (* Type inference *)
   308     val (inst_cs' :: css', ctxt') =
   309       (fold_burrow o fold_burrow) check (inst_cs :: elem_css @ [concl_cs]) ctxt;
   310     val (elem_css', [concl_cs']) = chop (length elem_css) css';
   311   in
   312     (map restore_inst (insts ~~ inst_cs'), map restore_elem (elems ~~ elem_css'),
   313       concl_cs', ctxt')
   314   end;
   315 
   316 end;
   317 
   318 
   319 (** Prepare locale elements **)
   320 
   321 fun declare_elem prep_vars (Fixes fixes) ctxt =
   322       let val (vars, _) = prep_vars fixes ctxt
   323       in ctxt |> ProofContext.add_fixes_i vars |> snd end
   324   | declare_elem prep_vars (Constrains csts) ctxt =
   325       ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
   326   | declare_elem _ (Assumes _) ctxt = ctxt
   327   | declare_elem _ (Defines _) ctxt = ctxt
   328   | declare_elem _ (Notes _) ctxt = ctxt;
   329 
   330 (** Finish locale elements, extract specification text **)
   331 
   332 val norm_term = Envir.beta_norm oo Term.subst_atomic;
   333 
   334 fun abstract_thm thy eq =
   335   Thm.assume (Thm.cterm_of thy eq) |> Drule.gen_all |> Drule.abs_def;
   336 
   337 fun bind_def ctxt eq (xs, env, ths) =
   338   let
   339     val ((y, T), b) = LocalDefs.abs_def eq;
   340     val b' = norm_term env b;
   341     val th = abstract_thm (ProofContext.theory_of ctxt) eq;
   342     fun err msg = error (msg ^ ": " ^ quote y);
   343   in
   344     exists (fn (x, _) => x = y) xs andalso
   345       err "Attempt to define previously specified variable";
   346     exists (fn (Free (y', _), _) => y = y' | _ => false) env andalso
   347       err "Attempt to redefine variable";
   348     (Term.add_frees b' xs, (Free (y, T), b') :: env, th :: ths)
   349   end;
   350 
   351 fun eval_text _ _ (Fixes _) text = text
   352   | eval_text _ _ (Constrains _) text = text
   353   | eval_text _ is_ext (Assumes asms)
   354         (((exts, exts'), (ints, ints')), (xs, env, defs)) =
   355       let
   356         val ts = maps (map #1 o #2) asms;
   357         val ts' = map (norm_term env) ts;
   358         val spec' =
   359           if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
   360           else ((exts, exts'), (ints @ ts, ints' @ ts'));
   361       in (spec', (fold Term.add_frees ts' xs, env, defs)) end
   362   | eval_text ctxt _ (Defines defs) (spec, binds) =
   363       (spec, fold (bind_def ctxt o #1 o #2) defs binds)
   364   | eval_text _ _ (Notes _) text = text;
   365 
   366 fun closeup _ _ false elem = elem
   367   | closeup ctxt parms true elem =
   368       let
   369         fun close_frees t =
   370           let
   371             val rev_frees =
   372               Term.fold_aterms (fn Free (x, T) =>
   373                 if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
   374           in Term.list_all_free (rev rev_frees, t) end;
   375 
   376         fun no_binds [] = []
   377           | no_binds _ = error "Illegal term bindings in context element";
   378       in
   379         (case elem of
   380           Assumes asms => Assumes (asms |> map (fn (a, propps) =>
   381             (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
   382         | Defines defs => Defines (defs |> map (fn (a, (t, ps)) =>
   383             (a, (close_frees (#2 (LocalDefs.cert_def ctxt t)), no_binds ps))))
   384         | e => e)
   385       end;
   386 
   387 fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
   388       let val x = Binding.base_name binding
   389       in (binding, AList.lookup (op =) parms x, mx) end) fixes)
   390   | finish_primitive _ _ (Constrains _) = Constrains []
   391   | finish_primitive _ close (Assumes asms) = close (Assumes asms)
   392   | finish_primitive _ close (Defines defs) = close (Defines defs)
   393   | finish_primitive _ _ (Notes facts) = Notes facts;
   394 
   395 fun finish_inst ctxt parms do_close (loc, (prfx, inst)) text =
   396   let
   397     val thy = ProofContext.theory_of ctxt;
   398     val (parm_names, parm_types) = NewLocale.params_of thy loc |>
   399       map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
   400     val (asm, defs) = NewLocale.specification_of thy loc;
   401     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
   402     val asm' = Option.map (Morphism.term morph) asm;
   403     val defs' = map (Morphism.term morph) defs;
   404     val text' = text |>
   405       (if is_some asm
   406         then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
   407         else I) |>
   408       (if not (null defs)
   409         then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
   410         else I)
   411 (* FIXME clone from new_locale.ML *)
   412   in ((loc, morph), text') end;
   413 
   414 fun finish_elem ctxt parms do_close elem text =
   415   let
   416     val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
   417     val text' = eval_text ctxt true elem' text;
   418   in (elem', text') end
   419   
   420 fun finish ctxt parms do_close insts elems text =
   421   let
   422     val (deps, text') = fold_map (finish_inst ctxt parms do_close) insts text;
   423     val (elems', text'') = fold_map (finish_elem ctxt parms do_close) elems text';
   424   in (deps, elems', text'') end;
   425 
   426 
   427 (** Process full context statement: instantiations + elements + conclusion **)
   428 
   429 (* Interleave incremental parsing and type inference over entire parsed stretch. *)
   430 
   431 local
   432 
   433 fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
   434     strict do_close context raw_import raw_elems raw_concl =
   435   let
   436     val thy = ProofContext.theory_of context;
   437 
   438     val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
   439 
   440     fun prep_inst (loc, (prfx, inst)) (i, insts, ctxt) =
   441       let
   442         val (parm_names, parm_types) = NewLocale.params_of thy loc |>
   443           map (fn (b, SOME T, _) => (Binding.base_name b, T)) |> split_list;
   444         val inst' = parse_inst parm_names inst ctxt;
   445         val parm_types' = map (TypeInfer.paramify_vars o
   446           Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
   447         val inst'' = map2 TypeInfer.constrain parm_types' inst';
   448         val insts' = insts @ [(loc, (prfx, inst''))];
   449         val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
   450         val inst''' = insts'' |> List.last |> snd |> snd;
   451         val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
   452         val ctxt'' = NewLocale.activate_declarations thy (loc, morph) ctxt;
   453       in (i+1, insts', ctxt'') end;
   454   
   455     fun prep_elem raw_elem (insts, elems, ctxt) =
   456       let
   457         val ctxt' = declare_elem prep_vars raw_elem ctxt;
   458         val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
   459         (* FIXME term bindings *)
   460         val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
   461       in (insts, elems', ctxt') end;
   462 
   463     fun prep_concl raw_concl (insts, elems, ctxt) =
   464       let
   465         val concl = (map o map) (fn (t, ps) =>
   466           (parse_prop ctxt t, map (parse_prop ctxt) ps)) raw_concl;
   467       in check_autofix insts elems concl ctxt end;
   468 
   469     val fors = prep_vars fixed context |> fst;
   470     val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
   471     val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], NewLocale.clear_local_idents ctxt);
   472     val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
   473     val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
   474 
   475     (* Retrieve parameter types *)
   476     val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes) |
   477       _ => fn ps => ps) (Fixes fors :: elems) [];
   478     val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
   479     val parms = xs ~~ Ts;  (* params from expression and elements *)
   480 
   481     val Fixes fors' = finish_primitive parms I (Fixes fors);
   482     val (deps, elems', text) = finish ctxt' parms do_close insts elems ((([], []), ([], [])), ([], [], []));
   483     (* text has the following structure:
   484            (((exts, exts'), (ints, ints')), (xs, env, defs))
   485        where
   486          exts: external assumptions (terms in assumes elements)
   487          exts': dito, normalised wrt. env
   488          ints: internal assumptions (terms in assumptions from insts)
   489          ints': dito, normalised wrt. env
   490          xs: the free variables in exts' and ints' and rhss of definitions,
   491            this includes parameters except defined parameters
   492          env: list of term pairs encoding substitutions, where the first term
   493            is a free variable; substitutions represent defines elements and
   494            the rhs is normalised wrt. the previous env
   495          defs: theorems representing the substitutions from defines elements
   496            (thms are normalised wrt. env).
   497        elems is an updated version of raw_elems:
   498          - type info added to Fixes and modified in Constrains
   499          - axiom and definition statement replaced by corresponding one
   500            from proppss in Assumes and Defines
   501          - Facts unchanged
   502        *)
   503 
   504   in ((fors', deps, elems', concl), (parms, text)) end
   505 
   506 in
   507 
   508 fun read_full_context_statement x =
   509   prep_full_context_statement Syntax.parse_typ Syntax.parse_prop parse_inst
   510   ProofContext.read_vars intern x;
   511 fun cert_full_context_statement x =
   512   prep_full_context_statement (K I) (K I) make_inst ProofContext.cert_vars (K I) x;
   513 
   514 end;
   515 
   516 
   517 (* Context statement: elements + conclusion *)
   518 
   519 local
   520 
   521 fun prep_statement prep activate raw_elems raw_concl context =
   522   let
   523      val ((_, _, elems, concl), _) = prep true false context ([], []) raw_elems raw_concl;
   524      val (_, context') = activate elems (ProofContext.set_stmt true context);
   525   in (concl, context') end;
   526 
   527 in
   528 
   529 fun read_statement x = prep_statement read_full_context_statement Element.activate x;
   530 fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
   531 
   532 end;
   533 
   534 
   535 (* Locale declaration: import + elements *)
   536 
   537 local
   538 
   539 fun prep_declaration prep activate raw_import raw_elems context =
   540   let
   541     val thy = ProofContext.theory_of context;
   542 
   543     val ((fixed, deps, elems, _), (parms, (spec, (_, _, defs)))) =
   544       prep false true context raw_import raw_elems [];
   545     (* Declare parameters and imported facts *)
   546     val context' = context |>
   547       ProofContext.add_fixes_i fixed |> snd |>
   548       NewLocale.clear_local_idents |> fold (NewLocale.activate_local_facts thy) deps;
   549     val ((elems', _), _) = activate elems (ProofContext.set_stmt true context');
   550   in ((fixed, deps, elems'), (parms, spec, defs)) end;
   551 
   552 in
   553 
   554 fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
   555 fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
   556 
   557 end;
   558 
   559 
   560 (* Locale expression to set up a goal *)
   561 
   562 local
   563 
   564 fun props_of thy (name, morph) =
   565   let
   566     val (asm, defs) = NewLocale.specification_of thy name;
   567   in
   568     (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
   569   end;
   570 
   571 fun prep_goal_expression prep expression context =
   572   let
   573     val thy = ProofContext.theory_of context;
   574 
   575     val ((fixed, deps, _, _), _) = prep true true context expression [] [];
   576     (* proof obligations *)
   577     val propss = map (props_of thy) deps;
   578 
   579     val goal_ctxt = context |>
   580       ProofContext.add_fixes_i fixed |> snd |>
   581       (fold o fold) Variable.auto_fixes propss;
   582 
   583     val export = Variable.export_morphism goal_ctxt context;
   584     val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
   585 (*    val exp_term = TermSubst.zero_var_indexes o Morphism.term export; *)
   586     val exp_term = Drule.term_rule thy (singleton exp_fact);
   587     val exp_typ = Logic.type_map exp_term;
   588     val export' =
   589       Morphism.morphism {binding = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   590   in ((propss, deps, export'), goal_ctxt) end;
   591     
   592 in
   593 
   594 fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
   595 fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
   596 
   597 end;
   598 
   599 
   600 (*** Locale declarations ***)
   601 
   602 (* axiomsN: name of theorem set with destruct rules for locale predicates,
   603      also name suffix of delta predicates and assumptions. *)
   604 
   605 val axiomsN = "axioms";
   606 
   607 local
   608 
   609 (* introN: name of theorems for introduction rules of locale and
   610      delta predicates *)
   611 
   612 val introN = "intro";
   613 
   614 fun atomize_spec thy ts =
   615   let
   616     val t = Logic.mk_conjunction_balanced ts;
   617     val body = ObjectLogic.atomize_term thy t;
   618     val bodyT = Term.fastype_of body;
   619   in
   620     if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
   621     else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
   622   end;
   623 
   624 (* achieve plain syntax for locale predicates (without "PROP") *)
   625 
   626 fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
   627   if length args = n then
   628     Syntax.const "_aprop" $
   629       Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
   630   else raise Match);
   631 
   632 (* define one predicate including its intro rule and axioms
   633    - bname: predicate name
   634    - parms: locale parameters
   635    - defs: thms representing substitutions from defines elements
   636    - ts: terms representing locale assumptions (not normalised wrt. defs)
   637    - norm_ts: terms representing locale assumptions (normalised wrt. defs)
   638    - thy: the theory
   639 *)
   640 
   641 fun def_pred bname parms defs ts norm_ts thy =
   642   let
   643     val name = Sign.full_bname thy bname;
   644 
   645     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
   646     val env = Term.add_term_free_names (body, []);
   647     val xs = filter (member (op =) env o #1) parms;
   648     val Ts = map #2 xs;
   649     val extraTs = (Term.term_tfrees body \\ fold Term.add_tfreesT Ts [])
   650       |> Library.sort_wrt #1 |> map TFree;
   651     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
   652 
   653     val args = map Logic.mk_type extraTs @ map Free xs;
   654     val head = Term.list_comb (Const (name, predT), args);
   655     val statement = ObjectLogic.ensure_propT thy head;
   656 
   657     val ([pred_def], defs_thy) =
   658       thy
   659       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
   660       |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
   661       |> PureThy.add_defs false
   662         [((Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
   663     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
   664 
   665     val cert = Thm.cterm_of defs_thy;
   666 
   667     val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
   668       MetaSimplifier.rewrite_goals_tac [pred_def] THEN
   669       Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
   670       Tactic.compose_tac (false,
   671         Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
   672 
   673     val conjuncts =
   674       (Drule.equal_elim_rule2 OF [body_eq,
   675         MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
   676       |> Conjunction.elim_balanced (length ts);
   677     val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
   678       Element.prove_witness defs_ctxt t
   679        (MetaSimplifier.rewrite_goals_tac defs THEN
   680         Tactic.compose_tac (false, ax, 0) 1));
   681   in ((statement, intro, axioms), defs_thy) end;
   682 
   683 in
   684 
   685 (* CB: main predicate definition function *)
   686 
   687 fun define_preds pname (parms, ((exts, exts'), (ints, ints')), defs) thy =
   688   let
   689     val (a_pred, a_intro, a_axioms, thy'') =
   690       if null exts then (NONE, NONE, [], thy)
   691       else
   692         let
   693           val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
   694           val ((statement, intro, axioms), thy') =
   695             thy
   696             |> def_pred aname parms defs exts exts';
   697           val (_, thy'') =
   698             thy'
   699             |> Sign.add_path aname
   700             |> Sign.no_base_names
   701             |> PureThy.note_thmss Thm.internalK
   702               [((Binding.name introN, []), [([intro], [NewLocale.unfold_attrib])])]
   703             ||> Sign.restore_naming thy';
   704           in (SOME statement, SOME intro, axioms, thy'') end;
   705     val (b_pred, b_intro, b_axioms, thy'''') =
   706       if null ints then (NONE, NONE, [], thy'')
   707       else
   708         let
   709           val ((statement, intro, axioms), thy''') =
   710             thy''
   711             |> def_pred pname parms defs (ints @ the_list a_pred) (ints' @ the_list a_pred);
   712           val (_, thy'''') =
   713             thy'''
   714             |> Sign.add_path pname
   715             |> Sign.no_base_names
   716             |> PureThy.note_thmss Thm.internalK
   717                  [((Binding.name introN, []), [([intro], [NewLocale.intro_attrib])]),
   718                   ((Binding.name axiomsN, []),
   719                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
   720             ||> Sign.restore_naming thy''';
   721         in (SOME statement, SOME intro, axioms, thy'''') end;
   722   in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
   723 
   724 end;
   725 
   726 
   727 local
   728 
   729 fun assumes_to_notes (Assumes asms) axms =
   730       fold_map (fn (a, spec) => fn axs =>
   731           let val (ps, qs) = chop (length spec) axs
   732           in ((a, [(ps, [])]), qs) end) asms axms
   733       |> apfst (curry Notes Thm.assumptionK)
   734   | assumes_to_notes e axms = (e, axms);
   735 
   736 fun defines_to_notes thy (Defines defs) defns =
   737     let
   738       val defs' = map (fn (_, (def, _)) => def) defs
   739       val notes = map (fn (a, (def, _)) =>
   740         (a, [([assume (cterm_of thy def)], [])])) defs
   741     in
   742       (Notes (Thm.definitionK, notes), defns @ defs')
   743     end
   744   | defines_to_notes _ e defns = (e, defns);
   745 
   746 fun gen_add_locale prep_decl
   747     bname predicate_name raw_imprt raw_body thy =
   748   let
   749     val thy_ctxt = ProofContext.init thy;
   750     val name = Sign.full_bname thy bname;
   751     val _ = NewLocale.test_locale thy name andalso
   752       error ("Duplicate definition of locale " ^ quote name);
   753 
   754     val ((fixed, deps, body_elems), text as (parms, ((_, exts'), _), defs)) =
   755       prep_decl raw_imprt raw_body thy_ctxt;
   756     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
   757       define_preds predicate_name text thy;
   758 
   759     val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
   760     val _ = if null extraTs then ()
   761       else warning ("Additional type variable(s) in locale specification " ^ quote bname);
   762 
   763     val satisfy = Element.satisfy_morphism b_axioms;
   764 
   765     val params = fixed @
   766       (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
   767     val asm = if is_some b_statement then b_statement else a_statement;
   768     val (body_elems', defns) = fold_map (defines_to_notes thy') body_elems [];
   769     val notes = body_elems' |>
   770       (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
   771       fst |> map (Element.morph_ctxt satisfy) |>
   772       map_filter (fn Notes notes => SOME notes | _ => NONE) |>
   773       (if is_some asm
   774         then cons (Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
   775           [([assume (cterm_of thy' (the asm))], [(Attrib.internal o K) NewLocale.witness_attrib])])])
   776         else I);
   777 
   778     val deps' = map (fn (l, morph) => (l, morph $> satisfy)) deps;
   779 
   780     val loc_ctxt = thy' |>
   781       NewLocale.register_locale bname (extraTs, params)
   782         (asm, map prop_of defs) ([], [])
   783         (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
   784       NewLocale.init name
   785   in (name, loc_ctxt) end;
   786 
   787 in
   788 
   789 val add_locale_cmd = gen_add_locale read_declaration;
   790 val add_locale = gen_add_locale cert_declaration;
   791 
   792 end;
   793 
   794 
   795 (*** Interpretation ***)
   796 
   797 (** Witnesses and goals **)
   798 
   799 fun prep_propp propss = propss |> map (map (rpair [] o Element.mark_witness));
   800 
   801 fun prep_result propps thmss =
   802   ListPair.map (fn (props, thms) => map2 Element.make_witness props thms) (propps, thmss);
   803 
   804 
   805 (** Interpretation between locales: declaring sublocale relationships **)
   806 
   807 local
   808 
   809 fun gen_sublocale prep_expr intern
   810     raw_target expression thy =
   811   let
   812     val target = intern thy raw_target;
   813     val target_ctxt = NewLocale.init target thy;
   814 
   815     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
   816     
   817     fun store_dep ((name, morph), thms) =
   818       NewLocale.add_dependency target (name, morph $> Element.satisfy_morphism thms $> export);
   819 
   820     fun after_qed results =
   821       ProofContext.theory (
   822         (* store dependencies *)
   823         fold store_dep (deps ~~ prep_result propss results) #>
   824         (* propagate registrations *)
   825         (fn thy => fold_rev (fn reg => NewLocale.activate_global_facts reg)
   826           (NewLocale.get_global_registrations thy) thy));
   827   in
   828     goal_ctxt |>
   829       Proof.theorem_i NONE after_qed (prep_propp propss) |>
   830       Element.refine_witness |> Seq.hd
   831   end;
   832 
   833 in
   834 
   835 fun sublocale_cmd x = gen_sublocale read_goal_expression NewLocale.intern x;
   836 fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
   837 
   838 end;
   839 
   840 (** Interpretation in theories **)
   841 
   842 local
   843 
   844 fun gen_interpretation prep_expr
   845     expression thy =
   846   let
   847     val ctxt = ProofContext.init thy;
   848 
   849     val ((propss, deps, export), goal_ctxt) = prep_expr expression ctxt;
   850     
   851     fun store_reg ((name, morph), thms) =
   852       let
   853         val morph' = morph $> Element.satisfy_morphism thms $> export;
   854       in
   855         NewLocale.add_global_registration (name, morph') #>
   856         NewLocale.activate_global_facts (name, morph')
   857       end;
   858 
   859     fun after_qed results =
   860       ProofContext.theory (fold store_reg (deps ~~ prep_result propss results));
   861   in
   862     goal_ctxt |>
   863       Proof.theorem_i NONE after_qed (prep_propp propss) |>
   864       Element.refine_witness |> Seq.hd
   865   end;
   866 
   867 in
   868 
   869 fun interpretation_cmd x = gen_interpretation read_goal_expression x;
   870 fun interpretation x = gen_interpretation cert_goal_expression x;
   871 
   872 end;
   873 
   874 end;
   875