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