src/Pure/Isar/proof_context.ML
author wenzelm
Tue, 14 Jul 2009 12:10:44 +0200
changeset 32003 befec6450fd6
parent 30817 e96498265a05
child 32043 a6a6e8031c14
permissions -rw-r--r--
tuned prepare_patternT: Term.exists_subtype;
     1 (*  Title:      Pure/Isar/proof_context.ML
     2     Author:     Markus Wenzel, TU Muenchen
     3 
     4 The key concept of Isar proof contexts: elevates primitive local
     5 reasoning Gamma |- phi to a structured concept, with generic context
     6 elements.  See also structure Variable and Assumption.
     7 *)
     8 
     9 signature PROOF_CONTEXT =
    10 sig
    11   val theory_of: Proof.context -> theory
    12   val init: theory -> Proof.context
    13   type mode
    14   val mode_default: mode
    15   val mode_stmt: mode
    16   val mode_pattern: mode
    17   val mode_schematic: mode
    18   val mode_abbrev: mode
    19   val set_mode: mode -> Proof.context -> Proof.context
    20   val get_mode: Proof.context -> mode
    21   val restore_mode: Proof.context -> Proof.context -> Proof.context
    22   val abbrev_mode: Proof.context -> bool
    23   val set_stmt: bool -> Proof.context -> Proof.context
    24   val naming_of: Proof.context -> NameSpace.naming
    25   val full_name: Proof.context -> binding -> string
    26   val consts_of: Proof.context -> Consts.T
    27   val const_syntax_name: Proof.context -> string -> string
    28   val the_const_constraint: Proof.context -> string -> typ
    29   val mk_const: Proof.context -> string * typ list -> term
    30   val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
    31   val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
    32   val facts_of: Proof.context -> Facts.T
    33   val transfer_syntax: theory -> Proof.context -> Proof.context
    34   val transfer: theory -> Proof.context -> Proof.context
    35   val theory: (theory -> theory) -> Proof.context -> Proof.context
    36   val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
    37   val extern_fact: Proof.context -> string -> xstring
    38   val pretty_term_abbrev: Proof.context -> term -> Pretty.T
    39   val pretty_thm_aux: Proof.context -> bool -> thm -> Pretty.T
    40   val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T
    41   val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T
    42   val pretty_thm: Proof.context -> thm -> Pretty.T
    43   val pretty_thms: Proof.context -> thm list -> Pretty.T
    44   val pretty_fact: Proof.context -> string * thm list -> Pretty.T
    45   val string_of_thm: Proof.context -> thm -> string
    46   val read_typ: Proof.context -> string -> typ
    47   val read_typ_syntax: Proof.context -> string -> typ
    48   val read_typ_abbrev: Proof.context -> string -> typ
    49   val cert_typ: Proof.context -> typ -> typ
    50   val cert_typ_syntax: Proof.context -> typ -> typ
    51   val cert_typ_abbrev: Proof.context -> typ -> typ
    52   val get_skolem: Proof.context -> string -> string
    53   val revert_skolem: Proof.context -> string -> string
    54   val infer_type: Proof.context -> string -> typ
    55   val inferred_param: string -> Proof.context -> typ * Proof.context
    56   val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
    57   val read_tyname: Proof.context -> string -> typ
    58   val read_const_proper: Proof.context -> string -> term
    59   val read_const: Proof.context -> string -> term
    60   val allow_dummies: Proof.context -> Proof.context
    61   val decode_term: Proof.context -> term -> term
    62   val standard_infer_types: Proof.context -> term list -> term list
    63   val read_term_pattern: Proof.context -> string -> term
    64   val read_term_schematic: Proof.context -> string -> term
    65   val read_term_abbrev: Proof.context -> string -> term
    66   val expand_abbrevs: Proof.context -> term -> term
    67   val cert_term: Proof.context -> term -> term
    68   val cert_prop: Proof.context -> term -> term
    69   val goal_export: Proof.context -> Proof.context -> thm list -> thm list
    70   val export: Proof.context -> Proof.context -> thm list -> thm list
    71   val export_morphism: Proof.context -> Proof.context -> morphism
    72   val norm_export_morphism: Proof.context -> Proof.context -> morphism
    73   val bind_terms: (indexname * term option) list -> Proof.context -> Proof.context
    74   val auto_bind_goal: term list -> Proof.context -> Proof.context
    75   val auto_bind_facts: term list -> Proof.context -> Proof.context
    76   val match_bind: bool -> (string list * string) list -> Proof.context -> term list * Proof.context
    77   val match_bind_i: bool -> (term list * term) list -> Proof.context -> term list * Proof.context
    78   val read_propp: Proof.context * (string * string list) list list
    79     -> Proof.context * (term * term list) list list
    80   val cert_propp: Proof.context * (term * term list) list list
    81     -> Proof.context * (term * term list) list list
    82   val read_propp_schematic: Proof.context * (string * string list) list list
    83     -> Proof.context * (term * term list) list list
    84   val cert_propp_schematic: Proof.context * (term * term list) list list
    85     -> Proof.context * (term * term list) list list
    86   val bind_propp: Proof.context * (string * string list) list list
    87     -> Proof.context * (term list list * (Proof.context -> Proof.context))
    88   val bind_propp_i: Proof.context * (term * term list) list list
    89     -> Proof.context * (term list list * (Proof.context -> Proof.context))
    90   val bind_propp_schematic: Proof.context * (string * string list) list list
    91     -> Proof.context * (term list list * (Proof.context -> Proof.context))
    92   val bind_propp_schematic_i: Proof.context * (term * term list) list list
    93     -> Proof.context * (term list list * (Proof.context -> Proof.context))
    94   val fact_tac: thm list -> int -> tactic
    95   val some_fact_tac: Proof.context -> int -> tactic
    96   val get_fact: Proof.context -> Facts.ref -> thm list
    97   val get_fact_single: Proof.context -> Facts.ref -> thm
    98   val get_thms: Proof.context -> xstring -> thm list
    99   val get_thm: Proof.context -> xstring -> thm
   100   val add_path: string -> Proof.context -> Proof.context
   101   val mandatory_path: string -> Proof.context -> Proof.context
   102   val restore_naming: Proof.context -> Proof.context -> Proof.context
   103   val reset_naming: Proof.context -> Proof.context
   104   val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list ->
   105     Proof.context -> (string * thm list) list * Proof.context
   106   val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context
   107   val read_vars: (binding * string option * mixfix) list -> Proof.context ->
   108     (binding * typ option * mixfix) list * Proof.context
   109   val cert_vars: (binding * typ option * mixfix) list -> Proof.context ->
   110     (binding * typ option * mixfix) list * Proof.context
   111   val add_fixes: (binding * typ option * mixfix) list -> Proof.context ->
   112     string list * Proof.context
   113   val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a)
   114   val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context
   115   val add_assms: Assumption.export ->
   116     (Thm.binding * (string * string list) list) list ->
   117     Proof.context -> (string * thm list) list * Proof.context
   118   val add_assms_i: Assumption.export ->
   119     (Thm.binding * (term * term list) list) list ->
   120     Proof.context -> (string * thm list) list * Proof.context
   121   val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context
   122   val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context
   123   val get_case: Proof.context -> string -> string option list -> RuleCases.T
   124   val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context
   125   val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->
   126     Context.generic -> Context.generic
   127   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
   128   val add_abbrev: string -> Properties.T ->
   129     binding * term -> Proof.context -> (term * term) * Proof.context
   130   val revert_abbrev: string -> string -> Proof.context -> Proof.context
   131   val verbose: bool ref
   132   val setmp_verbose: ('a -> 'b) -> 'a -> 'b
   133   val print_syntax: Proof.context -> unit
   134   val print_abbrevs: Proof.context -> unit
   135   val print_binds: Proof.context -> unit
   136   val print_lthms: Proof.context -> unit
   137   val print_cases: Proof.context -> unit
   138   val debug: bool ref
   139   val prems_limit: int ref
   140   val pretty_ctxt: Proof.context -> Pretty.T list
   141   val pretty_context: Proof.context -> Pretty.T list
   142   val query_type: Proof.context -> string -> Properties.T
   143   val query_const: Proof.context -> string -> Properties.T
   144   val query_class: Proof.context -> string -> Properties.T
   145 end;
   146 
   147 structure ProofContext: PROOF_CONTEXT =
   148 struct
   149 
   150 val theory_of = Context.theory_of_proof;
   151 val init = Context.init_proof;
   152 
   153 
   154 (** inner syntax mode **)
   155 
   156 datatype mode =
   157   Mode of
   158    {stmt: bool,                (*inner statement mode*)
   159     pattern: bool,             (*pattern binding schematic variables*)
   160     schematic: bool,           (*term referencing loose schematic variables*)
   161     abbrev: bool};             (*abbrev mode -- no normalization*)
   162 
   163 fun make_mode (stmt, pattern, schematic, abbrev) =
   164   Mode {stmt = stmt, pattern = pattern, schematic = schematic, abbrev = abbrev};
   165 
   166 val mode_default   = make_mode (false, false, false, false);
   167 val mode_stmt      = make_mode (true, false, false, false);
   168 val mode_pattern   = make_mode (false, true, false, false);
   169 val mode_schematic = make_mode (false, false, true, false);
   170 val mode_abbrev    = make_mode (false, false, false, true);
   171 
   172 
   173 
   174 (** Isar proof context information **)
   175 
   176 datatype ctxt =
   177   Ctxt of
   178    {mode: mode,                                       (*inner syntax mode*)
   179     naming: NameSpace.naming,                         (*local naming conventions*)
   180     syntax: LocalSyntax.T,                            (*local syntax*)
   181     consts: Consts.T * Consts.T,                      (*local/global consts*)
   182     facts: Facts.T,                                   (*local facts*)
   183     cases: (string * (RuleCases.T * bool)) list};     (*named case contexts*)
   184 
   185 fun make_ctxt (mode, naming, syntax, consts, facts, cases) =
   186   Ctxt {mode = mode, naming = naming, syntax = syntax,
   187     consts = consts, facts = facts, cases = cases};
   188 
   189 val local_naming = NameSpace.default_naming |> NameSpace.add_path "local";
   190 
   191 structure ContextData = ProofDataFun
   192 (
   193   type T = ctxt;
   194   fun init thy =
   195     make_ctxt (mode_default, local_naming, LocalSyntax.init thy,
   196       (Sign.consts_of thy, Sign.consts_of thy), Facts.empty, []);
   197 );
   198 
   199 fun rep_context ctxt = ContextData.get ctxt |> (fn Ctxt args => args);
   200 
   201 fun map_context f =
   202   ContextData.map (fn Ctxt {mode, naming, syntax, consts, facts, cases} =>
   203     make_ctxt (f (mode, naming, syntax, consts, facts, cases)));
   204 
   205 fun set_mode mode = map_context (fn (_, naming, syntax, consts, facts, cases) =>
   206   (mode, naming, syntax, consts, facts, cases));
   207 
   208 fun map_mode f =
   209   map_context (fn (Mode {stmt, pattern, schematic, abbrev}, naming, syntax, consts, facts, cases) =>
   210     (make_mode (f (stmt, pattern, schematic, abbrev)), naming, syntax, consts, facts, cases));
   211 
   212 fun map_naming f =
   213   map_context (fn (mode, naming, syntax, consts, facts, cases) =>
   214     (mode, f naming, syntax, consts, facts, cases));
   215 
   216 fun map_syntax f =
   217   map_context (fn (mode, naming, syntax, consts, facts, cases) =>
   218     (mode, naming, f syntax, consts, facts, cases));
   219 
   220 fun map_consts f =
   221   map_context (fn (mode, naming, syntax, consts, facts, cases) =>
   222     (mode, naming, syntax, f consts, facts, cases));
   223 
   224 fun map_facts f =
   225   map_context (fn (mode, naming, syntax, consts, facts, cases) =>
   226     (mode, naming, syntax, consts, f facts, cases));
   227 
   228 fun map_cases f =
   229   map_context (fn (mode, naming, syntax, consts, facts, cases) =>
   230     (mode, naming, syntax, consts, facts, f cases));
   231 
   232 val get_mode = #mode o rep_context;
   233 val restore_mode = set_mode o get_mode;
   234 val abbrev_mode = get_mode #> (fn Mode {abbrev, ...} => abbrev);
   235 
   236 fun set_stmt stmt =
   237   map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev));
   238 
   239 val naming_of = #naming o rep_context;
   240 val full_name = NameSpace.full_name o naming_of;
   241 
   242 val syntax_of = #syntax o rep_context;
   243 val syn_of = LocalSyntax.syn_of o syntax_of;
   244 val set_syntax_mode = map_syntax o LocalSyntax.set_mode;
   245 val restore_syntax_mode = map_syntax o LocalSyntax.restore_mode o syntax_of;
   246 
   247 val consts_of = #1 o #consts o rep_context;
   248 val const_syntax_name = Consts.syntax_name o consts_of;
   249 val the_const_constraint = Consts.the_constraint o consts_of;
   250 
   251 fun mk_const ctxt (c, Ts) = Const (c, Consts.instance (consts_of ctxt) (c, Ts));
   252 
   253 val facts_of = #facts o rep_context;
   254 val cases_of = #cases o rep_context;
   255 
   256 
   257 (* theory transfer *)
   258 
   259 fun transfer_syntax thy =
   260   map_syntax (LocalSyntax.rebuild thy) #>
   261   map_consts (fn consts as (local_consts, global_consts) =>
   262     let val thy_consts = Sign.consts_of thy in
   263       if Consts.eq_consts (thy_consts, global_consts) then consts
   264       else (Consts.merge (local_consts, thy_consts), thy_consts)
   265     end);
   266 
   267 fun transfer thy = Context.transfer_proof thy #> transfer_syntax thy;
   268 
   269 fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
   270 
   271 fun theory_result f ctxt =
   272   let val (res, thy') = f (theory_of ctxt)
   273   in (res, ctxt |> transfer thy') end;
   274 
   275 
   276 
   277 (** pretty printing **)
   278 
   279 (* extern *)
   280 
   281 fun extern_fact ctxt name =
   282   let
   283     val local_facts = facts_of ctxt;
   284     val global_facts = PureThy.facts_of (theory_of ctxt);
   285   in
   286     if is_some (Facts.lookup (Context.Proof ctxt) local_facts name)
   287     then Facts.extern local_facts name
   288     else Facts.extern global_facts name
   289   end;
   290 
   291 
   292 (* pretty *)
   293 
   294 fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt);
   295 
   296 fun pretty_thm_aux ctxt show_status th =
   297   let
   298     val flags = {quote = false, show_hyps = true, show_status = show_status};
   299     val asms = map Thm.term_of (Assumption.all_assms_of ctxt);
   300   in Display.pretty_thm_aux (Syntax.pp ctxt) flags asms th end;
   301 
   302 fun pretty_thms_aux ctxt flag [th] = pretty_thm_aux ctxt flag th
   303   | pretty_thms_aux ctxt flag ths =
   304       Pretty.blk (0, Pretty.fbreaks (map (pretty_thm_aux ctxt flag) ths));
   305 
   306 fun pretty_fact_name ctxt a = Pretty.block
   307   [Pretty.markup (Markup.fact a) [Pretty.str (extern_fact ctxt a)], Pretty.str ":"];
   308 
   309 fun pretty_fact_aux ctxt flag ("", ths) = pretty_thms_aux ctxt flag ths
   310   | pretty_fact_aux ctxt flag (a, [th]) = Pretty.block
   311       [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm_aux ctxt flag th]
   312   | pretty_fact_aux ctxt flag (a, ths) = Pretty.block
   313       (Pretty.fbreaks (pretty_fact_name ctxt a :: map (pretty_thm_aux ctxt flag) ths));
   314 
   315 fun pretty_thm ctxt = pretty_thm_aux ctxt true;
   316 fun pretty_thms ctxt = pretty_thms_aux ctxt true;
   317 fun pretty_fact ctxt = pretty_fact_aux ctxt true;
   318 
   319 val string_of_thm = Pretty.string_of oo pretty_thm;
   320 
   321 
   322 
   323 (** prepare types **)
   324 
   325 (* read_typ *)
   326 
   327 fun read_typ_mode mode ctxt s =
   328   Syntax.read_typ (Type.set_mode mode ctxt) s;
   329 
   330 val read_typ        = read_typ_mode Type.mode_default;
   331 val read_typ_syntax = read_typ_mode Type.mode_syntax;
   332 val read_typ_abbrev = read_typ_mode Type.mode_abbrev;
   333 
   334 
   335 (* cert_typ *)
   336 
   337 fun cert_typ_mode mode ctxt T =
   338   Sign.certify_typ_mode mode (theory_of ctxt) T
   339     handle TYPE (msg, _, _) => error msg;
   340 
   341 val cert_typ        = cert_typ_mode Type.mode_default;
   342 val cert_typ_syntax = cert_typ_mode Type.mode_syntax;
   343 val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev;
   344 
   345 
   346 
   347 (** prepare variables **)
   348 
   349 (* internalize Skolem constants *)
   350 
   351 val lookup_skolem = AList.lookup (op =) o Variable.fixes_of;
   352 fun get_skolem ctxt x = the_default x (lookup_skolem ctxt x);
   353 
   354 fun no_skolem internal x =
   355   if can Name.dest_skolem x then
   356     error ("Illegal reference to internal Skolem constant: " ^ quote x)
   357   else if not internal andalso can Name.dest_internal x then
   358     error ("Illegal reference to internal variable: " ^ quote x)
   359   else x;
   360 
   361 
   362 (* revert Skolem constants -- if possible *)
   363 
   364 fun revert_skolem ctxt x =
   365   (case find_first (fn (_, y) => y = x) (Variable.fixes_of ctxt) of
   366     SOME (x', _) => if lookup_skolem ctxt x' = SOME x then x' else x
   367   | NONE => x);
   368 
   369 
   370 (* default token translations *)
   371 
   372 local
   373 
   374 fun free_or_skolem ctxt x =
   375   (if can Name.dest_skolem x then Pretty.mark Markup.skolem (Pretty.str (revert_skolem ctxt x))
   376    else Pretty.mark Markup.free (Pretty.str x))
   377   |> Pretty.mark
   378     (if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt then Markup.fixed x
   379      else Markup.hilite);
   380 
   381 fun var_or_skolem _ s =
   382   (case Lexicon.read_variable s of
   383     SOME (x, i) =>
   384       (case try Name.dest_skolem x of
   385         NONE => Pretty.mark Markup.var (Pretty.str s)
   386       | SOME x' => Pretty.mark Markup.skolem
   387           (Pretty.str (setmp show_question_marks true Term.string_of_vname (x', i))))
   388   | NONE => Pretty.mark Markup.var (Pretty.str s));
   389 
   390 fun class_markup _ c =    (* FIXME authentic name *)
   391   Pretty.mark (Markup.tclassN, []) (Pretty.str c);
   392 
   393 fun plain_markup m _ s = Pretty.mark m (Pretty.str s);
   394 
   395 val token_trans =
   396  Syntax.tokentrans_mode ""
   397   [("class", class_markup),
   398    ("tfree", plain_markup Markup.tfree),
   399    ("tvar", plain_markup Markup.tvar),
   400    ("free", free_or_skolem),
   401    ("bound", plain_markup Markup.bound),
   402    ("var", var_or_skolem),
   403    ("numeral", plain_markup Markup.numeral),
   404    ("inner_string", plain_markup Markup.inner_string)];
   405 
   406 in val _ = Context.>> (Context.map_theory (Sign.add_tokentrfuns token_trans)) end;
   407 
   408 
   409 
   410 (** prepare terms and propositions **)
   411 
   412 (* inferred types of parameters *)
   413 
   414 fun infer_type ctxt x =
   415   Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt))
   416     (Free (x, dummyT)));
   417 
   418 fun inferred_param x ctxt =
   419   let val T = infer_type ctxt x
   420   in (T, ctxt |> Variable.declare_term (Free (x, T))) end;
   421 
   422 fun inferred_fixes ctxt =
   423   let
   424     val xs = rev (map #2 (Variable.fixes_of ctxt));
   425     val (Ts, ctxt') = fold_map inferred_param xs ctxt;
   426   in (xs ~~ Ts, ctxt') end;
   427 
   428 
   429 (* type and constant names *)
   430 
   431 local
   432 
   433 val token_content = Syntax.read_token #>> Symbol_Pos.content;
   434 
   435 fun prep_const_proper ctxt (c, pos) =
   436   let val t as (Const (d, _)) =
   437     (case Variable.lookup_const ctxt c of
   438       SOME d => Const (d, Consts.type_scheme (consts_of ctxt) d handle TYPE (msg, _, _) => error msg)
   439     | NONE => Consts.read_const (consts_of ctxt) c)
   440   in Position.report (Markup.const d) pos; t end;
   441 
   442 in
   443 
   444 fun read_tyname ctxt str =
   445   let
   446     val thy = theory_of ctxt;
   447     val (c, pos) = token_content str;
   448   in
   449     if Syntax.is_tid c then
   450      (Position.report Markup.tfree pos;
   451       TFree (c, the_default (Sign.defaultS thy) (Variable.def_sort ctxt (c, ~1))))
   452     else
   453       let
   454         val d = Sign.intern_type thy c;
   455         val _ = Position.report (Markup.tycon d) pos;
   456       in Type (d, replicate (Sign.arity_number thy d) dummyT) end
   457   end;
   458 
   459 fun read_const_proper ctxt = prep_const_proper ctxt o token_content;
   460 
   461 fun read_const ctxt str =
   462   let val (c, pos) = token_content str in
   463     (case (lookup_skolem ctxt c, Variable.is_const ctxt c) of
   464       (SOME x, false) =>
   465         (Position.report (Markup.name x
   466             (if can Name.dest_skolem x then Markup.skolem else Markup.free)) pos;
   467           Free (x, infer_type ctxt x))
   468     | _ => prep_const_proper ctxt (c, pos))
   469   end;
   470 
   471 end;
   472 
   473 
   474 (* read_term *)
   475 
   476 fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt);
   477 
   478 val read_term_pattern   = read_term_mode mode_pattern;
   479 val read_term_schematic = read_term_mode mode_schematic;
   480 val read_term_abbrev    = read_term_mode mode_abbrev;
   481 
   482 
   483 (* local abbreviations *)
   484 
   485 val tsig_of = Sign.tsig_of o ProofContext.theory_of;
   486 
   487 local
   488 
   489 fun certify_consts ctxt = Consts.certify (Syntax.pp ctxt) (tsig_of ctxt)
   490   (not (abbrev_mode ctxt)) (consts_of ctxt);
   491 
   492 fun reject_schematic (Var (xi, _)) =
   493       error ("Unbound schematic variable: " ^ Term.string_of_vname xi)
   494   | reject_schematic (Abs (_, _, t)) = reject_schematic t
   495   | reject_schematic (t $ u) = (reject_schematic t; reject_schematic u)
   496   | reject_schematic _ = ();
   497 
   498 fun expand_binds ctxt =
   499   let val Mode {pattern, schematic, ...} = get_mode ctxt in
   500     if pattern then I
   501     else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic)
   502   end;
   503 
   504 in
   505 
   506 fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt;
   507 
   508 end;
   509 
   510 
   511 fun contract_abbrevs ctxt t =
   512   let
   513     val thy = theory_of ctxt;
   514     val consts = consts_of ctxt;
   515     val Mode {abbrev, ...} = get_mode ctxt;
   516     val retrieve = Consts.retrieve_abbrevs consts (print_mode_value () @ [""]);
   517     fun match_abbrev u = Option.map #1 (get_first (Pattern.match_rew thy u) (retrieve u));
   518   in
   519     if abbrev orelse print_mode_active "no_abbrevs" orelse not (can Term.type_of t) then t
   520     else Pattern.rewrite_term thy [] [match_abbrev] t
   521   end;
   522 
   523 
   524 (* patterns *)
   525 
   526 fun prepare_patternT ctxt T =
   527   let
   528     val Mode {pattern, schematic, ...} = get_mode ctxt;
   529     val _ =
   530       pattern orelse schematic orelse
   531         T |> Term.exists_subtype
   532           (fn TVar (xi, _) =>
   533             not (TypeInfer.is_param xi) andalso
   534               error ("Illegal schematic type variable: " ^ Term.string_of_vname xi)
   535           | _ => false)
   536   in T end;
   537 
   538 
   539 local
   540 
   541 structure Allow_Dummies = ProofDataFun(type T = bool fun init _ = false);
   542 
   543 fun check_dummies ctxt t =
   544   if Allow_Dummies.get ctxt then t
   545   else Term.no_dummy_patterns t handle TERM _ => error "Illegal dummy pattern(s) in term";
   546 
   547 fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1);
   548 
   549 in
   550 
   551 val allow_dummies = Allow_Dummies.put true;
   552 
   553 fun prepare_patterns ctxt =
   554   let val Mode {pattern, ...} = get_mode ctxt in
   555     TypeInfer.fixate_params (Variable.names_of ctxt) #>
   556     pattern ? Variable.polymorphic ctxt #>
   557     (map o Term.map_types) (prepare_patternT ctxt) #>
   558     (if pattern then prepare_dummies else map (check_dummies ctxt))
   559   end;
   560 
   561 end;
   562 
   563 
   564 (* decoding raw terms (syntax trees) *)
   565 
   566 (* types *)
   567 
   568 fun get_sort thy def_sort raw_env =
   569   let
   570     val tsig = Sign.tsig_of thy;
   571 
   572     fun eq ((xi, S), (xi', S')) =
   573       Term.eq_ix (xi, xi') andalso Type.eq_sort tsig (S, S');
   574     val env = distinct eq raw_env;
   575     val _ = (case duplicates (eq_fst (op =)) env of [] => ()
   576       | dups => error ("Inconsistent sort constraints for type variable(s) "
   577           ^ commas_quote (map (Term.string_of_vname' o fst) dups)));
   578 
   579     fun get xi =
   580       (case (AList.lookup (op =) env xi, def_sort xi) of
   581         (NONE, NONE) => Type.defaultS tsig
   582       | (NONE, SOME S) => S
   583       | (SOME S, NONE) => S
   584       | (SOME S, SOME S') =>
   585           if Type.eq_sort tsig (S, S') then S'
   586           else error ("Sort constraint " ^ Syntax.string_of_sort_global thy S ^
   587             " inconsistent with default " ^ Syntax.string_of_sort_global thy S' ^
   588             " for type variable " ^ quote (Term.string_of_vname' xi)));
   589   in get end;
   590 
   591 local
   592 
   593 fun intern_skolem ctxt def_type x =
   594   let
   595     val _ = no_skolem false x;
   596     val sko = lookup_skolem ctxt x;
   597     val is_const = can (read_const_proper ctxt) x orelse Long_Name.is_qualified x;
   598     val is_declared = is_some (def_type (x, ~1));
   599   in
   600     if Variable.is_const ctxt x then NONE
   601     else if is_some sko then sko
   602     else if not is_const orelse is_declared then SOME x
   603     else NONE
   604   end;
   605 
   606 in
   607 
   608 fun term_context ctxt =
   609   let val thy = theory_of ctxt in
   610    {get_sort = get_sort thy (Variable.def_sort ctxt),
   611     map_const = fn a => ((true, #1 (Term.dest_Const (read_const_proper ctxt a)))
   612       handle ERROR _ => (false, Consts.intern (consts_of ctxt) a)),
   613     map_free = intern_skolem ctxt (Variable.def_type ctxt false),
   614     map_type = Sign.intern_tycons thy,
   615     map_sort = Sign.intern_sort thy}
   616   end;
   617 
   618 fun decode_term ctxt =
   619   let val {get_sort, map_const, map_free, map_type, map_sort} = term_context ctxt
   620   in Syntax.decode_term get_sort map_const map_free map_type map_sort end;
   621 
   622 end;
   623 
   624 
   625 (* certify terms *)
   626 
   627 local
   628 
   629 fun gen_cert prop ctxt t =
   630   t
   631   |> expand_abbrevs ctxt
   632   |> (fn t' => #1 (Sign.certify' prop (Syntax.pp ctxt) false (consts_of ctxt) (theory_of ctxt) t')
   633     handle TYPE (msg, _, _) => error msg
   634       | TERM (msg, _) => error msg);
   635 
   636 in
   637 
   638 val cert_term = gen_cert false;
   639 val cert_prop = gen_cert true;
   640 
   641 end;
   642 
   643 
   644 (* type checking/inference *)
   645 
   646 fun standard_infer_types ctxt ts =
   647   let val Mode {pattern, ...} = get_mode ctxt in
   648     TypeInfer.infer_types (Syntax.pp ctxt) (tsig_of ctxt) (Syntax.check_typs ctxt)
   649       (try (Consts.the_constraint (consts_of ctxt))) (Variable.def_type ctxt pattern)
   650       (Variable.names_of ctxt) (Variable.maxidx_of ctxt) ts
   651     handle TYPE (msg, _, _) => error msg
   652   end;
   653 
   654 local
   655 
   656 fun standard_typ_check ctxt =
   657   map (cert_typ_mode (Type.get_mode ctxt) ctxt) #>
   658   map (prepare_patternT ctxt);
   659 
   660 fun standard_term_check ctxt =
   661   standard_infer_types ctxt #>
   662   map (expand_abbrevs ctxt);
   663 
   664 fun standard_term_uncheck ctxt =
   665   map (contract_abbrevs ctxt);
   666 
   667 fun add eq what f = Context.>> (what (fn xs => fn ctxt =>
   668   let val xs' = f ctxt xs in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end));
   669 
   670 in
   671 
   672 val _ = add (op =) (Syntax.add_typ_check 0 "standard") standard_typ_check;
   673 val _ = add (op aconv) (Syntax.add_term_check 0 "standard") standard_term_check;
   674 val _ = add (op aconv) (Syntax.add_term_check 100 "fixate") prepare_patterns;
   675 
   676 val _ = add (op aconv) (Syntax.add_term_uncheck 0 "standard") standard_term_uncheck;
   677 
   678 end;
   679 
   680 
   681 
   682 (** inner syntax operations **)
   683 
   684 local
   685 
   686 fun parse_sort ctxt text =
   687   let
   688     val (syms, pos) = Syntax.parse_token Markup.sort text;
   689     val S = Syntax.standard_parse_sort ctxt (syn_of ctxt)
   690         (Sign.intern_sort (theory_of ctxt)) (syms, pos)
   691       handle ERROR msg => cat_error msg  ("Failed to parse sort" ^ Position.str_of pos)
   692   in S end;
   693 
   694 fun parse_typ ctxt text =
   695   let
   696     val thy = ProofContext.theory_of ctxt;
   697     val get_sort = get_sort thy (Variable.def_sort ctxt);
   698 
   699     val (syms, pos) = Syntax.parse_token Markup.typ text;
   700     val T = Sign.intern_tycons thy
   701         (Syntax.standard_parse_typ ctxt (syn_of ctxt) get_sort (Sign.intern_sort thy) (syms, pos))
   702       handle ERROR msg => cat_error msg  ("Failed to parse type" ^ Position.str_of pos);
   703   in T end;
   704 
   705 fun parse_term T ctxt text =
   706   let
   707     val thy = theory_of ctxt;
   708     val {get_sort, map_const, map_free, map_type, map_sort} = term_context ctxt;
   709 
   710     val (T', _) = TypeInfer.paramify_dummies T 0;
   711     val (markup, kind) = if T' = propT then (Markup.prop, "proposition") else (Markup.term, "term");
   712     val (syms, pos) = Syntax.parse_token markup text;
   713 
   714     fun check t = (Syntax.check_term ctxt (TypeInfer.constrain T' t); NONE)
   715       handle ERROR msg => SOME msg;
   716     val t = Syntax.standard_parse_term (Syntax.pp ctxt) check get_sort map_const map_free
   717         map_type map_sort ctxt (Sign.is_logtype thy) (syn_of ctxt) T' (syms, pos)
   718       handle ERROR msg => cat_error msg  ("Failed to parse " ^ kind ^ Position.str_of pos);
   719   in t end;
   720 
   721 
   722 fun unparse_sort ctxt S =
   723   Syntax.standard_unparse_sort ctxt (syn_of ctxt) (Sign.extern_sort (theory_of ctxt) S);
   724 
   725 fun unparse_typ ctxt T =
   726   Syntax.standard_unparse_typ ctxt (syn_of ctxt) (Sign.extern_typ (theory_of ctxt) T);
   727 
   728 fun unparse_term ctxt t =
   729   let
   730     val thy = theory_of ctxt;
   731     val syntax = syntax_of ctxt;
   732     val consts = consts_of ctxt;
   733   in
   734     t
   735     |> Sign.extern_term (Consts.extern_early consts) thy
   736     |> LocalSyntax.extern_term syntax
   737     |> Syntax.standard_unparse_term (Consts.extern consts) ctxt (LocalSyntax.syn_of syntax)
   738         (not (PureThy.old_appl_syntax thy))
   739   end;
   740 
   741 in
   742 
   743 val _ = Syntax.install_operations
   744   {parse_sort = parse_sort,
   745    parse_typ = parse_typ,
   746    parse_term = parse_term dummyT,
   747    parse_prop = parse_term propT,
   748    unparse_sort = unparse_sort,
   749    unparse_typ = unparse_typ,
   750    unparse_term = unparse_term};
   751 
   752 end;
   753 
   754 
   755 
   756 (** export results **)
   757 
   758 fun common_export is_goal inner outer =
   759   map (Assumption.export is_goal inner outer) #>
   760   Variable.export inner outer;
   761 
   762 val goal_export = common_export true;
   763 val export = common_export false;
   764 
   765 fun export_morphism inner outer =
   766   Assumption.export_morphism inner outer $>
   767   Variable.export_morphism inner outer;
   768 
   769 fun norm_export_morphism inner outer =
   770   export_morphism inner outer $>
   771   Morphism.thm_morphism Goal.norm_result;
   772 
   773 
   774 
   775 (** term bindings **)
   776 
   777 (* simult_matches *)
   778 
   779 fun simult_matches ctxt (t, pats) =
   780   (case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of
   781     NONE => error "Pattern match failed!"
   782   | SOME (env, _) => map (apsnd snd) (Envir.alist_of env));
   783 
   784 
   785 (* bind_terms *)
   786 
   787 val bind_terms = fold (fn (xi, t) => fn ctxt =>
   788   ctxt
   789   |> Variable.bind_term (xi, Option.map (cert_term (set_mode mode_default ctxt)) t));
   790 
   791 
   792 (* auto_bind *)
   793 
   794 fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b
   795   | drop_schematic b = b;
   796 
   797 fun auto_bind f ts ctxt = ctxt |> bind_terms (map drop_schematic (f (theory_of ctxt) ts));
   798 
   799 val auto_bind_goal = auto_bind AutoBind.goal;
   800 val auto_bind_facts = auto_bind AutoBind.facts;
   801 
   802 
   803 (* match_bind(_i) *)
   804 
   805 local
   806 
   807 fun gen_bind prep_terms gen raw_binds ctxt =
   808   let
   809     fun prep_bind (raw_pats, t) ctxt1 =
   810       let
   811         val T = Term.fastype_of t;
   812         val ctxt2 = Variable.declare_term t ctxt1;
   813         val pats = prep_terms (set_mode mode_pattern ctxt2) T raw_pats;
   814         val binds = simult_matches ctxt2 (t, pats);
   815       in (binds, ctxt2) end;
   816 
   817     val ts = prep_terms ctxt dummyT (map snd raw_binds);
   818     val (binds, ctxt') = apfst flat (fold_map prep_bind (map fst raw_binds ~~ ts) ctxt);
   819     val binds' =
   820       if gen then map #1 binds ~~ Variable.exportT_terms ctxt' ctxt (map #2 binds)
   821       else binds;
   822     val binds'' = map (apsnd SOME) binds';
   823     val ctxt'' =
   824       tap (Variable.warn_extra_tfrees ctxt)
   825        (if gen then
   826           ctxt (*sic!*) |> fold Variable.declare_term (map #2 binds') |> bind_terms binds''
   827         else ctxt' |> bind_terms binds'');
   828   in (ts, ctxt'') end;
   829 
   830 in
   831 
   832 fun read_terms ctxt T =
   833   map (Syntax.parse_term ctxt #> TypeInfer.constrain T) #> Syntax.check_terms ctxt;
   834 
   835 val match_bind = gen_bind read_terms;
   836 val match_bind_i = gen_bind (fn ctxt => fn _ => map (cert_term ctxt));
   837 
   838 end;
   839 
   840 
   841 (* propositions with patterns *)
   842 
   843 local
   844 
   845 fun prep_propp mode prep_props (context, args) =
   846   let
   847     fun prep (_, raw_pats) (ctxt, prop :: props) =
   848       let val ctxt' = Variable.declare_term prop ctxt
   849       in ((prop, prep_props (set_mode mode_pattern ctxt') raw_pats), (ctxt', props)) end;
   850 
   851     val (propp, (context', _)) = (fold_map o fold_map) prep args
   852       (context, prep_props (set_mode mode context) (maps (map fst) args));
   853   in (context', propp) end;
   854 
   855 fun gen_bind_propp mode parse_prop (ctxt, raw_args) =
   856   let
   857     val (ctxt', args) = prep_propp mode parse_prop (ctxt, raw_args);
   858     val binds = flat (flat (map (map (simult_matches ctxt')) args));
   859     val propss = map (map #1) args;
   860 
   861     (*generalize result: context evaluated now, binds added later*)
   862     val gen = Variable.exportT_terms ctxt' ctxt;
   863     fun gen_binds c = c |> bind_terms (map #1 binds ~~ map SOME (gen (map #2 binds)));
   864   in (ctxt' |> bind_terms (map (apsnd SOME) binds), (propss, gen_binds)) end;
   865 
   866 in
   867 
   868 val read_propp           = prep_propp mode_default Syntax.read_props;
   869 val cert_propp           = prep_propp mode_default (map o cert_prop);
   870 val read_propp_schematic = prep_propp mode_schematic Syntax.read_props;
   871 val cert_propp_schematic = prep_propp mode_schematic (map o cert_prop);
   872 
   873 val bind_propp             = gen_bind_propp mode_default Syntax.read_props;
   874 val bind_propp_i           = gen_bind_propp mode_default (map o cert_prop);
   875 val bind_propp_schematic   = gen_bind_propp mode_schematic Syntax.read_props;
   876 val bind_propp_schematic_i = gen_bind_propp mode_schematic (map o cert_prop);
   877 
   878 end;
   879 
   880 
   881 
   882 (** theorems **)
   883 
   884 (* fact_tac *)
   885 
   886 fun comp_incr_tac [] _ = no_tac
   887   | comp_incr_tac (th :: ths) i =
   888       (fn st => Goal.compose_hhf_tac (Drule.incr_indexes st th) i st) APPEND comp_incr_tac ths i;
   889 
   890 fun fact_tac facts = Goal.norm_hhf_tac THEN' comp_incr_tac facts;
   891 
   892 fun potential_facts ctxt prop =
   893   Facts.could_unify (facts_of ctxt) (Term.strip_all_body prop);
   894 
   895 fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => fact_tac (potential_facts ctxt goal) i);
   896 
   897 
   898 (* get_thm(s) *)
   899 
   900 local
   901 
   902 fun retrieve_thms pick ctxt (Facts.Fact s) =
   903       let
   904         val (_, pos) = Syntax.read_token s;
   905         val prop = Syntax.read_prop (set_mode mode_default ctxt) s
   906           |> singleton (Variable.polymorphic ctxt);
   907 
   908         fun prove_fact th =
   909           Goal.prove ctxt [] [] prop (K (ALLGOALS (fact_tac [th])))
   910           |> Thm.default_position_of th;
   911         val res =
   912           (case get_first (try prove_fact) (potential_facts ctxt prop) of
   913             SOME res => res
   914           | NONE => error ("Failed to retrieve literal fact" ^ Position.str_of pos ^ ":\n" ^
   915               Syntax.string_of_term ctxt prop))
   916       in pick "" [res] end
   917   | retrieve_thms pick ctxt xthmref =
   918       let
   919         val thy = theory_of ctxt;
   920         val local_facts = facts_of ctxt;
   921         val thmref = Facts.map_name_of_ref (Facts.intern local_facts) xthmref;
   922         val name = Facts.name_of_ref thmref;
   923         val pos = Facts.pos_of_ref xthmref;
   924         val thms =
   925           if name = "" then [Thm.transfer thy Drule.dummy_thm]
   926           else
   927             (case Facts.lookup (Context.Proof ctxt) local_facts name of
   928               SOME (_, ths) => (Position.report (Markup.local_fact name) pos;
   929                 map (Thm.transfer thy) (Facts.select thmref ths))
   930             | NONE => PureThy.get_fact (Context.Proof ctxt) thy xthmref);
   931       in pick name thms end;
   932 
   933 in
   934 
   935 val get_fact = retrieve_thms (K I);
   936 val get_fact_single = retrieve_thms Facts.the_single;
   937 
   938 fun get_thms ctxt = get_fact ctxt o Facts.named;
   939 fun get_thm ctxt = get_fact_single ctxt o Facts.named;
   940 
   941 end;
   942 
   943 
   944 (* name space operations *)
   945 
   946 val add_path        = map_naming o NameSpace.add_path;
   947 val mandatory_path  = map_naming o NameSpace.mandatory_path;
   948 val restore_naming  = map_naming o K o naming_of;
   949 val reset_naming    = map_naming (K local_naming);
   950 
   951 
   952 (* facts *)
   953 
   954 local
   955 
   956 fun update_thms _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b))
   957   | update_thms do_props (b, SOME ths) ctxt = ctxt |> map_facts
   958       (Facts.add_local do_props (naming_of ctxt) (b, ths) #> snd);
   959 
   960 in
   961 
   962 fun note_thmss kind = fold_map (fn ((b, more_attrs), raw_facts) => fn ctxt =>
   963   let
   964     val pos = Binding.pos_of b;
   965     val name = full_name ctxt b;
   966     val _ = ContextPosition.report_visible ctxt (Markup.local_fact_decl name) pos;
   967 
   968     val facts = PureThy.name_thmss false pos name raw_facts;
   969     fun app (th, attrs) x =
   970       swap (Library.foldl_map
   971         (Thm.proof_attributes (surround (Thm.kind kind) (attrs @ more_attrs))) (x, th));
   972     val (res, ctxt') = fold_map app facts ctxt;
   973     val thms = PureThy.name_thms false false pos name (flat res);
   974     val Mode {stmt, ...} = get_mode ctxt;
   975   in ((name, thms), ctxt' |> update_thms stmt (b, SOME thms)) end);
   976 
   977 fun put_thms do_props thms ctxt = ctxt
   978   |> map_naming (K local_naming)
   979   |> ContextPosition.set_visible false
   980   |> update_thms do_props (apfst Binding.name thms)
   981   |> ContextPosition.restore_visible ctxt
   982   |> restore_naming ctxt;
   983 
   984 end;
   985 
   986 
   987 
   988 (** parameters **)
   989 
   990 (* variables *)
   991 
   992 fun declare_var (x, opt_T, mx) ctxt =
   993   let val T = (case opt_T of SOME T => T | NONE => Syntax.mixfixT mx)
   994   in ((x, T, mx), ctxt |> Variable.declare_constraints (Free (x, T))) end;
   995 
   996 local
   997 
   998 fun prep_vars prep_typ internal =
   999   fold_map (fn (raw_b, raw_T, raw_mx) => fn ctxt =>
  1000     let
  1001       val raw_x = Name.of_binding raw_b;
  1002       val (x, mx) = Syntax.const_mixfix raw_x raw_mx;
  1003       val _ = Syntax.is_identifier (no_skolem internal x) orelse
  1004         error ("Illegal variable name: " ^ quote x);
  1005 
  1006       fun cond_tvars T =
  1007         if internal then T
  1008         else Type.no_tvars T handle TYPE (msg, _, _) => error msg;
  1009       val opt_T = Option.map (cond_tvars o cert_typ ctxt o prep_typ ctxt) raw_T;
  1010       val var = (Binding.map_name (K x) raw_b, opt_T, mx);
  1011     in (var, ctxt |> declare_var (x, opt_T, mx) |> #2) end);
  1012 
  1013 in
  1014 
  1015 val read_vars = prep_vars Syntax.parse_typ false;
  1016 val cert_vars = prep_vars (K I) true;
  1017 
  1018 end;
  1019 
  1020 
  1021 (* authentic constants *)
  1022 
  1023 local
  1024 
  1025 fun const_ast_tr intern ctxt [Syntax.Variable c] =
  1026       let
  1027         val Const (c', _) = read_const_proper ctxt c;
  1028         val d = if intern then const_syntax_name ctxt c' else c;
  1029       in Syntax.Constant d end
  1030   | const_ast_tr _ _ asts = raise Syntax.AST ("const_ast_tr", asts);
  1031 
  1032 in
  1033 
  1034 val _ = Context.>> (Context.map_theory
  1035  (Sign.add_syntax
  1036    [("_context_const", "id => logic", Delimfix "CONST _"),
  1037     ("_context_const", "id => aprop", Delimfix "CONST _"),
  1038     ("_context_const", "longid => logic", Delimfix "CONST _"),
  1039     ("_context_const", "longid => aprop", Delimfix "CONST _"),
  1040     ("_context_xconst", "id => logic", Delimfix "XCONST _"),
  1041     ("_context_xconst", "id => aprop", Delimfix "XCONST _"),
  1042     ("_context_xconst", "longid => logic", Delimfix "XCONST _"),
  1043     ("_context_xconst", "longid => aprop", Delimfix "XCONST _")] #>
  1044   Sign.add_advanced_trfuns
  1045     ([("_context_const", const_ast_tr true), ("_context_xconst", const_ast_tr false)], [], [], [])));
  1046 
  1047 end;
  1048 
  1049 
  1050 (* notation *)
  1051 
  1052 local
  1053 
  1054 fun const_syntax ctxt (Free (x, T), mx) = SOME (true, (x, T, mx))
  1055   | const_syntax ctxt (Const (c, _), mx) =
  1056       Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx))
  1057   | const_syntax _ _ = NONE;
  1058 
  1059 in
  1060 
  1061 fun notation add mode args ctxt =
  1062   ctxt |> map_syntax
  1063     (LocalSyntax.update_modesyntax (theory_of ctxt) add mode (map_filter (const_syntax ctxt) args));
  1064 
  1065 fun target_notation add mode args phi =
  1066   let val args' = filter (fn (t, _) => Type.similar_types (t, Morphism.term phi t)) args;
  1067   in Context.mapping (Sign.notation add mode args') (notation add mode args') end;
  1068 
  1069 end;
  1070 
  1071 
  1072 (* local constants *)
  1073 
  1074 fun add_const_constraint (c, opt_T) ctxt =
  1075   let
  1076     fun prepT raw_T =
  1077       let val T = cert_typ ctxt raw_T
  1078       in cert_term ctxt (Const (c, T)); T end;
  1079   in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end;
  1080 
  1081 fun add_abbrev mode tags (b, raw_t) ctxt =
  1082   let
  1083     val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t
  1084       handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.str_of b));
  1085     val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];
  1086     val ((lhs, rhs), consts') = consts_of ctxt
  1087       |> Consts.abbreviate (Syntax.pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode tags (b, t);
  1088   in
  1089     ctxt
  1090     |> (map_consts o apfst) (K consts')
  1091     |> Variable.declare_term rhs
  1092     |> pair (lhs, rhs)
  1093   end;
  1094 
  1095 fun revert_abbrev mode c = (map_consts o apfst) (Consts.revert_abbrev mode c);
  1096 
  1097 
  1098 (* fixes *)
  1099 
  1100 local
  1101 
  1102 fun prep_mixfix (x, T, mx) =
  1103   if mx <> NoSyn andalso mx <> Structure andalso
  1104       (can Name.dest_internal x orelse can Name.dest_skolem x) then
  1105     error ("Illegal mixfix syntax for internal/skolem constant " ^ quote x)
  1106   else (true, (x, T, mx));
  1107 
  1108 in
  1109 
  1110 fun add_fixes raw_vars ctxt =
  1111   let
  1112     val (vars, _) = cert_vars raw_vars ctxt;
  1113     val (xs', ctxt') = Variable.add_fixes (map (Name.of_binding o #1) vars) ctxt;
  1114     val ctxt'' =
  1115       ctxt'
  1116       |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)
  1117       |-> (map_syntax o LocalSyntax.add_syntax (theory_of ctxt) o map prep_mixfix);
  1118     val _ = (vars ~~ xs') |> List.app (fn ((b, _, _), x') =>
  1119       ContextPosition.report_visible ctxt (Markup.fixed_decl x') (Binding.pos_of b));
  1120   in (xs', ctxt'') end;
  1121 
  1122 end;
  1123 
  1124 
  1125 (* fixes vs. frees *)
  1126 
  1127 fun auto_fixes (arg as (ctxt, (propss, x))) =
  1128   ((fold o fold) Variable.auto_fixes propss ctxt, (propss, x));
  1129 
  1130 fun bind_fixes xs ctxt =
  1131   let
  1132     val (_, ctxt') = ctxt |> add_fixes (map (fn x => (Binding.name x, NONE, NoSyn)) xs);
  1133     fun bind (t as Free (x, T)) =
  1134           if member (op =) xs x then
  1135             (case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t)
  1136           else t
  1137       | bind (t $ u) = bind t $ bind u
  1138       | bind (Abs (x, T, t)) = Abs (x, T, bind t)
  1139       | bind a = a;
  1140   in (bind, ctxt') end;
  1141 
  1142 
  1143 
  1144 (** assumptions **)
  1145 
  1146 local
  1147 
  1148 fun gen_assms prepp exp args ctxt =
  1149   let
  1150     val cert = Thm.cterm_of (theory_of ctxt);
  1151     val (propss, ctxt1) = swap (prepp (ctxt, map snd args));
  1152     val _ = Variable.warn_extra_tfrees ctxt ctxt1;
  1153     val (premss, ctxt2) = fold_burrow (Assumption.add_assms exp o map cert) propss ctxt1;
  1154   in
  1155     ctxt2
  1156     |> auto_bind_facts (flat propss)
  1157     |> note_thmss Thm.assumptionK (map fst args ~~ map (map (fn th => ([th], []))) premss)
  1158   end;
  1159 
  1160 in
  1161 
  1162 val add_assms = gen_assms (apsnd #1 o bind_propp);
  1163 val add_assms_i = gen_assms (apsnd #1 o bind_propp_i);
  1164 
  1165 end;
  1166 
  1167 
  1168 
  1169 (** cases **)
  1170 
  1171 local
  1172 
  1173 fun rem_case name = remove (fn (x: string, (y, _)) => x = y) name;
  1174 
  1175 fun add_case _ ("", _) cases = cases
  1176   | add_case _ (name, NONE) cases = rem_case name cases
  1177   | add_case is_proper (name, SOME c) cases = (name, (c, is_proper)) :: rem_case name cases;
  1178 
  1179 fun prep_case name fxs c =
  1180   let
  1181     fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys
  1182       | replace [] ys = ys
  1183       | replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name);
  1184     val RuleCases.Case {fixes, assumes, binds, cases} = c;
  1185     val fixes' = replace fxs fixes;
  1186     val binds' = map drop_schematic binds;
  1187   in
  1188     if null (fold (Term.add_tvarsT o snd) fixes []) andalso
  1189       null (fold (fold Term.add_vars o snd) assumes []) then
  1190         RuleCases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases}
  1191     else error ("Illegal schematic variable(s) in case " ^ quote name)
  1192   end;
  1193 
  1194 fun fix (x, T) ctxt =
  1195   let
  1196     val (bind, ctxt') = bind_fixes [x] ctxt;
  1197     val t = bind (Free (x, T));
  1198   in (t, ctxt' |> Variable.declare_constraints t) end;
  1199 
  1200 in
  1201 
  1202 fun add_cases is_proper = map_cases o fold (add_case is_proper);
  1203 
  1204 fun case_result c ctxt =
  1205   let
  1206     val RuleCases.Case {fixes, ...} = c;
  1207     val (ts, ctxt') = ctxt |> fold_map fix fixes;
  1208     val RuleCases.Case {assumes, binds, cases, ...} = RuleCases.apply ts c;
  1209   in
  1210     ctxt'
  1211     |> bind_terms (map drop_schematic binds)
  1212     |> add_cases true (map (apsnd SOME) cases)
  1213     |> pair (assumes, (binds, cases))
  1214   end;
  1215 
  1216 val apply_case = apfst fst oo case_result;
  1217 
  1218 fun get_case ctxt name xs =
  1219   (case AList.lookup (op =) (cases_of ctxt) name of
  1220     NONE => error ("Unknown case: " ^ quote name)
  1221   | SOME (c, _) => prep_case name xs c);
  1222 
  1223 end;
  1224 
  1225 
  1226 
  1227 (** print context information **)
  1228 
  1229 val debug = ref false;
  1230 
  1231 val verbose = ref false;
  1232 fun verb f x = if ! verbose then f (x ()) else [];
  1233 
  1234 fun setmp_verbose f x = Library.setmp verbose true f x;
  1235 
  1236 
  1237 (* local syntax *)
  1238 
  1239 val print_syntax = Syntax.print_syntax o syn_of;
  1240 
  1241 
  1242 (* abbreviations *)
  1243 
  1244 fun pretty_abbrevs show_globals ctxt =
  1245   let
  1246     val ((space, consts), (_, globals)) =
  1247       pairself (#constants o Consts.dest) (#consts (rep_context ctxt));
  1248     fun add_abbr (_, (_, NONE)) = I
  1249       | add_abbr (c, (T, SOME t)) =
  1250           if not show_globals andalso Symtab.defined globals c then I
  1251           else cons (c, Logic.mk_equals (Const (c, T), t));
  1252     val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
  1253   in
  1254     if null abbrevs andalso not (! verbose) then []
  1255     else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
  1256   end;
  1257 
  1258 val print_abbrevs = Pretty.writeln o Pretty.chunks o pretty_abbrevs true;
  1259 
  1260 
  1261 (* term bindings *)
  1262 
  1263 fun pretty_binds ctxt =
  1264   let
  1265     val binds = Variable.binds_of ctxt;
  1266     fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));
  1267   in
  1268     if Vartab.is_empty binds andalso not (! verbose) then []
  1269     else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]
  1270   end;
  1271 
  1272 val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds;
  1273 
  1274 
  1275 (* local theorems *)
  1276 
  1277 fun pretty_lthms ctxt =
  1278   let
  1279     val local_facts = facts_of ctxt;
  1280     val props = Facts.props local_facts;
  1281     val facts =
  1282       (if null props then [] else [("unnamed", props)]) @
  1283       Facts.dest_static [] local_facts;
  1284   in
  1285     if null facts andalso not (! verbose) then []
  1286     else [Pretty.big_list "facts:" (map #1 (sort_wrt (#1 o #2) (map (`(pretty_fact ctxt)) facts)))]
  1287   end;
  1288 
  1289 val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms;
  1290 
  1291 
  1292 (* local contexts *)
  1293 
  1294 local
  1295 
  1296 fun pretty_case (name, (fixes, ((asms, (lets, cs)), ctxt))) =
  1297   let
  1298     val prt_term = Syntax.pretty_term ctxt;
  1299 
  1300     fun prt_let (xi, t) = Pretty.block
  1301       [Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1,
  1302         Pretty.quote (prt_term t)];
  1303 
  1304     fun prt_asm (a, ts) = Pretty.block (Pretty.breaks
  1305       ((if a = "" then [] else [Pretty.str (a ^ ":")]) @ map (Pretty.quote o prt_term) ts));
  1306 
  1307     fun prt_sect _ _ _ [] = []
  1308       | prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s ::
  1309             flat (Library.separate sep (map (Library.single o prt) xs))))];
  1310   in
  1311     Pretty.block (Pretty.fbreaks
  1312       (Pretty.str (name ^ ":") ::
  1313         prt_sect "fix" [] (Pretty.str o fst) fixes @
  1314         prt_sect "let" [Pretty.str "and"] prt_let
  1315           (map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @
  1316         (if forall (null o #2) asms then []
  1317           else prt_sect "assume" [Pretty.str "and"] prt_asm asms) @
  1318         prt_sect "subcases:" [] (Pretty.str o fst) cs))
  1319   end;
  1320 
  1321 in
  1322 
  1323 fun pretty_cases ctxt =
  1324   let
  1325     fun add_case (_, (_, false)) = I
  1326       | add_case (name, (c as RuleCases.Case {fixes, ...}, true)) =
  1327           cons (name, (fixes, case_result c ctxt));
  1328     val cases = fold add_case (cases_of ctxt) [];
  1329   in
  1330     if null cases andalso not (! verbose) then []
  1331     else [Pretty.big_list "cases:" (map pretty_case cases)]
  1332   end;
  1333 
  1334 val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases;
  1335 
  1336 end;
  1337 
  1338 
  1339 (* core context *)
  1340 
  1341 val prems_limit = ref ~1;
  1342 
  1343 fun pretty_ctxt ctxt =
  1344   if ! prems_limit < 0 andalso not (! debug) then []
  1345   else
  1346     let
  1347       val prt_term = Syntax.pretty_term ctxt;
  1348 
  1349       (*structures*)
  1350       val structs = LocalSyntax.structs_of (syntax_of ctxt);
  1351       val prt_structs = if null structs then []
  1352         else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::
  1353           Pretty.commas (map Pretty.str structs))];
  1354 
  1355       (*fixes*)
  1356       fun prt_fix (x, x') =
  1357         if x = x' then Pretty.str x
  1358         else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')];
  1359       val fixes =
  1360         rev (filter_out ((can Name.dest_internal orf member (op =) structs) o #1)
  1361           (Variable.fixes_of ctxt));
  1362       val prt_fixes = if null fixes then []
  1363         else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::
  1364           Pretty.commas (map prt_fix fixes))];
  1365 
  1366       (*prems*)
  1367       val prems = Assumption.all_prems_of ctxt;
  1368       val len = length prems;
  1369       val suppressed = len - ! prems_limit;
  1370       val prt_prems = if null prems then []
  1371         else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @
  1372           map (pretty_thm ctxt) (Library.drop (suppressed, prems)))];
  1373     in prt_structs @ prt_fixes @ prt_prems end;
  1374 
  1375 
  1376 (* main context *)
  1377 
  1378 fun pretty_context ctxt =
  1379   let
  1380     val prt_term = Syntax.pretty_term ctxt;
  1381     val prt_typ = Syntax.pretty_typ ctxt;
  1382     val prt_sort = Syntax.pretty_sort ctxt;
  1383 
  1384     (*theory*)
  1385     val pretty_thy = Pretty.block
  1386       [Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)];
  1387 
  1388     (*defaults*)
  1389     fun prt_atom prt prtT (x, X) = Pretty.block
  1390       [prt x, Pretty.str " ::", Pretty.brk 1, prtT X];
  1391 
  1392     fun prt_var (x, ~1) = prt_term (Syntax.free x)
  1393       | prt_var xi = prt_term (Syntax.var xi);
  1394 
  1395     fun prt_varT (x, ~1) = prt_typ (TFree (x, []))
  1396       | prt_varT xi = prt_typ (TVar (xi, []));
  1397 
  1398     val prt_defT = prt_atom prt_var prt_typ;
  1399     val prt_defS = prt_atom prt_varT prt_sort;
  1400 
  1401     val (types, sorts) = Variable.constraints_of ctxt;
  1402   in
  1403     verb single (K pretty_thy) @
  1404     pretty_ctxt ctxt @
  1405     verb (pretty_abbrevs false) (K ctxt) @
  1406     verb pretty_binds (K ctxt) @
  1407     verb pretty_lthms (K ctxt) @
  1408     verb pretty_cases (K ctxt) @
  1409     verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @
  1410     verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))
  1411   end;
  1412 
  1413 
  1414 (* query meta data *)
  1415 
  1416 val query_type = Type.the_tags o tsig_of;
  1417 
  1418 fun query_const ctxt name =
  1419   Consts.the_tags (consts_of ctxt) name handle TYPE (msg, _, _) => error msg;
  1420 
  1421 fun query_class ctxt name = query_const ctxt (Logic.const_of_class name);
  1422 
  1423 end;