src/Pure/Isar/expression.ML
author wenzelm
Sat, 16 Apr 2011 15:47:52 +0200
changeset 43231 da8817d01e7c
parent 43230 6ca5407863ed
child 43246 774df7c59508
permissions -rw-r--r--
modernized structure Proof_Context;
     1 (*  Title:      Pure/Isar/expression.ML
     2     Author:     Clemens Ballarin, TU Muenchen
     3 
     4 Locale expressions and user interface layer of locales.
     5 *)
     6 
     7 signature EXPRESSION =
     8 sig
     9   (* Locale expressions *)
    10   datatype 'term map = Positional of 'term option list | Named of (string * 'term) list
    11   type 'term expr = (string * ((string * bool) * 'term map)) list
    12   type expression_i = term expr * (binding * typ option * mixfix) list
    13   type expression = string expr * (binding * string option * mixfix) list
    14 
    15   (* Processing of context statements *)
    16   val cert_statement: Element.context_i list -> (term * term list) list list ->
    17     Proof.context -> (term * term list) list list * Proof.context
    18   val read_statement: Element.context list -> (string * string list) list list ->
    19     Proof.context -> (term * term list) list list * Proof.context
    20 
    21   (* Declaring locales *)
    22   val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list ->
    23     Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list
    24       * Element.context_i list) * ((string * typ) list * Proof.context)
    25   val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context list ->
    26     Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list
    27       * Element.context_i list) * ((string * typ) list * Proof.context)
    28       (*FIXME*)
    29   val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list ->
    30     Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list
    31       * Element.context_i list) * ((string * typ) list * Proof.context)
    32   val add_locale: (local_theory -> local_theory) -> binding -> binding ->
    33     expression_i -> Element.context_i list -> theory -> string * local_theory
    34   val add_locale_cmd: (local_theory -> local_theory) -> binding -> binding ->
    35     expression -> Element.context list -> theory -> string * local_theory
    36 
    37   (* Interpretation *)
    38   val cert_goal_expression: expression_i -> Proof.context ->
    39     (term list list * (string * morphism) list * morphism) * Proof.context
    40   val read_goal_expression: expression -> Proof.context ->
    41     (term list list * (string * morphism) list * morphism) * Proof.context
    42   val sublocale: (local_theory -> local_theory) -> string -> expression_i ->
    43     (Attrib.binding * term) list -> theory -> Proof.state
    44   val sublocale_cmd: (local_theory -> local_theory) -> string -> expression ->
    45     (Attrib.binding * string) list -> theory -> Proof.state
    46   val interpretation: expression_i -> (Attrib.binding * term) list ->
    47     theory -> Proof.state
    48   val interpretation_cmd: expression -> (Attrib.binding * string) list ->
    49     theory -> Proof.state
    50   val interpret: expression_i -> (Attrib.binding * term) list ->
    51     bool -> Proof.state -> Proof.state
    52   val interpret_cmd: expression -> (Attrib.binding * string) list ->
    53     bool -> Proof.state -> Proof.state
    54 
    55   (* Diagnostic *)
    56   val print_dependencies: Proof.context -> bool -> expression -> unit
    57 end;
    58 
    59 structure Expression : EXPRESSION =
    60 struct
    61 
    62 datatype ctxt = datatype Element.ctxt;
    63 
    64 
    65 (*** Expressions ***)
    66 
    67 datatype 'term map =
    68   Positional of 'term option list |
    69   Named of (string * 'term) list;
    70 
    71 type 'term expr = (string * ((string * bool) * 'term map)) list;
    72 
    73 type expression = string expr * (binding * string option * mixfix) list;
    74 type expression_i = term expr * (binding * typ option * mixfix) list;
    75 
    76 
    77 (** Internalise locale names in expr **)
    78 
    79 fun intern thy instances =  map (apfst (Locale.intern thy)) instances;
    80 
    81 
    82 (** Parameters of expression **)
    83 
    84 (*Sanity check of instantiations and extraction of implicit parameters.
    85   The latter only occurs iff strict = false.
    86   Positional instantiations are extended to match full length of parameter list
    87   of instantiated locale.*)
    88 
    89 fun parameters_of thy strict (expr, fixed) =
    90   let
    91     fun reject_dups message xs =
    92       (case duplicates (op =) xs of
    93         [] => ()
    94       | dups => error (message ^ commas dups));
    95 
    96     fun parm_eq ((p1: string, mx1: mixfix), (p2, mx2)) = p1 = p2 andalso
    97       (mx1 = mx2 orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression"));
    98 
    99     fun params_loc loc = Locale.params_of thy loc |> map (apfst #1);
   100     fun params_inst (loc, (prfx, Positional insts)) =
   101           let
   102             val ps = params_loc loc;
   103             val d = length ps - length insts;
   104             val insts' =
   105               if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
   106                 quote (Locale.extern thy loc))
   107               else insts @ replicate d NONE;
   108             val ps' = (ps ~~ insts') |>
   109               map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
   110           in (ps', (loc, (prfx, Positional insts'))) end
   111       | params_inst (loc, (prfx, Named insts)) =
   112           let
   113             val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
   114               (map fst insts);
   115             val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps =>
   116               if AList.defined (op =) ps p then AList.delete (op =) p ps
   117               else error (quote p ^ " not a parameter of instantiated expression"));
   118           in (ps', (loc, (prfx, Named insts))) end;
   119     fun params_expr is =
   120       let
   121         val (is', ps') = fold_map (fn i => fn ps =>
   122           let
   123             val (ps', i') = params_inst i;
   124             val ps'' = distinct parm_eq (ps @ ps');
   125           in (i', ps'') end) is []
   126       in (ps', is') end;
   127 
   128     val (implicit, expr') = params_expr expr;
   129 
   130     val implicit' = map #1 implicit;
   131     val fixed' = map (Variable.name o #1) fixed;
   132     val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
   133     val implicit'' =
   134       if strict then []
   135       else
   136         let val _ = reject_dups
   137           "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed')
   138         in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end;
   139 
   140   in (expr', implicit'' @ fixed) end;
   141 
   142 
   143 (** Read instantiation **)
   144 
   145 (* Parse positional or named instantiation *)
   146 
   147 local
   148 
   149 fun prep_inst prep_term ctxt parms (Positional insts) =
   150       (insts ~~ parms) |> map (fn
   151         (NONE, p) => Free (p, dummyT) |
   152         (SOME t, _) => prep_term ctxt t)
   153   | prep_inst prep_term ctxt parms (Named insts) =
   154       parms |> map (fn p => case AList.lookup (op =) insts p of
   155         SOME t => prep_term ctxt t |
   156         NONE => Free (p, dummyT));
   157 
   158 in
   159 
   160 fun parse_inst x = prep_inst Syntax.parse_term x;
   161 fun make_inst x = prep_inst (K I) x;
   162 
   163 end;
   164 
   165 
   166 (* Instantiation morphism *)
   167 
   168 fun inst_morph (parm_names, parm_types) ((prfx, mandatory), insts') ctxt =
   169   let
   170     (* parameters *)
   171     val type_parm_names = fold Term.add_tfreesT parm_types [] |> map fst;
   172 
   173     (* type inference and contexts *)
   174     val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types;
   175     val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar);
   176     val arg = type_parms @ map2 Type.constraint parm_types' insts';
   177     val res = Syntax.check_terms ctxt arg;
   178     val ctxt' = ctxt |> fold Variable.auto_fixes res;
   179 
   180     (* instantiation *)
   181     val (type_parms'', res') = chop (length type_parms) res;
   182     val insts'' = (parm_names ~~ res') |> map_filter
   183       (fn inst as (x, Free (y, _)) => if x = y then NONE else SOME inst
   184         | inst => SOME inst);
   185     val instT = Symtab.make (type_parm_names ~~ map Logic.dest_type type_parms'');
   186     val inst = Symtab.make insts'';
   187   in
   188     (Element.inst_morphism (Proof_Context.theory_of ctxt) (instT, inst) $>
   189       Morphism.binding_morphism (Binding.prefix mandatory prfx), ctxt')
   190   end;
   191 
   192 
   193 (*** Locale processing ***)
   194 
   195 (** Parsing **)
   196 
   197 fun parse_elem prep_typ prep_term ctxt =
   198   Element.map_ctxt
   199    {binding = I,
   200     typ = prep_typ ctxt,
   201     term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt),
   202     pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt),
   203     fact = I,
   204     attrib = I};
   205 
   206 fun parse_concl prep_term ctxt concl =
   207   (map o map) (fn (t, ps) =>
   208     (prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t,
   209       map (prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) concl;
   210 
   211 
   212 (** Simultaneous type inference: instantiations + elements + conclusion **)
   213 
   214 local
   215 
   216 fun mk_type T = (Logic.mk_type T, []);
   217 fun mk_term t = (t, []);
   218 fun mk_propp (p, pats) = (Type.constraint propT p, pats);
   219 
   220 fun dest_type (T, []) = Logic.dest_type T;
   221 fun dest_term (t, []) = t;
   222 fun dest_propp (p, pats) = (p, pats);
   223 
   224 fun extract_inst (_, (_, ts)) = map mk_term ts;
   225 fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs));
   226 
   227 fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes
   228   | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts
   229   | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms
   230   | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs
   231   | extract_elem (Notes _) = [];
   232 
   233 fun restore_elem (Fixes fixes, css) =
   234       (fixes ~~ css) |> map (fn ((x, _, mx), cs) =>
   235         (x, cs |> map dest_type |> try hd, mx)) |> Fixes
   236   | restore_elem (Constrains csts, css) =
   237       (csts ~~ css) |> map (fn ((x, _), cs) =>
   238         (x, cs |> map dest_type |> hd)) |> Constrains
   239   | restore_elem (Assumes asms, css) =
   240       (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes
   241   | restore_elem (Defines defs, css) =
   242       (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines
   243   | restore_elem (Notes notes, _) = Notes notes;
   244 
   245 fun check cs context =
   246   let
   247     fun prep (_, pats) (ctxt, t :: ts) =
   248       let val ctxt' = Variable.auto_fixes t ctxt
   249       in
   250         ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats),
   251           (ctxt', ts))
   252       end;
   253     val (cs', (context', _)) = fold_map prep cs
   254       (context, Syntax.check_terms
   255         (Proof_Context.set_mode Proof_Context.mode_schematic context) (map fst cs));
   256   in (cs', context') end;
   257 
   258 in
   259 
   260 fun check_autofix insts elems concl ctxt =
   261   let
   262     val inst_cs = map extract_inst insts;
   263     val elem_css = map extract_elem elems;
   264     val concl_cs = (map o map) mk_propp concl;
   265     (* Type inference *)
   266     val (inst_cs' :: css', ctxt') =
   267       (fold_burrow o fold_burrow) check (inst_cs :: elem_css @ [concl_cs]) ctxt;
   268     val (elem_css', [concl_cs']) = chop (length elem_css) css';
   269   in
   270     (map restore_inst (insts ~~ inst_cs'),
   271       map restore_elem (elems ~~ elem_css'),
   272       concl_cs', ctxt')
   273   end;
   274 
   275 end;
   276 
   277 
   278 (** Prepare locale elements **)
   279 
   280 fun declare_elem prep_vars (Fixes fixes) ctxt =
   281       let val (vars, _) = prep_vars fixes ctxt
   282       in ctxt |> Proof_Context.add_fixes vars |> snd end
   283   | declare_elem prep_vars (Constrains csts) ctxt =
   284       ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
   285   | declare_elem _ (Assumes _) ctxt = ctxt
   286   | declare_elem _ (Defines _) ctxt = ctxt
   287   | declare_elem _ (Notes _) ctxt = ctxt;
   288 
   289 
   290 (** Finish locale elements **)
   291 
   292 fun closeup _ _ false elem = elem
   293   | closeup ctxt parms true elem =
   294       let
   295         (* FIXME consider closing in syntactic phase -- before type checking *)
   296         fun close_frees t =
   297           let
   298             val rev_frees =
   299               Term.fold_aterms (fn Free (x, T) =>
   300                 if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
   301           in fold (Logic.all o Free) rev_frees t end;
   302 
   303         fun no_binds [] = []
   304           | no_binds _ = error "Illegal term bindings in context element";
   305       in
   306         (case elem of
   307           Assumes asms => Assumes (asms |> map (fn (a, propps) =>
   308             (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
   309         | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) =>
   310             let val ((c, _), t') = Local_Defs.cert_def ctxt (close_frees t)
   311             in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end))
   312         | e => e)
   313       end;
   314 
   315 fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
   316       let val x = Binding.name_of binding
   317       in (binding, AList.lookup (op =) parms x, mx) end) fixes)
   318   | finish_primitive _ _ (Constrains _) = Constrains []
   319   | finish_primitive _ close (Assumes asms) = close (Assumes asms)
   320   | finish_primitive _ close (Defines defs) = close (Defines defs)
   321   | finish_primitive _ _ (Notes facts) = Notes facts;
   322 
   323 fun finish_inst ctxt (loc, (prfx, inst)) =
   324   let
   325     val thy = Proof_Context.theory_of ctxt;
   326     val (parm_names, parm_types) = Locale.params_of thy loc |> map #1 |> split_list;
   327     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
   328   in (loc, morph) end;
   329 
   330 fun finish_elem ctxt parms do_close elem =
   331   finish_primitive parms (closeup ctxt parms do_close) elem;
   332 
   333 fun finish ctxt parms do_close insts elems =
   334   let
   335     val deps = map (finish_inst ctxt) insts;
   336     val elems' = map (finish_elem ctxt parms do_close) elems;
   337   in (deps, elems') end;
   338 
   339 
   340 (** Process full context statement: instantiations + elements + conclusion **)
   341 
   342 (* Interleave incremental parsing and type inference over entire parsed stretch. *)
   343 
   344 local
   345 
   346 fun prep_full_context_statement parse_typ parse_prop prep_vars_elem prep_inst prep_vars_inst prep_expr
   347     {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_concl ctxt1 =
   348   let
   349     val thy = Proof_Context.theory_of ctxt1;
   350 
   351     val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
   352 
   353     fun prep_insts_cumulative (loc, (prfx, inst)) (i, insts, ctxt) =
   354       let
   355         val (parm_names, parm_types) = Locale.params_of thy loc |> map #1 |> split_list;
   356         val inst' = prep_inst ctxt parm_names inst;
   357         val parm_types' = map (Type_Infer.paramify_vars o
   358           Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT_global) parm_types;
   359         val inst'' = map2 Type.constraint parm_types' inst';
   360         val insts' = insts @ [(loc, (prfx, inst''))];
   361         val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
   362         val inst''' = insts'' |> List.last |> snd |> snd;
   363         val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
   364         val ctxt'' = Locale.activate_declarations (loc, morph) ctxt;
   365       in (i + 1, insts', ctxt'') end;
   366 
   367     fun prep_elem insts raw_elem (elems, ctxt) =
   368       let
   369         val ctxt' = declare_elem prep_vars_elem raw_elem ctxt;
   370         val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
   371         val (_, _, _, ctxt'' (* FIXME not used *) ) = check_autofix insts elems' [] ctxt';
   372       in (elems', ctxt') end;
   373 
   374     fun prep_concl raw_concl (insts, elems, ctxt) =
   375       let
   376         val concl = parse_concl parse_prop ctxt raw_concl;
   377       in check_autofix insts elems concl ctxt end;
   378 
   379     val fors = prep_vars_inst fixed ctxt1 |> fst;
   380     val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd;
   381     val (_, insts', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], ctxt2);
   382 
   383     val add_free = fold_aterms
   384       (fn Free (x, T) => not (Variable.is_fixed ctxt3 x) ? insert (op =) (x, T) | _ => I);
   385     val _ =
   386       if fixed_frees then ()
   387       else
   388         (case fold (fold add_free o snd o snd) insts' [] of
   389           [] => ()
   390         | frees => error ("Illegal free variables in expression: " ^
   391             commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees))));
   392 
   393     val ctxt4 = init_body ctxt3;
   394     val (elems, ctxt5) = fold (prep_elem insts') raw_elems ([], ctxt4);
   395     val (insts, elems', concl, ctxt6) = prep_concl raw_concl (insts', elems, ctxt5);
   396 
   397     (* Retrieve parameter types *)
   398     val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Variable.name o #1) fixes)
   399       | _ => fn ps => ps) (Fixes fors :: elems') [];
   400     val (Ts, ctxt7) = fold_map Proof_Context.inferred_param xs ctxt6;
   401     val parms = xs ~~ Ts;  (* params from expression and elements *)
   402 
   403     val Fixes fors' = finish_primitive parms I (Fixes fors);
   404     val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors';
   405     val (deps, elems'') = finish ctxt6 parms do_close insts elems';
   406 
   407   in ((fixed, deps, elems'', concl), (parms, ctxt7)) end
   408 
   409 in
   410 
   411 fun cert_full_context_statement x =
   412   prep_full_context_statement (K I) (K I) Proof_Context.cert_vars
   413   make_inst Proof_Context.cert_vars (K I) x;
   414 
   415 fun cert_read_full_context_statement x =
   416   prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Proof_Context.read_vars
   417   make_inst Proof_Context.cert_vars (K I) x;
   418 
   419 fun read_full_context_statement x =
   420   prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Proof_Context.read_vars
   421   parse_inst Proof_Context.read_vars intern x;
   422 
   423 end;
   424 
   425 
   426 (* Context statement: elements + conclusion *)
   427 
   428 local
   429 
   430 fun prep_statement prep activate raw_elems raw_concl context =
   431   let
   432      val ((_, _, elems, concl), _) =
   433        prep {strict = true, do_close = false, fixed_frees = true}
   434         ([], []) I raw_elems raw_concl context;
   435      val (_, context') = context |>
   436        Proof_Context.set_stmt true |>
   437        fold_map activate elems;
   438   in (concl, context') end;
   439 
   440 in
   441 
   442 fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
   443 fun read_statement x = prep_statement read_full_context_statement Element.activate x;
   444 
   445 end;
   446 
   447 
   448 (* Locale declaration: import + elements *)
   449 
   450 fun fix_params params =
   451   Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd;
   452 
   453 local
   454 
   455 fun prep_declaration prep activate raw_import init_body raw_elems context =
   456   let
   457     val ((fixed, deps, elems, _), (parms, ctxt')) =
   458       prep {strict = false, do_close = true, fixed_frees = false}
   459         raw_import init_body raw_elems [] context;
   460     (* Declare parameters and imported facts *)
   461     val context' = context |>
   462       fix_params fixed |>
   463       fold (Context.proof_map o Locale.activate_facts NONE) deps;
   464     val (elems', _) = context' |>
   465       Proof_Context.set_stmt true |>
   466       fold_map activate elems;
   467   in ((fixed, deps, elems'), (parms, ctxt')) end;
   468 
   469 in
   470 
   471 fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
   472 fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x;
   473 fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
   474 
   475 end;
   476 
   477 
   478 (* Locale expression to set up a goal *)
   479 
   480 local
   481 
   482 fun props_of thy (name, morph) =
   483   let
   484     val (asm, defs) = Locale.specification_of thy name;
   485   in
   486     (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
   487   end;
   488 
   489 fun prep_goal_expression prep expression context =
   490   let
   491     val thy = Proof_Context.theory_of context;
   492 
   493     val ((fixed, deps, _, _), _) =
   494       prep {strict = true, do_close = true, fixed_frees = true} expression I [] [] context;
   495     (* proof obligations *)
   496     val propss = map (props_of thy) deps;
   497 
   498     val goal_ctxt = context |>
   499       fix_params fixed |>
   500       (fold o fold) Variable.auto_fixes propss;
   501 
   502     val export = Variable.export_morphism goal_ctxt context;
   503     val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
   504     val exp_term = Term_Subst.zero_var_indexes o Morphism.term export;
   505     val exp_typ = Logic.type_map exp_term;
   506     val export' = Morphism.morphism {binding = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   507   in ((propss, deps, export'), goal_ctxt) end;
   508 
   509 in
   510 
   511 fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
   512 fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
   513 
   514 end;
   515 
   516 
   517 (*** Locale declarations ***)
   518 
   519 (* extract specification text *)
   520 
   521 val norm_term = Envir.beta_norm oo Term.subst_atomic;
   522 
   523 fun bind_def ctxt eq (xs, env, eqs) =
   524   let
   525     val _ = Local_Defs.cert_def ctxt eq;
   526     val ((y, T), b) = Local_Defs.abs_def eq;
   527     val b' = norm_term env b;
   528     fun err msg = error (msg ^ ": " ^ quote y);
   529   in
   530     case filter (fn (Free (y', _), _) => y = y' | _ => false) env of
   531       [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) |
   532       dups => if forall (fn (_, b'') => b' aconv b'') dups
   533         then (xs, env, eqs)
   534         else err "Attempt to redefine variable"
   535   end;
   536 
   537 (* text has the following structure:
   538        (((exts, exts'), (ints, ints')), (xs, env, defs))
   539    where
   540      exts: external assumptions (terms in assumes elements)
   541      exts': dito, normalised wrt. env
   542      ints: internal assumptions (terms in assumptions from insts)
   543      ints': dito, normalised wrt. env
   544      xs: the free variables in exts' and ints' and rhss of definitions,
   545        this includes parameters except defined parameters
   546      env: list of term pairs encoding substitutions, where the first term
   547        is a free variable; substitutions represent defines elements and
   548        the rhs is normalised wrt. the previous env
   549      defs: the equations from the defines elements
   550    *)
   551 
   552 fun eval_text _ _ (Fixes _) text = text
   553   | eval_text _ _ (Constrains _) text = text
   554   | eval_text _ is_ext (Assumes asms)
   555         (((exts, exts'), (ints, ints')), (xs, env, defs)) =
   556       let
   557         val ts = maps (map #1 o #2) asms;
   558         val ts' = map (norm_term env) ts;
   559         val spec' =
   560           if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
   561           else ((exts, exts'), (ints @ ts, ints' @ ts'));
   562       in (spec', (fold Term.add_frees ts' xs, env, defs)) end
   563   | eval_text ctxt _ (Defines defs) (spec, binds) =
   564       (spec, fold (bind_def ctxt o #1 o #2) defs binds)
   565   | eval_text _ _ (Notes _) text = text;
   566 
   567 fun eval_inst ctxt (loc, morph) text =
   568   let
   569     val thy = Proof_Context.theory_of ctxt;
   570     val (asm, defs) = Locale.specification_of thy loc;
   571     val asm' = Option.map (Morphism.term morph) asm;
   572     val defs' = map (Morphism.term morph) defs;
   573     val text' = text |>
   574       (if is_some asm
   575         then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
   576         else I) |>
   577       (if not (null defs)
   578         then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
   579         else I)
   580 (* FIXME clone from locale.ML *)
   581   in text' end;
   582 
   583 fun eval_elem ctxt elem text =
   584   eval_text ctxt true elem text;
   585 
   586 fun eval ctxt deps elems =
   587   let
   588     val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [], []));
   589     val ((spec, (_, _, defs))) = fold (eval_elem ctxt) elems text';
   590   in (spec, defs) end;
   591 
   592 (* axiomsN: name of theorem set with destruct rules for locale predicates,
   593      also name suffix of delta predicates and assumptions. *)
   594 
   595 val axiomsN = "axioms";
   596 
   597 local
   598 
   599 (* introN: name of theorems for introduction rules of locale and
   600      delta predicates *)
   601 
   602 val introN = "intro";
   603 
   604 fun atomize_spec thy ts =
   605   let
   606     val t = Logic.mk_conjunction_balanced ts;
   607     val body = Object_Logic.atomize_term thy t;
   608     val bodyT = Term.fastype_of body;
   609   in
   610     if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
   611     else (body, bodyT, Object_Logic.atomize (Thm.cterm_of thy t))
   612   end;
   613 
   614 (* achieve plain syntax for locale predicates (without "PROP") *)
   615 
   616 fun aprop_tr' n c = (Lexicon.mark_const c, fn ctxt => fn args =>
   617   if length args = n then
   618     Syntax.const "_aprop" $   (* FIXME avoid old-style early externing (!??) *)
   619       Term.list_comb (Syntax.free (Proof_Context.extern_const ctxt c), args)
   620   else raise Match);
   621 
   622 (* define one predicate including its intro rule and axioms
   623    - binding: predicate name
   624    - parms: locale parameters
   625    - defs: thms representing substitutions from defines elements
   626    - ts: terms representing locale assumptions (not normalised wrt. defs)
   627    - norm_ts: terms representing locale assumptions (normalised wrt. defs)
   628    - thy: the theory
   629 *)
   630 
   631 fun def_pred binding parms defs ts norm_ts thy =
   632   let
   633     val name = Sign.full_name thy binding;
   634 
   635     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
   636     val env = Term.add_free_names body [];
   637     val xs = filter (member (op =) env o #1) parms;
   638     val Ts = map #2 xs;
   639     val extraTs =
   640       (subtract (op =) (fold Term.add_tfreesT Ts []) (Term.add_tfrees body []))
   641       |> Library.sort_wrt #1 |> map TFree;
   642     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
   643 
   644     val args = map Logic.mk_type extraTs @ map Free xs;
   645     val head = Term.list_comb (Const (name, predT), args);
   646     val statement = Object_Logic.ensure_propT thy head;
   647 
   648     val ([pred_def], defs_thy) =
   649       thy
   650       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
   651       |> Sign.declare_const ((Binding.conceal binding, predT), NoSyn) |> snd
   652       |> Global_Theory.add_defs false
   653         [((Binding.conceal (Thm.def_binding binding), Logic.mk_equals (head, body)), [])];
   654     val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head;
   655 
   656     val cert = Thm.cterm_of defs_thy;
   657 
   658     val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
   659       rewrite_goals_tac [pred_def] THEN
   660       Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
   661       Tactic.compose_tac (false,
   662         Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
   663 
   664     val conjuncts =
   665       (Drule.equal_elim_rule2 OF [body_eq,
   666         Raw_Simplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
   667       |> Conjunction.elim_balanced (length ts);
   668     val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
   669       Element.prove_witness defs_ctxt t
   670        (rewrite_goals_tac defs THEN
   671         Tactic.compose_tac (false, ax, 0) 1));
   672   in ((statement, intro, axioms), defs_thy) end;
   673 
   674 in
   675 
   676 (* main predicate definition function *)
   677 
   678 fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy =
   679   let
   680     val defs' = map (cterm_of thy #> Assumption.assume #> Drule.abs_def) defs;
   681 
   682     val (a_pred, a_intro, a_axioms, thy'') =
   683       if null exts then (NONE, NONE, [], thy)
   684       else
   685         let
   686           val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding;
   687           val ((statement, intro, axioms), thy') =
   688             thy
   689             |> def_pred abinding parms defs' exts exts';
   690           val (_, thy'') =
   691             thy'
   692             |> Sign.qualified_path true abinding
   693             |> Global_Theory.note_thmss ""
   694               [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.unfold_add])])]
   695             ||> Sign.restore_naming thy';
   696           in (SOME statement, SOME intro, axioms, thy'') end;
   697     val (b_pred, b_intro, b_axioms, thy'''') =
   698       if null ints then (NONE, NONE, [], thy'')
   699       else
   700         let
   701           val ((statement, intro, axioms), thy''') =
   702             thy''
   703             |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred);
   704           val (_, thy'''') =
   705             thy'''
   706             |> Sign.qualified_path true binding
   707             |> Global_Theory.note_thmss ""
   708                  [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.intro_add])]),
   709                   ((Binding.conceal (Binding.name axiomsN), []),
   710                     [(map (Drule.export_without_context o Element.conclude_witness) axioms, [])])]
   711             ||> Sign.restore_naming thy''';
   712         in (SOME statement, SOME intro, axioms, thy'''') end;
   713   in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
   714 
   715 end;
   716 
   717 
   718 local
   719 
   720 fun assumes_to_notes (Assumes asms) axms =
   721       fold_map (fn (a, spec) => fn axs =>
   722           let val (ps, qs) = chop (length spec) axs
   723           in ((a, [(ps, [])]), qs) end) asms axms
   724       |> apfst (curry Notes "")
   725   | assumes_to_notes e axms = (e, axms);
   726 
   727 fun defines_to_notes thy (Defines defs) =
   728       Notes (Thm.definitionK, map (fn (a, (def, _)) =>
   729         (a, [([Assumption.assume (cterm_of thy def)],
   730           [(Attrib.internal o K) Locale.witness_add])])) defs)
   731   | defines_to_notes _ e = e;
   732 
   733 fun gen_add_locale prep_decl
   734     before_exit binding raw_predicate_binding raw_import raw_body thy =
   735   let
   736     val name = Sign.full_name thy binding;
   737     val _ = Locale.defined thy name andalso
   738       error ("Duplicate definition of locale " ^ quote name);
   739 
   740     val ((fixed, deps, body_elems), (parms, ctxt')) =
   741       prep_decl raw_import I raw_body (Proof_Context.init_global thy);
   742     val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems;
   743 
   744     val extraTs =
   745       subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []);
   746     val _ =
   747       if null extraTs then ()
   748       else warning ("Additional type variable(s) in locale specification " ^
   749         quote (Binding.str_of binding) ^ ": " ^
   750           commas (map (Syntax.string_of_typ ctxt' o TFree) (sort_wrt #1 extraTs)));
   751 
   752     val predicate_binding =
   753       if Binding.is_empty raw_predicate_binding then binding
   754       else raw_predicate_binding;
   755     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
   756       define_preds predicate_binding parms text thy;
   757 
   758     val a_satisfy = Element.satisfy_morphism a_axioms;
   759     val b_satisfy = Element.satisfy_morphism b_axioms;
   760 
   761     val params = fixed @
   762       maps (fn Fixes fixes =>
   763         map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems;
   764     val asm = if is_some b_statement then b_statement else a_statement;
   765 
   766     val notes =
   767       if is_some asm then
   768         [("", [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []),
   769           [([Assumption.assume (cterm_of thy' (the asm))],
   770             [(Attrib.internal o K) Locale.witness_add])])])]
   771       else [];
   772 
   773     val notes' = body_elems |>
   774       map (defines_to_notes thy') |>
   775       map (Element.morph_ctxt a_satisfy) |>
   776       (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
   777       fst |>
   778       map (Element.morph_ctxt b_satisfy) |>
   779       map_filter (fn Notes notes => SOME notes | _ => NONE);
   780 
   781     val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps;
   782     val axioms = map Element.conclude_witness b_axioms;
   783 
   784     val loc_ctxt = thy'
   785       |> Locale.register_locale binding (extraTs, params)
   786           (asm, rev defs) (a_intro, b_intro) axioms [] (rev notes) (rev deps')
   787       |> Named_Target.init before_exit name
   788       |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes';
   789 
   790   in (name, loc_ctxt) end;
   791 
   792 in
   793 
   794 val add_locale = gen_add_locale cert_declaration;
   795 val add_locale_cmd = gen_add_locale read_declaration;
   796 
   797 end;
   798 
   799 
   800 (*** Interpretation ***)
   801 
   802 (** Interpretation in theories and proof contexts **)
   803 
   804 local
   805 
   806 fun note_eqns_register deps witss attrss eqns export export' context =
   807   let
   808     fun meta_rewrite context =
   809       map (Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def) o
   810         maps snd;
   811   in
   812     context
   813     |> Element.generic_note_thmss Thm.lemmaK
   814       (attrss ~~ map (fn eqn => [([Morphism.thm (export' $> export) eqn], [])]) eqns)
   815     |-> (fn facts => `(fn context => meta_rewrite context facts))
   816     |-> (fn eqns => fold (fn ((dep, morph), wits) =>
   817       fn context =>
   818         Locale.add_registration (dep, morph $> Element.satisfy_morphism
   819             (map (Element.morph_witness export') wits))
   820           (Element.eq_morphism (Context.theory_of context) eqns |>
   821             Option.map (rpair true))
   822           export context) (deps ~~ witss))
   823   end;
   824 
   825 fun gen_interpretation prep_expr parse_prop prep_attr
   826     expression equations theory =
   827   let
   828     val ((propss, deps, export), expr_ctxt) = Proof_Context.init_global theory
   829       |> prep_expr expression;
   830 
   831     val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
   832     val attrss = map ((apsnd o map) (prep_attr theory) o fst) equations;
   833     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
   834     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
   835 
   836     fun after_qed witss eqns =
   837       (Proof_Context.background_theory o Context.theory_map)
   838         (note_eqns_register deps witss attrss eqns export export');
   839 
   840   in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
   841 
   842 fun gen_interpret prep_expr parse_prop prep_attr
   843     expression equations int state =
   844   let
   845     val _ = Proof.assert_forward_or_chain state;
   846     val ctxt = Proof.context_of state;
   847     val theory = Proof_Context.theory_of ctxt;
   848 
   849     val ((propss, deps, export), expr_ctxt) = prep_expr expression ctxt;
   850 
   851     val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
   852     val attrss = map ((apsnd o map) (prep_attr theory) o fst) equations;
   853     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
   854     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
   855 
   856     fun after_qed witss eqns = (Proof.map_context o Context.proof_map)
   857       (note_eqns_register deps witss attrss eqns export export');
   858   in
   859     state
   860     |> Element.witness_local_proof_eqs after_qed "interpret" propss eqns goal_ctxt int
   861   end;
   862 
   863 in
   864 
   865 fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
   866 fun interpretation_cmd x = gen_interpretation read_goal_expression
   867   Syntax.parse_prop Attrib.intern_src x;
   868 
   869 fun interpret x = gen_interpret cert_goal_expression (K I) (K I) x;
   870 fun interpret_cmd x = gen_interpret read_goal_expression
   871   Syntax.parse_prop Attrib.intern_src x;
   872 
   873 end;
   874 
   875 
   876 (** Interpretation between locales: declaring sublocale relationships **)
   877 
   878 local
   879 
   880 fun note_eqns_dependency target deps witss attrss eqns export export' ctxt =
   881   let
   882     fun meta_rewrite ctxt =
   883       map (Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def) o maps snd;
   884     val facts =
   885       (attrss ~~ map (fn eqn => [([Morphism.thm (export' $> export) eqn], [])]) eqns);
   886     val eqns' = ctxt |> Context.Proof
   887       |> Element.generic_note_thmss Thm.lemmaK facts
   888       |> apsnd Context.the_proof  (* FIXME Context.proof_map_result *)
   889       |-> (fn facts => `(fn ctxt => meta_rewrite ctxt facts))
   890       |> fst;  (* FIXME duplication to add_thmss *)
   891   in
   892     ctxt
   893     |> Locale.add_thmss target Thm.lemmaK facts
   894     |> Proof_Context.background_theory (fold (fn ((dep, morph), wits) =>
   895       fn theory =>
   896         Locale.add_dependency target
   897           (dep, morph $> Element.satisfy_morphism
   898             (map (Element.morph_witness export') wits))
   899           (Element.eq_morphism theory eqns' |> Option.map (rpair true))
   900           export theory) (deps ~~ witss))
   901   end;
   902 
   903 fun gen_sublocale prep_expr intern parse_prop prep_attr
   904     before_exit raw_target expression equations thy =
   905   let
   906     val target = intern thy raw_target;
   907     val target_ctxt = Named_Target.init before_exit target thy;
   908     val ((propss, deps, export), expr_ctxt) = prep_expr expression target_ctxt;
   909 
   910     val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
   911     val attrss = map ((apsnd o map) (prep_attr thy) o fst) equations;
   912     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
   913     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
   914 
   915     fun after_qed witss eqns =
   916       note_eqns_dependency target deps witss attrss eqns export export';
   917 
   918   in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
   919 in
   920 
   921 fun sublocale x = gen_sublocale cert_goal_expression (K I) (K I) (K I) x;
   922 fun sublocale_cmd x =
   923   gen_sublocale read_goal_expression Locale.intern Syntax.parse_prop Attrib.intern_src x;
   924 
   925 end;
   926 
   927 
   928 (** Print the instances that would be activated by an interpretation
   929   of the expression in the current context (clean = false) or in an
   930   empty context (clean = true). **)
   931 
   932 fun print_dependencies ctxt clean expression =
   933   let
   934     val ((_, deps, export), expr_ctxt) = read_goal_expression expression ctxt;
   935   in
   936     Locale.print_dependencies expr_ctxt clean export deps
   937   end;
   938 
   939 end;
   940