src/Pure/Isar/expression.ML
author blanchet
Wed, 04 Mar 2009 11:05:29 +0100
changeset 30242 aea5d7fa7ef5
parent 30240 5b25fee0362c
parent 30227 24d975352879
child 30341 b3ef64cadcad
permissions -rw-r--r--
Merge.
     1 (*  Title:      Pure/Isar/expression.ML
     2     Author:     Clemens Ballarin, TU Muenchen
     3 
     4 Locale expressions.
     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 -> ((binding * typ option * 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 -> ((binding * typ option * 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 -> ((binding * typ option * mixfix) list * (string * morphism) list
    31       * Element.context_i list) * ((string * typ) list * Proof.context)
    32   val add_locale: bstring -> bstring -> expression_i -> Element.context_i list ->
    33     theory -> string * local_theory
    34   val add_locale_cmd: bstring -> bstring -> expression -> Element.context list ->
    35     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: string -> expression_i -> theory -> Proof.state
    43   val sublocale_cmd: string -> expression -> theory -> Proof.state
    44   val interpretation: expression_i -> (Attrib.binding * term) list ->
    45     theory -> Proof.state
    46   val interpretation_cmd: expression -> (Attrib.binding * string) list ->
    47     theory -> Proof.state
    48   val interpret: expression_i -> bool -> Proof.state -> Proof.state
    49   val interpret_cmd: expression -> bool -> Proof.state -> Proof.state
    50 end;
    51 
    52 
    53 structure Expression : EXPRESSION =
    54 struct
    55 
    56 datatype ctxt = datatype Element.ctxt;
    57 
    58 
    59 (*** Expressions ***)
    60 
    61 datatype 'term map =
    62   Positional of 'term option list |
    63   Named of (string * 'term) list;
    64 
    65 type 'term expr = (string * ((string * bool) * 'term map)) list;
    66 
    67 type expression = string expr * (binding * string option * mixfix) list;
    68 type expression_i = term expr * (binding * typ option * mixfix) list;
    69 
    70 
    71 (** Internalise locale names in expr **)
    72 
    73 fun intern thy instances =  map (apfst (Locale.intern thy)) instances;
    74 
    75 
    76 (** Parameters of expression.
    77 
    78    Sanity check of instantiations and extraction of implicit parameters.
    79    The latter only occurs iff strict = false.
    80    Positional instantiations are extended to match full length of parameter list
    81    of instantiated locale. **)
    82 
    83 fun parameters_of thy strict (expr, fixed) =
    84   let
    85     fun reject_dups message xs =
    86       let val dups = duplicates (op =) xs
    87       in
    88         if null dups then () else error (message ^ commas dups)
    89       end;
    90 
    91     fun match_bind (n, b) = (n = Binding.name_of b);
    92     fun parm_eq ((b1, mx1: mixfix), (b2, mx2)) =
    93       (* FIXME: cannot compare bindings for equality, instead check for equal name and syntax *)
    94       Binding.name_of b1 = Binding.name_of b2 andalso
    95         (mx1 = mx2 orelse
    96           error ("Conflicting syntax for parameter " ^ quote (Binding.str_of b1) ^ " in expression"));
    97       
    98     fun params_loc loc =
    99       (Locale.params_of thy loc |> map (fn (p, _, mx) => (p, mx)), loc);
   100     fun params_inst (expr as (loc, (prfx, Positional insts))) =
   101           let
   102             val (ps, loc') = 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 (expr as (loc, (prfx, Named insts))) =
   112           let
   113             val _ = reject_dups "Duplicate instantiation of the following parameter(s): "
   114               (map fst insts);
   115 
   116             val (ps, loc') = params_loc loc;
   117             val ps' = fold (fn (p, _) => fn ps =>
   118               if AList.defined match_bind ps p then AList.delete match_bind p ps
   119               else error (quote p ^" not a parameter of instantiated expression.")) insts ps;
   120           in (ps', (loc', (prfx, Named insts))) end;
   121     fun params_expr is =
   122           let
   123             val (is', ps') = fold_map (fn i => fn ps =>
   124               let
   125                 val (ps', i') = params_inst i;
   126                 val ps'' = distinct parm_eq (ps @ ps');
   127               in (i', ps'') end) is []
   128           in (ps', is') end;
   129 
   130     val (implicit, expr') = params_expr expr;
   131 
   132     val implicit' = map (#1 #> Binding.name_of) implicit;
   133     val fixed' = map (#1 #> Binding.name_of) fixed;
   134     val _ = reject_dups "Duplicate fixed parameter(s): " fixed';
   135     val implicit'' = if strict then []
   136       else let val _ = reject_dups
   137           "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed')
   138         in map (fn (b, mx) => (b, 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, strict), 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 (TypeInfer.paramify_vars o Logic.varifyT) parm_types;
   175     val type_parms = fold Term.add_tvarsT parm_types' [] |> map (Logic.mk_type o TVar);
   176     val arg = type_parms @ map2 TypeInfer.constrain 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 (ProofContext.theory_of ctxt) (instT, inst) $>
   189       Morphism.binding_morphism (Binding.add_prefix strict 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 (ProofContext.set_mode ProofContext.mode_schematic ctxt),
   202     pattern = prep_term (ProofContext.set_mode ProofContext.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 (ProofContext.set_mode ProofContext.mode_schematic ctxt) t, (* FIXME ?? *)
   209       map (prep_term (ProofContext.set_mode ProofContext.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) = (Syntax.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 (ProofContext.set_mode ProofContext.mode_pattern ctxt') pats),
   251           (ctxt', ts))
   252       end
   253     val (cs', (context', _)) = fold_map prep cs
   254       (context, Syntax.check_terms
   255         (ProofContext.set_mode ProofContext.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'), map restore_elem (elems ~~ elem_css'),
   271       concl_cs', ctxt')
   272   end;
   273 
   274 end;
   275 
   276 
   277 (** Prepare locale elements **)
   278 
   279 fun declare_elem prep_vars (Fixes fixes) ctxt =
   280       let val (vars, _) = prep_vars fixes ctxt
   281       in ctxt |> ProofContext.add_fixes_i vars |> snd end
   282   | declare_elem prep_vars (Constrains csts) ctxt =
   283       ctxt |> prep_vars (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) csts) |> snd
   284   | declare_elem _ (Assumes _) ctxt = ctxt
   285   | declare_elem _ (Defines _) ctxt = ctxt
   286   | declare_elem _ (Notes _) ctxt = ctxt;
   287 
   288 (** Finish locale elements **)
   289 
   290 fun closeup _ _ false elem = elem
   291   | closeup ctxt parms true elem =
   292       let
   293         fun close_frees t =
   294           let
   295             val rev_frees =
   296               Term.fold_aterms (fn Free (x, T) =>
   297                 if AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t [];
   298           in Term.list_all_free (rev rev_frees, t) end; (* FIXME use fold Logic.all *)
   299   (* FIXME consider closing in syntactic phase *)
   300 
   301         fun no_binds [] = []
   302           | no_binds _ = error "Illegal term bindings in context element";
   303       in
   304         (case elem of
   305           Assumes asms => Assumes (asms |> map (fn (a, propps) =>
   306             (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps)))
   307         | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) =>
   308             let val ((c, _), t') = LocalDefs.cert_def ctxt (close_frees t)
   309             in ((Binding.map_name (Thm.def_name_optional c) name, atts), (t', no_binds ps)) end))
   310         | e => e)
   311       end;
   312 
   313 fun finish_primitive parms _ (Fixes fixes) = Fixes (map (fn (binding, _, mx) =>
   314       let val x = Binding.name_of binding
   315       in (binding, AList.lookup (op =) parms x, mx) end) fixes)
   316   | finish_primitive _ _ (Constrains _) = Constrains []
   317   | finish_primitive _ close (Assumes asms) = close (Assumes asms)
   318   | finish_primitive _ close (Defines defs) = close (Defines defs)
   319   | finish_primitive _ _ (Notes facts) = Notes facts;
   320 
   321 fun finish_inst ctxt parms do_close (loc, (prfx, inst)) =
   322   let
   323     val thy = ProofContext.theory_of ctxt;
   324     val (parm_names, parm_types) = Locale.params_of thy loc |>
   325       map_split (fn (b, SOME T, _) => (Binding.name_of b, T));
   326     val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst) ctxt;
   327   in (loc, morph) end;
   328 
   329 fun finish_elem ctxt parms do_close elem =
   330   let
   331     val elem' = finish_primitive parms (closeup ctxt parms do_close) elem;
   332   in elem' end
   333   
   334 fun finish ctxt parms do_close insts elems =
   335   let
   336     val deps = map (finish_inst ctxt parms do_close) insts;
   337     val elems' = map (finish_elem ctxt parms do_close) elems;
   338   in (deps, elems') end;
   339 
   340 
   341 (** Process full context statement: instantiations + elements + conclusion **)
   342 
   343 (* Interleave incremental parsing and type inference over entire parsed stretch. *)
   344 
   345 local
   346 
   347 fun prep_full_context_statement parse_typ parse_prop prep_vars_elem prep_inst prep_vars_inst prep_expr
   348     strict do_close raw_import init_body raw_elems raw_concl ctxt1 =
   349   let
   350     val thy = ProofContext.theory_of ctxt1;
   351 
   352     val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
   353 
   354     fun prep_insts (loc, (prfx, inst)) (i, insts, ctxt) =
   355       let
   356         val (parm_names, parm_types) = Locale.params_of thy loc |>
   357           map_split (fn (b, SOME T, _) => (Binding.name_of b, T))
   358             (*FIXME return value of Locale.params_of has strange type*)
   359         val inst' = prep_inst ctxt parm_names inst;
   360         val parm_types' = map (TypeInfer.paramify_vars o
   361           Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) o Logic.varifyT) parm_types;
   362         val inst'' = map2 TypeInfer.constrain parm_types' inst';
   363         val insts' = insts @ [(loc, (prfx, inst''))];
   364         val (insts'', _, _, ctxt' (* FIXME not used *) ) = check_autofix insts' [] [] ctxt;
   365         val inst''' = insts'' |> List.last |> snd |> snd;
   366         val (morph, _) = inst_morph (parm_names, parm_types) (prfx, inst''') ctxt;
   367         val ctxt'' = Locale.activate_declarations thy (loc, morph) ctxt;
   368       in (i+1, insts', ctxt'') end;
   369   
   370     fun prep_elem insts raw_elem (elems, ctxt) =
   371       let
   372         val ctxt' = declare_elem prep_vars_elem raw_elem ctxt;
   373         val elems' = elems @ [parse_elem parse_typ parse_prop ctxt' raw_elem];
   374         val (_, _, _, ctxt'') = check_autofix insts elems' [] ctxt';
   375       in (elems', ctxt') end;
   376 
   377     fun prep_concl raw_concl (insts, elems, ctxt) =
   378       let
   379         val concl = parse_concl parse_prop ctxt raw_concl;
   380       in check_autofix insts elems concl ctxt end;
   381 
   382     val fors = prep_vars_inst fixed ctxt1 |> fst;
   383     val ctxt2 = ctxt1 |> ProofContext.add_fixes_i fors |> snd;
   384     val (_, insts', ctxt3) = fold prep_insts raw_insts (0, [], ctxt2);
   385     val ctxt4 = init_body ctxt3;
   386     val (elems, ctxt5) = fold (prep_elem insts') raw_elems ([], ctxt4);
   387     val (insts, elems', concl, ctxt6) =
   388       prep_concl raw_concl (insts', elems, ctxt5);
   389 
   390     (* Retrieve parameter types *)
   391     val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.name_of o #1) fixes)
   392       | _ => fn ps => ps) (Fixes fors :: elems') [];
   393     val (Ts, ctxt7) = fold_map ProofContext.inferred_param xs ctxt6; 
   394     val parms = xs ~~ Ts;  (* params from expression and elements *)
   395 
   396     val Fixes fors' = finish_primitive parms I (Fixes fors);
   397     val (deps, elems'') = finish ctxt6 parms do_close insts elems';
   398 
   399   in ((fors', deps, elems'', concl), (parms, ctxt7)) end
   400 
   401 in
   402 
   403 fun cert_full_context_statement x =
   404   prep_full_context_statement (K I) (K I) ProofContext.cert_vars
   405   make_inst ProofContext.cert_vars (K I) x;
   406 fun cert_read_full_context_statement x =
   407   prep_full_context_statement Syntax.parse_typ Syntax.parse_prop ProofContext.read_vars
   408   make_inst ProofContext.cert_vars (K I) x;
   409 fun read_full_context_statement x =
   410   prep_full_context_statement Syntax.parse_typ Syntax.parse_prop ProofContext.read_vars
   411   parse_inst ProofContext.read_vars intern x;
   412 
   413 end;
   414 
   415 
   416 (* Context statement: elements + conclusion *)
   417 
   418 local
   419 
   420 fun prep_statement prep activate raw_elems raw_concl context =
   421   let
   422      val ((_, _, elems, concl), _) =
   423        prep true false ([], []) I raw_elems raw_concl context;
   424      val (_, context') = context |>
   425        ProofContext.set_stmt true |>
   426        activate elems;
   427   in (concl, context') end;
   428 
   429 in
   430 
   431 fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x;
   432 fun read_statement x = prep_statement read_full_context_statement Element.activate x;
   433 
   434 end;
   435 
   436 
   437 (* Locale declaration: import + elements *)
   438 
   439 local
   440 
   441 fun prep_declaration prep activate raw_import init_body raw_elems context =
   442   let
   443     val ((fixed, deps, elems, _), (parms, ctxt')) =
   444       prep false true raw_import init_body raw_elems [] context ;
   445     (* Declare parameters and imported facts *)
   446     val context' = context |>
   447       ProofContext.add_fixes_i fixed |> snd |>
   448       fold Locale.activate_local_facts deps;
   449     val (elems', _) = context' |>
   450       ProofContext.set_stmt true |>
   451       activate elems;
   452   in ((fixed, deps, elems'), (parms, ctxt')) end;
   453 
   454 in
   455 
   456 fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x;
   457 fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x;
   458 fun read_declaration x = prep_declaration read_full_context_statement Element.activate x;
   459 
   460 end;
   461 
   462 
   463 (* Locale expression to set up a goal *)
   464 
   465 local
   466 
   467 fun props_of thy (name, morph) =
   468   let
   469     val (asm, defs) = Locale.specification_of thy name;
   470   in
   471     (case asm of NONE => defs | SOME asm => asm :: defs) |> map (Morphism.term morph)
   472   end;
   473 
   474 fun prep_goal_expression prep expression context =
   475   let
   476     val thy = ProofContext.theory_of context;
   477 
   478     val ((fixed, deps, _, _), _) =
   479       prep true true expression I [] [] context;
   480     (* proof obligations *)
   481     val propss = map (props_of thy) deps;
   482 
   483     val goal_ctxt = context |>
   484       ProofContext.add_fixes_i fixed |> snd |>
   485       (fold o fold) Variable.auto_fixes propss;
   486 
   487     val export = Variable.export_morphism goal_ctxt context;
   488     val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export;
   489     val exp_term = TermSubst.zero_var_indexes o Morphism.term export;
   490     val exp_typ = Logic.type_map exp_term;
   491     val export' = Morphism.morphism {binding = I, typ = exp_typ, term = exp_term, fact = exp_fact};
   492   in ((propss, deps, export'), goal_ctxt) end;
   493     
   494 in
   495 
   496 fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x;
   497 fun read_goal_expression x = prep_goal_expression read_full_context_statement x;
   498 
   499 end;
   500 
   501 
   502 (*** Locale declarations ***)
   503 
   504 (* extract specification text *)
   505 
   506 val norm_term = Envir.beta_norm oo Term.subst_atomic;
   507 
   508 fun bind_def ctxt eq (xs, env, eqs) =
   509   let
   510     val _ = LocalDefs.cert_def ctxt eq;
   511     val ((y, T), b) = LocalDefs.abs_def eq;
   512     val b' = norm_term env b;
   513     fun err msg = error (msg ^ ": " ^ quote y);
   514   in
   515     case filter (fn (Free (y', _), _) => y = y' | _ => false) env of
   516       [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) |
   517       dups => if forall (fn (_, b'') => b' aconv b'') dups
   518         then (xs, env, eqs)
   519         else err "Attempt to redefine variable"
   520   end;
   521 
   522 (* text has the following structure:
   523        (((exts, exts'), (ints, ints')), (xs, env, defs))
   524    where
   525      exts: external assumptions (terms in assumes elements)
   526      exts': dito, normalised wrt. env
   527      ints: internal assumptions (terms in assumptions from insts)
   528      ints': dito, normalised wrt. env
   529      xs: the free variables in exts' and ints' and rhss of definitions,
   530        this includes parameters except defined parameters
   531      env: list of term pairs encoding substitutions, where the first term
   532        is a free variable; substitutions represent defines elements and
   533        the rhs is normalised wrt. the previous env
   534      defs: the equations from the defines elements
   535    *)
   536 
   537 fun eval_text _ _ (Fixes _) text = text
   538   | eval_text _ _ (Constrains _) text = text
   539   | eval_text _ is_ext (Assumes asms)
   540         (((exts, exts'), (ints, ints')), (xs, env, defs)) =
   541       let
   542         val ts = maps (map #1 o #2) asms;
   543         val ts' = map (norm_term env) ts;
   544         val spec' =
   545           if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints'))
   546           else ((exts, exts'), (ints @ ts, ints' @ ts'));
   547       in (spec', (fold Term.add_frees ts' xs, env, defs)) end
   548   | eval_text ctxt _ (Defines defs) (spec, binds) =
   549       (spec, fold (bind_def ctxt o #1 o #2) defs binds)
   550   | eval_text _ _ (Notes _) text = text;
   551 
   552 fun eval_inst ctxt (loc, morph) text =
   553   let
   554     val thy = ProofContext.theory_of ctxt;
   555     val (asm, defs) = Locale.specification_of thy loc;
   556     val asm' = Option.map (Morphism.term morph) asm;
   557     val defs' = map (Morphism.term morph) defs;
   558     val text' = text |>
   559       (if is_some asm
   560         then eval_text ctxt false (Assumes [(Attrib.empty_binding, [(the asm', [])])])
   561         else I) |>
   562       (if not (null defs)
   563         then eval_text ctxt false (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs'))
   564         else I)
   565 (* FIXME clone from locale.ML *)
   566   in text' end;
   567 
   568 fun eval_elem ctxt elem text =
   569   let
   570     val text' = eval_text ctxt true elem text;
   571   in text' end;
   572 
   573 fun eval ctxt deps elems =
   574   let
   575     val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [], []));
   576     val ((spec, (_, _, defs))) = fold (eval_elem ctxt) elems text';
   577   in (spec, defs) end;
   578 
   579 (* axiomsN: name of theorem set with destruct rules for locale predicates,
   580      also name suffix of delta predicates and assumptions. *)
   581 
   582 val axiomsN = "axioms";
   583 
   584 local
   585 
   586 (* introN: name of theorems for introduction rules of locale and
   587      delta predicates *)
   588 
   589 val introN = "intro";
   590 
   591 fun atomize_spec thy ts =
   592   let
   593     val t = Logic.mk_conjunction_balanced ts;
   594     val body = ObjectLogic.atomize_term thy t;
   595     val bodyT = Term.fastype_of body;
   596   in
   597     if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of thy t))
   598     else (body, bodyT, ObjectLogic.atomize (Thm.cterm_of thy t))
   599   end;
   600 
   601 (* achieve plain syntax for locale predicates (without "PROP") *)
   602 
   603 fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
   604   if length args = n then
   605     Syntax.const "_aprop" $
   606       Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
   607   else raise Match);
   608 
   609 (* define one predicate including its intro rule and axioms
   610    - bname: predicate name
   611    - parms: locale parameters
   612    - defs: thms representing substitutions from defines elements
   613    - ts: terms representing locale assumptions (not normalised wrt. defs)
   614    - norm_ts: terms representing locale assumptions (normalised wrt. defs)
   615    - thy: the theory
   616 *)
   617 
   618 fun def_pred bname parms defs ts norm_ts thy =
   619   let
   620     val name = Sign.full_bname thy bname;
   621 
   622     val (body, bodyT, body_eq) = atomize_spec thy norm_ts;
   623     val env = Term.add_free_names body [];
   624     val xs = filter (member (op =) env o #1) parms;
   625     val Ts = map #2 xs;
   626     val extraTs =
   627       (Term.add_tfrees body [] \\ fold Term.add_tfreesT Ts [])
   628       |> Library.sort_wrt #1 |> map TFree;
   629     val predT = map Term.itselfT extraTs ---> Ts ---> bodyT;
   630 
   631     val args = map Logic.mk_type extraTs @ map Free xs;
   632     val head = Term.list_comb (Const (name, predT), args);
   633     val statement = ObjectLogic.ensure_propT thy head;
   634 
   635     val ([pred_def], defs_thy) =
   636       thy
   637       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
   638       |> Sign.declare_const [] ((Binding.name bname, predT), NoSyn) |> snd
   639       |> PureThy.add_defs false
   640         [((Binding.name (Thm.def_name bname), Logic.mk_equals (head, body)), [Thm.kind_internal])];
   641     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
   642 
   643     val cert = Thm.cterm_of defs_thy;
   644 
   645     val intro = Goal.prove_global defs_thy [] norm_ts statement (fn _ =>
   646       MetaSimplifier.rewrite_goals_tac [pred_def] THEN
   647       Tactic.compose_tac (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN
   648       Tactic.compose_tac (false,
   649         Conjunction.intr_balanced (map (Thm.assume o cert) norm_ts), 0) 1);
   650 
   651     val conjuncts =
   652       (Drule.equal_elim_rule2 OF [body_eq,
   653         MetaSimplifier.rewrite_rule [pred_def] (Thm.assume (cert statement))])
   654       |> Conjunction.elim_balanced (length ts);
   655     val axioms = ts ~~ conjuncts |> map (fn (t, ax) =>
   656       Element.prove_witness defs_ctxt t
   657        (MetaSimplifier.rewrite_goals_tac defs THEN
   658         Tactic.compose_tac (false, ax, 0) 1));
   659   in ((statement, intro, axioms), defs_thy) end;
   660 
   661 in
   662 
   663 (* CB: main predicate definition function *)
   664 
   665 fun define_preds pname parms (((exts, exts'), (ints, ints')), defs) thy =
   666   let
   667     val defs' = map (cterm_of thy #> Assumption.assume #> Drule.gen_all #> Drule.abs_def) defs;
   668 
   669     val (a_pred, a_intro, a_axioms, thy'') =
   670       if null exts then (NONE, NONE, [], thy)
   671       else
   672         let
   673           val aname = if null ints then pname else pname ^ "_" ^ axiomsN;
   674           val ((statement, intro, axioms), thy') =
   675             thy
   676             |> def_pred aname parms defs' exts exts';
   677           val (_, thy'') =
   678             thy'
   679             |> Sign.add_path aname
   680             |> Sign.no_base_names
   681             |> PureThy.note_thmss Thm.internalK
   682               [((Binding.name introN, []), [([intro], [Locale.unfold_attrib])])]
   683             ||> Sign.restore_naming thy';
   684           in (SOME statement, SOME intro, axioms, thy'') end;
   685     val (b_pred, b_intro, b_axioms, thy'''') =
   686       if null ints then (NONE, NONE, [], thy'')
   687       else
   688         let
   689           val ((statement, intro, axioms), thy''') =
   690             thy''
   691             |> def_pred pname parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred);
   692           val (_, thy'''') =
   693             thy'''
   694             |> Sign.add_path pname
   695             |> Sign.no_base_names
   696             |> PureThy.note_thmss Thm.internalK
   697                  [((Binding.name introN, []), [([intro], [Locale.intro_attrib])]),
   698                   ((Binding.name axiomsN, []),
   699                     [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
   700             ||> Sign.restore_naming thy''';
   701         in (SOME statement, SOME intro, axioms, thy'''') end;
   702   in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
   703 
   704 end;
   705 
   706 
   707 local
   708 
   709 fun assumes_to_notes (Assumes asms) axms =
   710       fold_map (fn (a, spec) => fn axs =>
   711           let val (ps, qs) = chop (length spec) axs
   712           in ((a, [(ps, [])]), qs) end) asms axms
   713       |> apfst (curry Notes Thm.assumptionK)
   714   | assumes_to_notes e axms = (e, axms);
   715 
   716 fun defines_to_notes thy (Defines defs) =
   717       Notes (Thm.definitionK, map (fn (a, (def, _)) =>
   718         (a, [([Assumption.assume (cterm_of thy def)],
   719           [(Attrib.internal o K) Locale.witness_attrib])])) defs)
   720   | defines_to_notes _ e = e;
   721 
   722 fun gen_add_locale prep_decl
   723     bname raw_predicate_bname raw_import raw_body thy =
   724   let
   725     val name = Sign.full_bname thy bname;
   726     val _ = Locale.defined thy name andalso
   727       error ("Duplicate definition of locale " ^ quote name);
   728 
   729     val ((fixed, deps, body_elems), (parms, ctxt')) =
   730       prep_decl raw_import I raw_body (ProofContext.init thy);
   731     val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems;
   732 
   733     val predicate_bname = if raw_predicate_bname = "" then bname
   734       else raw_predicate_bname;
   735     val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') =
   736       define_preds predicate_bname parms text thy;
   737 
   738     val extraTs = fold Term.add_tfrees exts' [] \\ fold Term.add_tfreesT (map snd parms) [];
   739     val _ = if null extraTs then ()
   740       else warning ("Additional type variable(s) in locale specification " ^ quote bname);
   741 
   742     val a_satisfy = Element.satisfy_morphism a_axioms;
   743     val b_satisfy = Element.satisfy_morphism b_axioms;
   744 
   745     val params = fixed @
   746       (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
   747     val asm = if is_some b_statement then b_statement else a_statement;
   748 
   749     val notes =
   750         if is_some asm
   751         then [(Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
   752           [([Assumption.assume (cterm_of thy' (the asm))],
   753             [(Attrib.internal o K) Locale.witness_attrib])])])]
   754         else [];
   755 
   756     val notes' = body_elems |>
   757       map (defines_to_notes thy') |>
   758       map (Element.morph_ctxt a_satisfy) |>
   759       (fn elems => fold_map assumes_to_notes elems (map Element.conclude_witness a_axioms)) |>
   760       fst |>
   761       map (Element.morph_ctxt b_satisfy) |>
   762       map_filter (fn Notes notes => SOME notes | _ => NONE);
   763 
   764     val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps;
   765     val axioms = map Element.conclude_witness b_axioms;
   766 
   767     val loc_ctxt = thy'
   768       |> Locale.register_locale bname (extraTs, params)
   769           (asm, rev defs) (a_intro, b_intro) axioms ([], []) 
   770           (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev)
   771       |> TheoryTarget.init (SOME name)
   772       |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes';
   773 
   774   in (name, loc_ctxt) end;
   775 
   776 in
   777 
   778 val add_locale = gen_add_locale cert_declaration;
   779 val add_locale_cmd = gen_add_locale read_declaration;
   780 
   781 end;
   782 
   783 
   784 (*** Interpretation ***)
   785 
   786 (** Interpretation between locales: declaring sublocale relationships **)
   787 
   788 local
   789 
   790 fun gen_sublocale prep_expr intern raw_target expression thy =
   791   let
   792     val target = intern thy raw_target;
   793     val target_ctxt = Locale.init target thy;
   794 
   795     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
   796     
   797     fun after_qed witss = ProofContext.theory (
   798       fold2 (fn (name, morph) => fn wits => Locale.add_dependency target
   799         (name, morph $> Element.satisfy_morphism wits $> export)) deps witss #>
   800       (fn thy => fold_rev Locale.activate_global_facts
   801           (Locale.get_registrations thy) thy));
   802   in Element.witness_proof after_qed propss goal_ctxt end;
   803 
   804 in
   805 
   806 fun sublocale x = gen_sublocale cert_goal_expression (K I) x;
   807 fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern x;
   808 
   809 end;
   810 
   811 
   812 (** Interpretation in theories **)
   813 
   814 local
   815 
   816 fun gen_interpretation prep_expr parse_prop prep_attr
   817     expression equations theory =
   818   let
   819     val ((propss, regs, export), expr_ctxt) = ProofContext.init theory
   820       |> prep_expr expression;
   821 
   822     val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt;
   823     val eq_attns = map ((apsnd o map) (prep_attr theory) o fst) equations;
   824     val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
   825     val export' = Variable.export_morphism goal_ctxt expr_ctxt;
   826 
   827     fun store_reg ((name, morph), wits) thy =
   828       let
   829         val wits' = map (Element.morph_witness export') wits;
   830         val morph' = morph $> Element.satisfy_morphism wits';
   831       in
   832         thy
   833         |> Locale.add_registration (name, (morph', export))
   834         |> pair (name, morph')
   835       end;
   836 
   837     fun store_eqns_activate regs [] thy =
   838           thy
   839           |> fold (fn (name, morph) =>
   840                Locale.activate_global_facts (name, morph $> export)) regs
   841       | store_eqns_activate regs eqs thy =
   842           let
   843             val eqs' = eqs |> map (Morphism.thm (export' $> export) #>
   844               LocalDefs.meta_rewrite_rule (ProofContext.init thy) #>
   845               Drule.abs_def);
   846             val eq_morph = Element.eq_morphism thy eqs';
   847             val eq_attns' = map ((apsnd o map) (Attrib.attribute_i thy)) eq_attns;
   848           in
   849             thy
   850             |> fold (fn (name, morph) =>
   851                 Locale.amend_registration eq_morph (name, morph) #>
   852                 Locale.activate_global_facts (name, morph $> eq_morph $> export)) regs
   853             |> PureThy.note_thmss Thm.lemmaK (eq_attns' ~~ map (fn eq => [([eq], [])]) eqs')
   854             |> snd
   855           end;
   856 
   857     fun after_qed wits eqs = ProofContext.theory (fold_map store_reg (regs ~~ wits)
   858         #-> (fn regs => store_eqns_activate regs eqs));
   859 
   860   in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
   861 
   862 in
   863 
   864 fun interpretation x = gen_interpretation cert_goal_expression (K I) (K I) x;
   865 fun interpretation_cmd x = gen_interpretation read_goal_expression
   866   Syntax.parse_prop Attrib.intern_src x;
   867 
   868 end;
   869 
   870 
   871 (** Interpretation in proof contexts **)
   872 
   873 local
   874 
   875 fun gen_interpret prep_expr
   876     expression int state =
   877   let
   878     val _ = Proof.assert_forward_or_chain state;
   879     val ctxt = Proof.context_of state;
   880 
   881     val ((propss, regs, export), goal_ctxt) = prep_expr expression ctxt;
   882     
   883     fun store_reg (name, morph) thms =
   884       let
   885         val morph' = morph $> Element.satisfy_morphism thms $> export;
   886       in Locale.activate_local_facts (name, morph') end;
   887 
   888     fun after_qed wits =
   889       Proof.map_context (fold2 store_reg regs wits);
   890   in
   891     state
   892     |> Element.witness_local_proof after_qed "interpret" propss goal_ctxt int
   893   end;
   894 
   895 in
   896 
   897 fun interpret x = gen_interpret cert_goal_expression x;
   898 fun interpret_cmd x = gen_interpret read_goal_expression x;
   899 
   900 end;
   901 
   902 end;
   903