src/Pure/Isar/proof_context.ML
author wenzelm
Sun, 09 Feb 2025 12:35:29 +0100
changeset 82120 a4aa45999dd7
parent 81596 af21a61dadad
child 82587 7415414bd9d8
permissions -rw-r--r--
clarified modules: more robust Isabelle symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/proof_context.ML
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     3
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
     4
The key concept of Isar proof contexts: elevates primitive local
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
     5
reasoning Gamma |- phi to a structured concept, with generic context
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
     6
elements.  See also structure Variable and Assumption.
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     7
*)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     8
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     9
signature PROOF_CONTEXT =
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
    10
sig
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    11
  val theory_of: Proof.context -> theory
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36542
diff changeset
    12
  val init_global: theory -> Proof.context
77889
5db014c36f42 clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
wenzelm
parents: 77723
diff changeset
    13
  val get_global: {long: bool} -> theory -> string -> Proof.context
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
    14
  type mode
24501
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
    15
  val mode_default: mode
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
    16
  val mode_pattern: mode
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
    17
  val mode_schematic: mode
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
    18
  val mode_abbrev: mode
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
    19
  val set_mode: mode -> Proof.context -> Proof.context
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
    20
  val get_mode: Proof.context -> mode
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
    21
  val restore_mode: Proof.context -> Proof.context -> Proof.context
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
    22
  val abbrev_mode: Proof.context -> bool
80074
951c371c1cd9 clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents: 79471
diff changeset
    23
  val syntax_of: Proof.context -> Syntax.syntax
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
    24
  val tsig_of: Proof.context -> Type.tsig
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36450
diff changeset
    25
  val set_defsort: sort -> Proof.context -> Proof.context
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
    26
  val default_sort: Proof.context -> indexname -> sort
59840
0ab8750c9342 proper local Proof_Context.arity_sorts;
wenzelm
parents: 59795
diff changeset
    27
  val arity_sorts: Proof.context -> string -> sort -> sort list
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    28
  val consts_of: Proof.context -> Consts.T
20784
eece9aaaf352 Syntax.mode;
wenzelm
parents: 20631
diff changeset
    29
  val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    30
  val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
    31
  val map_naming: (Name_Space.naming -> Name_Space.naming) -> Proof.context -> Proof.context
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
    32
  val naming_of: Proof.context -> Name_Space.naming
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
    33
  val restore_naming: Proof.context -> Proof.context -> Proof.context
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
    34
  val full_name: Proof.context -> binding -> string
59886
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
    35
  val get_scope: Proof.context -> Binding.scope option
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
    36
  val new_scope: Proof.context -> Binding.scope * Proof.context
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59917
diff changeset
    37
  val private_scope: Binding.scope -> Proof.context -> Proof.context
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59917
diff changeset
    38
  val private: Position.T -> Proof.context -> Proof.context
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59970
diff changeset
    39
  val qualified_scope: Binding.scope -> Proof.context -> Proof.context
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59970
diff changeset
    40
  val qualified: Position.T -> Proof.context -> Proof.context
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
    41
  val concealed: Proof.context -> Proof.context
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    42
  val class_space: Proof.context -> Name_Space.T
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    43
  val type_space: Proof.context -> Name_Space.T
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    44
  val const_space: Proof.context -> Name_Space.T
81537
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
    45
  val lookup_free: Proof.context -> string -> string option
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
    46
  val printable_const: Proof.context -> string -> bool
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
    47
  val exclude_consts: Symset.T -> Proof.context -> Proof.context
61261
ddb2da7cb2e4 more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents: 61168
diff changeset
    48
  val defs_context: Proof.context -> Defs.context
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    49
  val intern_class: Proof.context -> xstring -> string
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    50
  val intern_type: Proof.context -> xstring -> string
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    51
  val intern_const: Proof.context -> xstring -> string
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    52
  val extern_class: Proof.context -> string -> xstring
55304
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    53
  val markup_class: Proof.context -> string -> string
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    54
  val pretty_class: Proof.context -> string -> Pretty.T
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    55
  val extern_type: Proof.context -> string -> xstring
55304
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    56
  val markup_type: Proof.context -> string -> string
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    57
  val pretty_type: Proof.context -> string -> Pretty.T
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
    58
  val extern_const: Proof.context -> string -> xstring
81539
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
    59
  val extern_fixed: Proof.context -> string -> string
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
    60
  val extern_entity: Proof.context -> string -> string
55304
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    61
  val markup_const: Proof.context -> string -> string
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
    62
  val pretty_const: Proof.context -> string -> Pretty.T
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    63
  val transfer: theory -> Proof.context -> Proof.context
57928
f4ff42c5b05b transfer result of Global_Theory.add_thms_dynamic to context stack;
wenzelm
parents: 57887
diff changeset
    64
  val transfer_facts: theory -> Proof.context -> Proof.context
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38328
diff changeset
    65
  val background_theory: (theory -> theory) -> Proof.context -> Proof.context
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38328
diff changeset
    66
  val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
    67
  val facts_of: Proof.context -> Facts.T
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
    68
  val facts_of_fact: Proof.context -> string -> Facts.T
63080
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 63057
diff changeset
    69
  val markup_extern_fact: Proof.context -> string -> Markup.T list * xstring
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
    70
  val pretty_thm_name: Proof.context -> Thm_Name.T -> Pretty.T
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
    71
  val print_thm_name: Proof.context -> Thm_Name.T -> string
70308
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
    72
  val augment: term -> Proof.context -> Proof.context
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
    73
  val pretty_term_abbrev: Proof.context -> term -> Pretty.T
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    74
  val pretty_fact: Proof.context -> string * thm list -> Pretty.T
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
    75
  val check_class: Proof.context -> xstring * Position.T -> class * Position.report list
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
    76
  val read_class: Proof.context -> string -> class
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    77
  val read_typ: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    78
  val read_typ_syntax: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    79
  val read_typ_abbrev: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    80
  val cert_typ: Proof.context -> typ -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    81
  val cert_typ_syntax: Proof.context -> typ -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    82
  val cert_typ_abbrev: Proof.context -> typ -> typ
36505
79c1d2bbe5a9 ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents: 36503
diff changeset
    83
  val infer_type: Proof.context -> string * typ -> typ
60407
53ef7b78e78a tuned signature;
wenzelm
parents: 60401
diff changeset
    84
  val inferred_param: string -> Proof.context -> (string * typ) * Proof.context
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
    85
  val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
    86
  val check_type_name: {proper: bool, strict: bool} -> Proof.context ->
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
    87
    xstring * Position.T -> typ * Position.report list
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
    88
  val read_type_name: {proper: bool, strict: bool} -> Proof.context -> string -> typ
55956
94d384d621b0 reject internal term names outright, and complete consts instead;
wenzelm
parents: 55955
diff changeset
    89
  val consts_completion_message: Proof.context -> xstring * Position.T list -> string
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
    90
  val check_const: {proper: bool, strict: bool} -> Proof.context ->
55959
c3b458435f4f more decisive commitment to get_free vs. the_const;
wenzelm
parents: 55956
diff changeset
    91
    xstring * Position.T list -> term * Position.report list
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
    92
  val read_const: {proper: bool, strict: bool} -> Proof.context -> string -> term
46922
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
    93
  val read_arity: Proof.context -> xstring * string list * string -> arity
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
    94
  val cert_arity: Proof.context -> arity -> arity
27259
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
    95
  val allow_dummies: Proof.context -> Proof.context
56333
38f1422ef473 support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents: 56294
diff changeset
    96
  val prepare_sortsT: Proof.context -> typ list -> string list * typ list
38f1422ef473 support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents: 56294
diff changeset
    97
  val prepare_sorts: Proof.context -> term list -> string list * term list
36152
34d1ce2d746d get_sort: suppress dummyS from input;
wenzelm
parents: 35680
diff changeset
    98
  val check_tfree: Proof.context -> string * sort -> string * sort
24684
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
    99
  val read_term_pattern: Proof.context -> string -> term
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   100
  val read_term_schematic: Proof.context -> string -> term
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   101
  val read_term_abbrev: Proof.context -> string -> term
40879
ca132ef44944 configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents: 39557
diff changeset
   102
  val show_abbrevs: bool Config.T
81236
7c6665613190 tuned signature;
wenzelm
parents: 81232
diff changeset
   103
  val contract_abbrevs: Proof.context -> term -> term
25345
dd5b851f8ef0 renamed ProofContext.read_const' to ProofContext.read_const_proper;
wenzelm
parents: 25332
diff changeset
   104
  val expand_abbrevs: Proof.context -> term -> term
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   105
  val cert_term: Proof.context -> term -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   106
  val cert_prop: Proof.context -> term -> term
35616
b342390d296f provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents: 35429
diff changeset
   107
  val def_type: Proof.context -> indexname -> typ option
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   108
  val standard_typ_check: Proof.context -> typ list -> typ list
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   109
  val standard_term_check_finish: Proof.context -> term list -> term list
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   110
  val standard_term_uncheck: Proof.context -> term list -> term list
78086
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   111
  val export_: {goal: bool} -> Proof.context -> Proof.context -> thm list -> thm list
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   112
  val export: Proof.context -> Proof.context -> thm list -> thm list
78086
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   113
  val export_goal: Proof.context -> Proof.context -> thm list -> thm list
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   114
  val export_morphism: Proof.context -> Proof.context -> morphism
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   115
  val auto_bind_goal: term list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   116
  val auto_bind_facts: term list -> Proof.context -> Proof.context
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   117
  val simult_matches: Proof.context -> term * term list -> (indexname * term) list
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   118
  val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   119
  val bind_term: indexname * term -> Proof.context -> Proof.context
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   120
  val cert_propp: Proof.context -> (term * term list) list list ->
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   121
    (term list list * (indexname * term) list)
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   122
  val read_propp: Proof.context -> (string * string list) list list ->
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   123
    (term list list * (indexname * term) list)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54740
diff changeset
   124
  val fact_tac: Proof.context -> thm list -> int -> tactic
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   125
  val some_fact_tac: Proof.context -> int -> tactic
63255
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
   126
  val lookup_fact: Proof.context -> string -> {dynamic: bool, thms: thm list} option
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
   127
  val dynamic_facts_dummy: bool Config.T
57942
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
   128
  val get_fact_generic: Context.generic -> Facts.ref -> string option * thm list
26346
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
   129
  val get_fact: Proof.context -> Facts.ref -> thm list
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
   130
  val get_fact_single: Proof.context -> Facts.ref -> thm
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
   131
  val get_thms: Proof.context -> xstring -> thm list
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
   132
  val get_thm: Proof.context -> xstring -> thm
67679
8fd84fe1d60b tuned signature;
wenzelm
parents: 67671
diff changeset
   133
  val is_stmt: Proof.context -> bool
63083
wenzelm
parents: 63080
diff changeset
   134
  val set_stmt: bool -> Proof.context -> Proof.context
wenzelm
parents: 63080
diff changeset
   135
  val restore_stmt: Proof.context -> Proof.context -> Proof.context
61811
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   136
  val add_thms_dynamic: binding * (Context.generic -> thm list) ->
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   137
    Proof.context -> string * Proof.context
67671
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
   138
  val add_thms_lazy: string -> (binding * thm list lazy) -> Proof.context -> Proof.context
67713
041898537c19 tuned signature;
wenzelm
parents: 67679
diff changeset
   139
  val note_thms: string -> Thm.binding * (thm list * attribute list) list ->
041898537c19 tuned signature;
wenzelm
parents: 67679
diff changeset
   140
    Proof.context -> (string * thm list) * Proof.context
30761
ac7570d80c3d renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents: 30758
diff changeset
   141
  val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list ->
30211
556d1810cdad Thm.binding;
wenzelm
parents: 30190
diff changeset
   142
    Proof.context -> (string * thm list) list * Proof.context
28082
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
   143
  val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context
66246
c2c18b6b48da tuned signature;
wenzelm
parents: 66245
diff changeset
   144
  val alias_fact: binding -> string -> Proof.context -> Proof.context
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
   145
  val read_var: binding * string option * mixfix -> Proof.context ->
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
   146
    (binding * typ option * mixfix) * Proof.context
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
   147
  val cert_var: binding * typ option * mixfix -> Proof.context ->
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
   148
    (binding * typ option * mixfix) * Proof.context
30763
6976521b4263 renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents: 30761
diff changeset
   149
  val add_fixes: (binding * typ option * mixfix) list -> Proof.context ->
6976521b4263 renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents: 30761
diff changeset
   150
    string list * Proof.context
60469
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
   151
  val add_fixes_cmd: (binding * string option * mixfix) list -> Proof.context ->
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
   152
    string list * Proof.context
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   153
  val add_assms: Assumption.export ->
60377
234b7da8542e tuned signature;
wenzelm
parents: 60367
diff changeset
   154
    (Thm.binding * (term * term list) list) list ->
28082
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
   155
    Proof.context -> (string * thm list) list * Proof.context
60377
234b7da8542e tuned signature;
wenzelm
parents: 60367
diff changeset
   156
  val add_assms_cmd: Assumption.export ->
234b7da8542e tuned signature;
wenzelm
parents: 60367
diff changeset
   157
    (Thm.binding * (string * string list) list) list ->
28082
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
   158
    Proof.context -> (string * thm list) list * Proof.context
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
   159
  val dest_cases: Proof.context option -> Proof.context -> (string * Rule_Cases.T) list
60573
e549969355b2 tuned signature;
wenzelm
parents: 60565
diff changeset
   160
  val update_cases: (string * Rule_Cases.T option) list -> Proof.context -> Proof.context
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60489
diff changeset
   161
  val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context
57486
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
   162
  val check_case: Proof.context -> bool ->
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
   163
    string * Position.T -> binding option list -> Rule_Cases.T
80749
232a839ef8e6 clarified signature: more operations;
wenzelm
parents: 80744
diff changeset
   164
  val is_syntax_const: Proof.context -> string -> bool
74331
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
   165
  val check_syntax_const: Proof.context -> string * Position.T -> string
74333
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
   166
  val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list ->
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
   167
    Proof.context -> Proof.context
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
   168
  val generic_syntax: bool -> Syntax.mode -> (string * typ * mixfix) list ->
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
   169
    Context.generic -> Context.generic
81594
7e1b66416b7f commands 'syntax_types' and 'syntax_consts' now work in a local theory context;
wenzelm
parents: 81591
diff changeset
   170
  val generic_syntax_deps: (string * string list) list -> Context.generic -> Context.generic
81591
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
   171
  val translations: bool -> Ast.ast Syntax.trrule list -> Proof.context -> Proof.context
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
   172
  val generic_translations: bool -> Ast.ast Syntax.trrule list -> Context.generic -> Context.generic
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
   173
  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
   174
  val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context
47247
23654b331d5c tuned signature;
wenzelm
parents: 47005
diff changeset
   175
  val generic_type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> morphism ->
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
   176
    Context.generic -> Context.generic
47247
23654b331d5c tuned signature;
wenzelm
parents: 47005
diff changeset
   177
  val generic_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   178
    Context.generic -> Context.generic
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
   179
  val type_alias: binding -> string -> Proof.context -> Proof.context
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
   180
  val const_alias: binding -> string -> Proof.context -> Proof.context
24767
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
   181
  val add_const_constraint: string * typ option -> Proof.context -> Proof.context
33173
b8ca12f6681a eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents: 33165
diff changeset
   182
  val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
25052
a014d544f54d added revert_abbrev;
wenzelm
parents: 25039
diff changeset
   183
  val revert_abbrev: string -> string -> Proof.context -> Proof.context
47275
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
   184
  val generic_add_abbrev: string -> binding * term -> Context.generic ->
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
   185
    (term * term) * Context.generic
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
   186
  val generic_revert_abbrev: string -> string -> Context.generic -> Context.generic
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   187
  type stmt =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   188
   {vars: ((binding * typ option * mixfix) * (string * term)) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   189
    propss: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   190
    binds: (indexname * term) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   191
    result_binds: (indexname * term) list}
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   192
  val cert_stmt: (binding * typ option * mixfix) list -> (term * term list) list list ->
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   193
    Proof.context -> stmt * Proof.context
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   194
  val read_stmt: (binding * string option * mixfix) list -> (string * string list) list list ->
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   195
    Proof.context -> stmt * Proof.context
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   196
  type statement =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   197
   {fixes: (string * term) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   198
    assumes: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   199
    shows: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   200
    result_binds: (indexname * term option) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   201
    text: term,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   202
    result_text: term}
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   203
  val cert_statement: (binding * typ option * mixfix) list ->
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   204
    (term * term list) list list -> (term * term list) list list -> Proof.context ->
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   205
    statement * Proof.context
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   206
  val read_statement: (binding * string option * mixfix) list ->
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   207
    (string * string list) list list -> (string * string list) list list -> Proof.context ->
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
   208
    statement * Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   209
  val print_syntax: Proof.context -> unit
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
   210
  val print_abbrevs: bool -> Proof.context -> unit
57415
e721124f1b1e command 'print_term_bindings' supersedes 'print_binds';
wenzelm
parents: 56867
diff changeset
   211
  val pretty_term_bindings: Proof.context -> Pretty.T list
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
   212
  val pretty_local_facts: bool -> Proof.context -> Pretty.T list
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
   213
  val print_local_facts: bool -> Proof.context -> unit
56867
224109105008 more print operations;
wenzelm
parents: 56864
diff changeset
   214
  val pretty_cases: Proof.context -> Pretty.T list
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
   215
  val print_cases_proof: Proof.context -> Proof.context -> string
42717
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42502
diff changeset
   216
  val debug: bool Config.T
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42502
diff changeset
   217
  val verbose: bool Config.T
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   218
  val pretty_ctxt: Proof.context -> Pretty.T list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   219
  val pretty_context: Proof.context -> Pretty.T list
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   220
end;
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   221
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   222
structure Proof_Context: PROOF_CONTEXT =
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   223
struct
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   224
42363
e52e3f697510 do not open ML structures;
wenzelm
parents: 42360
diff changeset
   225
val theory_of = Proof_Context.theory_of;
e52e3f697510 do not open ML structures;
wenzelm
parents: 42360
diff changeset
   226
val init_global = Proof_Context.init_global;
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 51601
diff changeset
   227
val get_global = Proof_Context.get_global;
42363
e52e3f697510 do not open ML structures;
wenzelm
parents: 42360
diff changeset
   228
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
   229
7270
2b64729d9acb warn_vars;
wenzelm
parents: 7200
diff changeset
   230
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   231
(** inner syntax mode **)
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   232
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   233
datatype mode =
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   234
  Mode of
63083
wenzelm
parents: 63080
diff changeset
   235
   {pattern: bool,             (*pattern binding schematic variables*)
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   236
    schematic: bool,           (*term referencing loose schematic variables*)
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   237
    abbrev: bool};             (*abbrev mode -- no normalization*)
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   238
63083
wenzelm
parents: 63080
diff changeset
   239
fun make_mode (pattern, schematic, abbrev) =
wenzelm
parents: 63080
diff changeset
   240
  Mode {pattern = pattern, schematic = schematic, abbrev = abbrev};
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   241
63083
wenzelm
parents: 63080
diff changeset
   242
val mode_default   = make_mode (false, false, false);
wenzelm
parents: 63080
diff changeset
   243
val mode_pattern   = make_mode (true, false, false);
wenzelm
parents: 63080
diff changeset
   244
val mode_schematic = make_mode (false, true, false);
wenzelm
parents: 63080
diff changeset
   245
val mode_abbrev    = make_mode (false, false, true);
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   246
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   247
70307
53d21039518a tuned whitespace;
wenzelm
parents: 70255
diff changeset
   248
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   249
(** Isar proof context information **)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   250
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
   251
type cases = Rule_Cases.T Name_Space.table;
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
   252
val empty_cases: cases = Name_Space.empty_table Markup.caseN;
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
   253
45650
wenzelm
parents: 45453
diff changeset
   254
datatype data =
wenzelm
parents: 45453
diff changeset
   255
  Data of
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36450
diff changeset
   256
   {mode: mode,                  (*inner syntax mode*)
ddc965e172c4 localized default sort;
wenzelm
parents: 36450
diff changeset
   257
    syntax: Local_Syntax.T,      (*local syntax*)
ddc965e172c4 localized default sort;
wenzelm
parents: 36450
diff changeset
   258
    tsig: Type.tsig * Type.tsig, (*local/global type signature -- local name space / defsort only*)
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   259
    consts: Symset.T * Consts.T * Consts.T,
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   260
      (*exclude consts, local/global consts -- local name space / abbrevs only*)
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   261
    facts: Facts.T,              (*local facts, based on initial global facts*)
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
   262
    cases: cases};               (*named case contexts*)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   263
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   264
fun make_data (mode, syntax, tsig, consts, facts, cases) =
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   265
  Data {mode = mode, syntax = syntax, tsig = tsig, consts = consts, facts = facts, cases = cases};
19079
9a7678a0736d added put_thms_internal: local_naming, no fact index;
wenzelm
parents: 19062
diff changeset
   266
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37145
diff changeset
   267
structure Data = Proof_Data
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   268
(
45650
wenzelm
parents: 45453
diff changeset
   269
  type T = data;
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   270
  fun init thy =
56139
b7add947a6ef more frugal recording of changes: join merely requires information from one side;
wenzelm
parents: 56062
diff changeset
   271
    make_data (mode_default,
b7add947a6ef more frugal recording of changes: join merely requires information from one side;
wenzelm
parents: 56062
diff changeset
   272
      Local_Syntax.init thy,
b7add947a6ef more frugal recording of changes: join merely requires information from one side;
wenzelm
parents: 56062
diff changeset
   273
      (Type.change_ignore (Sign.tsig_of thy), Sign.tsig_of thy),
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   274
      (Symset.empty, Consts.change_ignore (Sign.consts_of thy), Sign.consts_of thy),
56140
ed92ce2ac88e just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents: 56139
diff changeset
   275
      Global_Theory.facts_of thy,
56139
b7add947a6ef more frugal recording of changes: join merely requires information from one side;
wenzelm
parents: 56062
diff changeset
   276
      empty_cases);
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   277
);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   278
45650
wenzelm
parents: 45453
diff changeset
   279
fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   280
61811
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   281
fun map_data_result f ctxt =
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   282
  let
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   283
    val Data {mode, syntax, tsig, consts, facts, cases} = Data.get ctxt;
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   284
    val (res, data') = f (mode, syntax, tsig, consts, facts, cases) ||> make_data;
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   285
  in (res, Data.put data' ctxt) end;
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   286
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   287
fun map_data f = snd o map_data_result (pair () o f);
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   288
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   289
fun set_mode mode = map_data (fn (_, syntax, tsig, consts, facts, cases) =>
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   290
  (mode, syntax, tsig, consts, facts, cases));
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   291
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   292
fun map_syntax f =
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   293
  map_data (fn (mode, syntax, tsig, consts, facts, cases) =>
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   294
    (mode, f syntax, tsig, consts, facts, cases));
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   295
59152
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   296
fun map_syntax_idents f ctxt =
60382
8111a4d538ec tuned (see also 66e6c539a36d);
wenzelm
parents: 60381
diff changeset
   297
  let val (opt_idents', syntax') = f (#syntax (rep_data ctxt)) in
59152
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   298
    ctxt
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   299
    |> map_syntax (K syntax')
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   300
    |> (case opt_idents' of NONE => I | SOME idents' => Syntax_Trans.put_idents idents')
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   301
  end;
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
   302
36450
62eaaffe6e47 more systematic naming of tsig operations;
wenzelm
parents: 36448
diff changeset
   303
fun map_tsig f =
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   304
  map_data (fn (mode, syntax, tsig, consts, facts, cases) =>
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   305
    (mode, syntax, f tsig, consts, facts, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   306
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   307
fun map_consts f =
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   308
  map_data (fn (mode, syntax, tsig, consts, facts, cases) =>
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   309
    (mode, syntax, tsig, f consts, facts, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   310
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   311
fun map_local_consts f = map_consts (fn (a, b, c) => (a, f b, c));
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   312
61811
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   313
fun map_facts_result f =
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   314
  map_data_result (fn (mode, syntax, tsig, consts, facts, cases) =>
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   315
    let val (res, facts') = f facts
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   316
    in (res, (mode, syntax, tsig, consts, facts', cases)) end);
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   317
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
   318
fun map_facts f = snd o map_facts_result (pair () o f);
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   319
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   320
fun map_cases f =
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   321
  map_data (fn (mode, syntax, tsig, consts, facts, cases) =>
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   322
    (mode, syntax, tsig, consts, facts, f cases));
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   323
45650
wenzelm
parents: 45453
diff changeset
   324
val get_mode = #mode o rep_data;
28407
f9db1da584ac ContextPosition.report;
wenzelm
parents: 28396
diff changeset
   325
val restore_mode = set_mode o get_mode;
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
   326
val abbrev_mode = get_mode #> (fn Mode {abbrev, ...} => abbrev);
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   327
80074
951c371c1cd9 clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents: 79471
diff changeset
   328
val syntax_of = Local_Syntax.syntax_of o #syntax o rep_data;
33387
acea2f336721 modernized structure Local_Syntax;
wenzelm
parents: 33386
diff changeset
   329
val set_syntax_mode = map_syntax o Local_Syntax.set_mode;
80074
951c371c1cd9 clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents: 79471
diff changeset
   330
val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o #syntax o rep_data;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   331
45650
wenzelm
parents: 45453
diff changeset
   332
val tsig_of = #1 o #tsig o rep_data;
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36450
diff changeset
   333
val set_defsort = map_tsig o apfst o Type.set_defsort;
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
   334
fun default_sort ctxt = the_default (Type.defaultS (tsig_of ctxt)) o Variable.def_sort ctxt;
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61261
diff changeset
   335
fun arity_sorts ctxt = Type.arity_sorts (Context.Proof ctxt) (tsig_of ctxt);
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   336
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   337
val exclude_consts_of = #1 o #consts o rep_data;
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   338
val consts_of = #2 o #consts o rep_data;
45650
wenzelm
parents: 45453
diff changeset
   339
val cases_of = #cases o rep_data;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   340
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   341
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   342
(* naming *)
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   343
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   344
val naming_of = Name_Space.naming_of o Context.Proof;
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   345
val map_naming = Context.proof_map o Name_Space.map_naming;
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   346
val restore_naming = map_naming o K o naming_of;
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   347
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   348
val full_name = Name_Space.full_name o naming_of;
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   349
59886
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   350
val get_scope = Name_Space.get_scope o naming_of;
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   351
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   352
fun new_scope ctxt =
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   353
  let
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   354
    val (scope, naming') = Name_Space.new_scope (naming_of ctxt);
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   355
    val ctxt' = map_naming (K naming') ctxt;
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   356
  in (scope, ctxt') end;
e0dc738eb08c support for explicit scope of private entries;
wenzelm
parents: 59880
diff changeset
   357
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59917
diff changeset
   358
val private_scope = map_naming o Name_Space.private_scope;
59896
e563b0ee0331 tuned signature;
wenzelm
parents: 59886
diff changeset
   359
val private = map_naming o Name_Space.private;
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59970
diff changeset
   360
val qualified_scope = map_naming o Name_Space.qualified_scope;
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59970
diff changeset
   361
val qualified = map_naming o Name_Space.qualified;
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
   362
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   363
val concealed = map_naming Name_Space.concealed;
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   364
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   365
81537
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   366
(* const vs. free *)
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   367
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   368
val const_space = Consts.space_of o consts_of;
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   369
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   370
fun is_const_declared ctxt x =
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   371
  let val space = const_space ctxt
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   372
  in Name_Space.declared space (Name_Space.intern space x) end;
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   373
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   374
fun lookup_free ctxt x =
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   375
  if Variable.is_const ctxt x then NONE
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   376
  else
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   377
    (case Variable.lookup_fixed ctxt x of
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   378
      NONE =>
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   379
        let
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   380
          val is_const = Long_Name.is_qualified x orelse is_const_declared ctxt x;
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   381
          val is_free_declared = is_some (Variable.default_type ctxt x);
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   382
        in if not is_const orelse is_free_declared then SOME x else NONE end
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   383
    | fixed => fixed);
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   384
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   385
fun printable_const ctxt x =
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   386
  is_none (lookup_free ctxt x) andalso
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   387
    not (Symset.member (exclude_consts_of ctxt) x);
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   388
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   389
fun exclude_consts names2 =
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   390
  map_consts (fn (names1, consts, thy_consts) => (Symset.union names1 names2, consts, thy_consts));
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   391
81537
d230683a35fc more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents: 81535
diff changeset
   392
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   393
(* name spaces *)
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   394
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   395
val class_space = Type.class_space o tsig_of;
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   396
val type_space = Type.type_space o tsig_of;
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   397
61265
wenzelm
parents: 61262
diff changeset
   398
fun defs_context ctxt = (ctxt, (const_space ctxt, type_space ctxt));
61261
ddb2da7cb2e4 more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents: 61168
diff changeset
   399
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   400
val intern_class = Name_Space.intern o class_space;
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   401
val intern_type = Name_Space.intern o type_space;
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   402
val intern_const = Name_Space.intern o const_space;
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   403
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   404
fun extern_class ctxt = Name_Space.extern ctxt (class_space ctxt);
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   405
fun extern_type ctxt = Name_Space.extern ctxt (type_space ctxt);
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   406
fun extern_const ctxt = Name_Space.extern_if (printable_const ctxt) ctxt (const_space ctxt);
81539
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   407
fun extern_fixed ctxt x = if Name.is_skolem x then Variable.revert_fixed ctxt x else x;
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   408
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   409
fun extern_entity ctxt =
81596
af21a61dadad clarified signature (see also 2157039256d3);
wenzelm
parents: 81594
diff changeset
   410
  Lexicon.unmark_entity
81539
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   411
   {case_class = extern_class ctxt,
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   412
    case_type = extern_type ctxt,
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   413
    case_const = extern_const ctxt,
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   414
    case_fixed = extern_fixed ctxt,
8e4ca2c87e86 clarified modules;
wenzelm
parents: 81538
diff changeset
   415
    case_default = I};
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   416
55304
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   417
fun markup_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |-> Markup.markup;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   418
fun markup_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |-> Markup.markup;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   419
fun markup_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |-> Markup.markup;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   420
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   421
fun pretty_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |> Pretty.mark_str;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   422
fun pretty_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |> Pretty.mark_str;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   423
fun pretty_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |> Pretty.mark_str;
55ac31bc08a4 more formal markup;
wenzelm
parents: 54883
diff changeset
   424
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42358
diff changeset
   425
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   426
(* theory transfer *)
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   427
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   428
fun transfer_syntax thy ctxt =
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   429
  let
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   430
    val thy_ctxt =
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   431
      Proof_Context.init_global thy
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   432
      |> Context_Position.restore_visible ctxt;
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   433
  in
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   434
    ctxt |>
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   435
    map_syntax (Local_Syntax.rebuild thy_ctxt) |>
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   436
    map_tsig (fn tsig as (local_tsig, global_tsig) =>
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   437
      let val thy_tsig = Sign.tsig_of thy in
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   438
        if Type.eq_tsig (thy_tsig, global_tsig) then tsig
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   439
        else (Type.merge_tsig (Context.Proof ctxt) (local_tsig, thy_tsig), thy_tsig)  (*historic merge order*)
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   440
      end) |>
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   441
    map_consts (fn consts as (exclude_consts, local_consts, global_consts) =>
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   442
      let val thy_consts = Sign.consts_of thy in
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   443
        if Consts.eq_consts (thy_consts, global_consts) then consts
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
   444
        else (exclude_consts, Consts.merge (local_consts, thy_consts), thy_consts)  (*historic merge order*)
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   445
      end)
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
   446
  end;
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   447
33031
b75c35574e04 backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents: 32966
diff changeset
   448
fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   449
57928
f4ff42c5b05b transfer result of Global_Theory.add_thms_dynamic to context stack;
wenzelm
parents: 57887
diff changeset
   450
fun transfer_facts thy =
f4ff42c5b05b transfer result of Global_Theory.add_thms_dynamic to context stack;
wenzelm
parents: 57887
diff changeset
   451
  map_facts (fn local_facts => Facts.merge (Global_Theory.facts_of thy, local_facts));
f4ff42c5b05b transfer result of Global_Theory.add_thms_dynamic to context stack;
wenzelm
parents: 57887
diff changeset
   452
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38328
diff changeset
   453
fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   454
38756
d07959fabde6 renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38328
diff changeset
   455
fun background_theory_result f ctxt =
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   456
  let val (res, thy') = f (theory_of ctxt)
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   457
  in (res, ctxt |> transfer thy') end;
19019
412009243085 added map_theory;
wenzelm
parents: 19001
diff changeset
   458
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   459
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   460
(* hybrid facts *)
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   461
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   462
val facts_of = #facts o rep_data;
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   463
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   464
fun facts_of_fact ctxt name =
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   465
  let
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   466
    val local_facts = facts_of ctxt;
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   467
    val global_facts = Global_Theory.facts_of (theory_of ctxt);
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   468
  in
56143
ed2b660a52a1 more accurate resolution of hybrid facts, which actually changes the sort order of results;
wenzelm
parents: 56141
diff changeset
   469
    if Facts.defined local_facts name
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   470
    then local_facts else global_facts
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   471
  end;
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   472
80329
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   473
fun markup_extern_fact _ "" = ([], "")
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   474
  | markup_extern_fact ctxt name =
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   475
      let
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   476
        val facts = facts_of_fact ctxt name;
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   477
        val (markup, xname) = Facts.markup_extern ctxt facts name;
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   478
        val markups =
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   479
          if Facts.is_dynamic facts name then [markup, Markup.dynamic_fact name]
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   480
          else [markup];
d90a96894644 more robust / permissive;
wenzelm
parents: 80328
diff changeset
   481
      in (markups, xname) end;
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   482
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
   483
fun pretty_thm_name ctxt (name, i) =
81535
db073d1733ab clarified signature: uniform context;
wenzelm
parents: 81532
diff changeset
   484
  Facts.pretty_thm_name ctxt (facts_of_fact ctxt name) (name, i);
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
   485
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
   486
val print_thm_name = Pretty.string_of oo pretty_thm_name;
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80074
diff changeset
   487
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   488
70308
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
   489
(* augment context by implicit term declarations *)
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
   490
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
   491
fun augment t ctxt = ctxt
70733
ce1afe0f3071 clarified modules;
wenzelm
parents: 70730
diff changeset
   492
  |> Variable.add_fixes_implicit t
70364
b2bedb022a75 support for a soft-type system within the Isabelle logical framework;
wenzelm
parents: 70308
diff changeset
   493
  |> Variable.declare_term t
b2bedb022a75 support for a soft-type system within the Isabelle logical framework;
wenzelm
parents: 70308
diff changeset
   494
  |> Soft_Type_System.augment t;
70308
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
   495
7f568724d67e clarified signature;
wenzelm
parents: 70307
diff changeset
   496
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   497
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   498
(** pretty printing **)
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   499
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   500
fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt);
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   501
80326
53f12ab896e6 clarified operations, following pretty_thm_name;
wenzelm
parents: 80307
diff changeset
   502
fun pretty_fact_name ctxt a = Pretty.marks_str (markup_extern_fact ctxt a);
28209
02f3222a392d pretty_fact: extern fact name wrt. the given context, assuming that is the proper one for presentation;
wenzelm
parents: 28087
diff changeset
   503
51583
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50239
diff changeset
   504
fun pretty_fact ctxt =
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50239
diff changeset
   505
  let
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61265
diff changeset
   506
    val pretty_thm = Thm.pretty_thm ctxt;
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61265
diff changeset
   507
    val pretty_thms = map (Thm.pretty_thm_item ctxt);
51583
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50239
diff changeset
   508
  in
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50239
diff changeset
   509
    fn ("", [th]) => pretty_thm th
80328
559909bd7715 clarified signature: more operations;
wenzelm
parents: 80326
diff changeset
   510
     | ("", ths) => Pretty.block0 (Pretty.fbreaks (pretty_thms ths))
80326
53f12ab896e6 clarified operations, following pretty_thm_name;
wenzelm
parents: 80307
diff changeset
   511
     | (a, [th]) =>
53f12ab896e6 clarified operations, following pretty_thm_name;
wenzelm
parents: 80307
diff changeset
   512
        Pretty.block [pretty_fact_name ctxt a, Pretty.str ":", Pretty.brk 1, pretty_thm th]
53f12ab896e6 clarified operations, following pretty_thm_name;
wenzelm
parents: 80307
diff changeset
   513
     | (a, ths) =>
80918
e87d96ac5277 proper fbreaks (amending 53f12ab896e6);
wenzelm
parents: 80910
diff changeset
   514
        Pretty.block (pretty_fact_name ctxt a :: Pretty.fbreaks (Pretty.str ":" :: pretty_thms ths))
51583
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50239
diff changeset
   515
  end;
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   516
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   517
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   518
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   519
(** prepare types **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   520
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   521
(* classes *)
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   522
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   523
fun check_class ctxt (xname, pos) =
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   524
  let
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   525
    val tsig = tsig_of ctxt;
55839
ee71b2687c4b more markup for read_class: imitate Name_Space.check despite lack of Name_Space.table;
wenzelm
parents: 55829
diff changeset
   526
    val class_space = Type.class_space tsig;
ee71b2687c4b more markup for read_class: imitate Name_Space.check despite lack of Name_Space.table;
wenzelm
parents: 55829
diff changeset
   527
56007
wenzelm
parents: 56002
diff changeset
   528
    val name = Type.cert_class tsig (Name_Space.intern class_space xname)
55839
ee71b2687c4b more markup for read_class: imitate Name_Space.check despite lack of Name_Space.table;
wenzelm
parents: 55829
diff changeset
   529
      handle TYPE (msg, _, _) =>
55840
2982d233d798 consider completion report as part of error message -- less stateful, may get handled;
wenzelm
parents: 55839
diff changeset
   530
        error (msg ^ Position.here pos ^
69289
bf6937af7fe8 clarified signature;
wenzelm
parents: 69185
diff changeset
   531
          Completion.markup_report
bf6937af7fe8 clarified signature;
wenzelm
parents: 69185
diff changeset
   532
            [Name_Space.completion (Context.Proof ctxt) class_space (K true) (xname, pos)]);
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   533
    val reports =
55923
wenzelm
parents: 55922
diff changeset
   534
      if Context_Position.is_reported ctxt pos
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   535
      then [(pos, Name_Space.markup class_space name), (pos, Markup.tclass)] else [];
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   536
  in (name, reports) end;
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   537
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   538
fun read_class ctxt text =
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   539
  let
59795
d453c69596cc clarified input source;
wenzelm
parents: 59621
diff changeset
   540
    val source = Syntax.read_input text;
69349
7cef9e386ffe more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents: 69289
diff changeset
   541
    val (c, reports) = check_class ctxt (Input.source_content source);
71674
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   542
    val _ = Context_Position.reports ctxt reports;
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   543
  in c end;
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   544
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   545
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   546
(* types *)
24277
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   547
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   548
fun read_typ_mode mode ctxt s =
24486
1dbf377c2e9a moved type_mode to type.ML;
wenzelm
parents: 24388
diff changeset
   549
  Syntax.read_typ (Type.set_mode mode ctxt) s;
24277
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   550
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   551
val read_typ = read_typ_mode Type.mode_default;
24277
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   552
val read_typ_syntax = read_typ_mode Type.mode_syntax;
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   553
val read_typ_abbrev = read_typ_mode Type.mode_abbrev;
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   554
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   555
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   556
fun cert_typ_mode mode ctxt T =
79455
d7f32f04bd13 clarified signature;
wenzelm
parents: 79452
diff changeset
   557
  Type.certify_typ mode (tsig_of ctxt) T
24277
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   558
    handle TYPE (msg, _, _) => error msg;
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   559
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   560
val cert_typ = cert_typ_mode Type.mode_default;
24277
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   561
val cert_typ_syntax = cert_typ_mode Type.mode_syntax;
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   562
val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev;
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   563
6442fde2daaa added implicit type mode (cf. Type.mode);
wenzelm
parents: 24012
diff changeset
   564
24388
cf24894b81ff added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents: 24371
diff changeset
   565
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   566
(** prepare terms and propositions **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   567
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   568
(* inferred types of parameters *)
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   569
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   570
fun infer_type ctxt x =
36505
79c1d2bbe5a9 ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents: 36503
diff changeset
   571
  Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt)) (Free x));
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   572
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   573
fun inferred_param x ctxt =
60407
53ef7b78e78a tuned signature;
wenzelm
parents: 60401
diff changeset
   574
  let val p = (x, infer_type ctxt (x, dummyT))
53ef7b78e78a tuned signature;
wenzelm
parents: 60401
diff changeset
   575
  in (p, ctxt |> Variable.declare_term (Free p)) end;
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   576
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   577
fun inferred_fixes ctxt =
60407
53ef7b78e78a tuned signature;
wenzelm
parents: 60401
diff changeset
   578
  fold_map inferred_param (map #2 (Variable.dest_fixes ctxt)) ctxt;
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   579
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   580
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   581
(* type names *)
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   582
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
   583
fun check_type_name {proper, strict} ctxt (c, pos) =
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   584
  if Lexicon.is_tid c then
56040
98d466e6de8a tuned message;
wenzelm
parents: 56025
diff changeset
   585
    if proper then error ("Not a type constructor: " ^ quote c ^ Position.here pos)
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   586
    else
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   587
      let
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   588
        val reports =
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   589
          if Context_Position.is_reported ctxt pos
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   590
          then [(pos, Markup.tfree)] else [];
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   591
      in (TFree (c, default_sort ctxt (c, ~1)), reports) end
55922
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   592
  else
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   593
    let
710bc66f432c more markup for inner syntax class/type names (notably for completion);
wenzelm
parents: 55843
diff changeset
   594
      val ((d, reports), decl) = Type.check_decl (Context.Proof ctxt) (tsig_of ctxt) (c, pos);
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   595
      val reports' =
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   596
        if Context_Position.is_reported ctxt pos
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   597
        then reports @ [(pos, Markup.tconst)] else reports;
79470
9fcf73580c62 clarified signature;
wenzelm
parents: 79469
diff changeset
   598
      val _ =
9fcf73580c62 clarified signature;
wenzelm
parents: 79469
diff changeset
   599
        if strict andalso not (Type.decl_logical decl)
9fcf73580c62 clarified signature;
wenzelm
parents: 79469
diff changeset
   600
        then error ("Bad type name: " ^ quote d ^ Position.here pos) else ();
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   601
    in (Type (d, replicate (Type.decl_args decl) dummyT), reports') end;
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   602
71674
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   603
fun read_type_name flags ctxt text =
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   604
  let
59795
d453c69596cc clarified input source;
wenzelm
parents: 59621
diff changeset
   605
    val source = Syntax.read_input text;
71674
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   606
    val (T, reports) = check_type_name flags ctxt (Input.source_content source);
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   607
    val _ = Context_Position.reports ctxt reports;
55951
c07d184aebe9 tuned signature -- more uniform check_type_name/read_type_name;
wenzelm
parents: 55950
diff changeset
   608
  in T end;
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   609
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   610
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   611
(* constant names *)
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   612
55956
94d384d621b0 reject internal term names outright, and complete consts instead;
wenzelm
parents: 55955
diff changeset
   613
fun consts_completion_message ctxt (c, ps) =
94d384d621b0 reject internal term names outright, and complete consts instead;
wenzelm
parents: 55955
diff changeset
   614
  ps |> map (fn pos =>
69289
bf6937af7fe8 clarified signature;
wenzelm
parents: 69185
diff changeset
   615
    Name_Space.completion (Context.Proof ctxt) (Consts.space_of (consts_of ctxt)) (K true) (c, pos))
bf6937af7fe8 clarified signature;
wenzelm
parents: 69185
diff changeset
   616
  |> Completion.markup_report;
55956
94d384d621b0 reject internal term names outright, and complete consts instead;
wenzelm
parents: 55955
diff changeset
   617
56002
2028467b4df4 tuned signature;
wenzelm
parents: 55959
diff changeset
   618
fun check_const {proper, strict} ctxt (c, ps) =
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   619
  let
55956
94d384d621b0 reject internal term names outright, and complete consts instead;
wenzelm
parents: 55955
diff changeset
   620
    val _ =
55959
c3b458435f4f more decisive commitment to get_free vs. the_const;
wenzelm
parents: 55956
diff changeset
   621
      Name.reject_internal (c, ps) handle ERROR msg =>
c3b458435f4f more decisive commitment to get_free vs. the_const;
wenzelm
parents: 55956
diff changeset
   622
        error (msg ^ consts_completion_message ctxt (c, ps));
c3b458435f4f more decisive commitment to get_free vs. the_const;
wenzelm
parents: 55956
diff changeset
   623
    fun err msg = error (msg ^ Position.here_list ps);
81532
wenzelm
parents: 81253
diff changeset
   624
wenzelm
parents: 81253
diff changeset
   625
    val fixed = if proper then NONE else Variable.lookup_fixed ctxt c;
wenzelm
parents: 81253
diff changeset
   626
    val const = Variable.lookup_const ctxt c;
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   627
    val consts = consts_of ctxt;
81532
wenzelm
parents: 81253
diff changeset
   628
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   629
    val is_reported = Context_Position.is_reported ctxt;
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   630
    val (t, reports) =
81532
wenzelm
parents: 81253
diff changeset
   631
      if is_some fixed andalso is_none const then
wenzelm
parents: 81253
diff changeset
   632
        let
wenzelm
parents: 81253
diff changeset
   633
          val x = the fixed;
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   634
          val reports =
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   635
            ps |> maps (fn pos =>
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   636
              if is_reported pos then [(pos, Markup.name x (Variable.markup ctxt x))]
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   637
              else []);
81532
wenzelm
parents: 81253
diff changeset
   638
        in (Free (x, infer_type ctxt (x, dummyT)), reports) end
wenzelm
parents: 81253
diff changeset
   639
      else if is_some const then
wenzelm
parents: 81253
diff changeset
   640
        let
wenzelm
parents: 81253
diff changeset
   641
          val d = the const;
wenzelm
parents: 81253
diff changeset
   642
          val T = Consts.type_scheme consts d handle TYPE (msg, _, _) => err msg;
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   643
          val reports =
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   644
            ps |> maps (fn pos =>
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   645
              if is_reported pos then
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   646
                [(pos, Name_Space.markup (Consts.space_of consts) d), (pos, Markup.const)]
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   647
              else []);
81532
wenzelm
parents: 81253
diff changeset
   648
        in (Const (d, T), reports) end
81565
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   649
      else
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   650
        let
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   651
          val (d, reports1) = Consts.check_const (Context.Proof ctxt) consts (c, ps);
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   652
          val reports2 =
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   653
            ps |> maps (fn pos => if is_reported pos then [(pos, Markup.const)] else []);
bf19ea589f99 more markup (without rendering): class, type constructor, term constant --- similar to free, bound, etc.;
wenzelm
parents: 81558
diff changeset
   654
        in (d, reports1 @ reports2) end;
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   655
    val _ =
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   656
      (case (strict, t) of
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   657
        (true, Const (d, _)) =>
79463
7d10708bbc32 clarified signature;
wenzelm
parents: 79455
diff changeset
   658
          (ignore (Consts.the_const_type consts d) handle TYPE (msg, _, _) => err msg)
55954
a29aefc88c8d more uniform check_const/read_const;
wenzelm
parents: 55951
diff changeset
   659
      | _ => ());
55950
3bb5c7179234 clarified treatment of consts -- prefer value-oriented reports;
wenzelm
parents: 55949
diff changeset
   660
  in (t, reports) end;
55842
ea540323c444 tuned source structure;
wenzelm
parents: 55841
diff changeset
   661
71674
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   662
fun read_const flags ctxt text =
55950
3bb5c7179234 clarified treatment of consts -- prefer value-oriented reports;
wenzelm
parents: 55949
diff changeset
   663
  let
59795
d453c69596cc clarified input source;
wenzelm
parents: 59621
diff changeset
   664
    val source = Syntax.read_input text;
69349
7cef9e386ffe more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents: 69289
diff changeset
   665
    val (c, pos) = Input.source_content source;
71674
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   666
    val (t, reports) = check_const flags ctxt (c, [pos]);
48ff625687f5 more accurate context position reports;
wenzelm
parents: 70734
diff changeset
   667
    val _ = Context_Position.reports ctxt reports;
55950
3bb5c7179234 clarified treatment of consts -- prefer value-oriented reports;
wenzelm
parents: 55949
diff changeset
   668
  in t end;
25328
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   669
f71105742e2c discontinued ProofContext.read_prop_legacy;
wenzelm
parents: 25319
diff changeset
   670
46922
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   671
(* type arities *)
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   672
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   673
local
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   674
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   675
fun prep_arity prep_tycon prep_sort ctxt (t, Ss, S) =
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   676
  let val arity = (prep_tycon ctxt t, map (prep_sort ctxt) Ss, prep_sort ctxt S)
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61261
diff changeset
   677
  in Type.add_arity (Context.Proof ctxt) arity (tsig_of ctxt); arity end;
46922
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   678
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   679
in
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   680
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   681
val read_arity =
80632
3a196e63a80d tuned signature: more operations;
wenzelm
parents: 80329
diff changeset
   682
  prep_arity (dest_Type_name oo read_type_name {proper = true, strict = true}) Syntax.read_sort;
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
   683
46922
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   684
val cert_arity = prep_arity (K I) (Type.cert_sort o tsig_of);
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   685
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   686
end;
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   687
3717f3878714 source positions for locale and class expressions;
wenzelm
parents: 46849
diff changeset
   688
24684
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   689
(* read_term *)
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   690
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   691
fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt);
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   692
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   693
val read_term_pattern   = read_term_mode mode_pattern;
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   694
val read_term_schematic = read_term_mode mode_schematic;
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   695
val read_term_abbrev    = read_term_mode mode_abbrev;
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   696
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   697
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   698
(* local abbreviations *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   699
24501
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   700
local
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   701
61262
7bd1eb4b056e tuned signature: eliminated pointless type Context.pretty;
wenzelm
parents: 61261
diff changeset
   702
fun certify_consts ctxt =
79471
593fdddc6d98 clarified signature;
wenzelm
parents: 79470
diff changeset
   703
  Consts.certify {normalize = not (abbrev_mode ctxt)}
79451
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   704
    (Context.Proof ctxt) (tsig_of ctxt) (consts_of ctxt);
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   705
38979
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   706
fun expand_binds ctxt =
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   707
  let
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   708
    val Mode {pattern, schematic, ...} = get_mode ctxt;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   709
38979
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   710
    fun reject_schematic (t as Var _) =
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   711
          error ("Unbound schematic variable: " ^ Syntax.string_of_term ctxt t)
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   712
      | reject_schematic (Abs (_, _, t)) = reject_schematic t
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   713
      | reject_schematic (t $ u) = (reject_schematic t; reject_schematic u)
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   714
      | reject_schematic _ = ();
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   715
  in
24495
eab1b52a47d0 added join_mode;
wenzelm
parents: 24486
diff changeset
   716
    if pattern then I
eab1b52a47d0 added join_mode;
wenzelm
parents: 24486
diff changeset
   717
    else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic)
eab1b52a47d0 added join_mode;
wenzelm
parents: 24486
diff changeset
   718
  end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   719
24501
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   720
in
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   721
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   722
fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt;
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   723
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   724
end;
7a2b20145888 export various inner syntax modes;
wenzelm
parents: 24495
diff changeset
   725
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
   726
val show_abbrevs = Config.declare_bool ("show_abbrevs", \<^here>) (K true);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   727
81220
3d09d6f4c5b1 clarified signature;
wenzelm
parents: 81116
diff changeset
   728
fun contract_abbrevs ctxt tm =
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   729
  let
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   730
    val thy = theory_of ctxt;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   731
    val consts = consts_of ctxt;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   732
    val Mode {abbrev, ...} = get_mode ctxt;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   733
  in
81220
3d09d6f4c5b1 clarified signature;
wenzelm
parents: 81116
diff changeset
   734
    if abbrev orelse not (Config.get ctxt show_abbrevs) orelse not (can Term.type_of tm) then tm
81222
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   735
    else
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   736
      let
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   737
        val inst = #1 (Variable.import_inst true [tm] ctxt);
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   738
        val nets = Consts.revert_abbrevs consts (print_mode_value () @ [""]);
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   739
        val rew = Option.map #1 oo Pattern.match_rew thy;
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   740
        fun match_abbrev t = get_first (fn net => get_first (rew t) (Item_Net.retrieve net t)) nets;
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   741
      in
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   742
        Term_Subst.instantiate inst tm
81253
bbed9f218158 prefer rewrite_term_yoyo for improved performance and occasionally better results (conforming to Ast.normalize);
wenzelm
parents: 81236
diff changeset
   743
        |> Pattern.rewrite_term_yoyo thy [] [match_abbrev]
81222
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   744
        |> Term_Subst.instantiate_frees (Variable.import_inst_revert inst)
b61abd1e5027 notable performance tuning (amending a59d9b81be24 and 8976c5bc9e97): avoid costly could_beta_eta_contract, which traverses the whole term;
wenzelm
parents: 81221
diff changeset
   745
      end
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   746
  end;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   747
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   748
24518
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   749
(* patterns *)
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   750
32003
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   751
fun prepare_patternT ctxt T =
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   752
  let
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   753
    val Mode {pattern, schematic, ...} = get_mode ctxt;
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   754
    val _ =
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   755
      pattern orelse schematic orelse
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   756
        T |> Term.exists_subtype
38979
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   757
          (fn T as TVar (xi, _) =>
37145
01aa36932739 renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
wenzelm
parents: 36610
diff changeset
   758
            not (Type_Infer.is_param xi) andalso
38979
60dbf0b3f6c7 prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents: 38756
diff changeset
   759
              error ("Illegal schematic type variable: " ^ Syntax.string_of_typ ctxt T)
32003
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   760
          | _ => false)
befec6450fd6 tuned prepare_patternT: Term.exists_subtype;
wenzelm
parents: 30815
diff changeset
   761
  in T end;
24518
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   762
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   763
24505
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   764
local
6550
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   765
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
   766
val dummies = Config.declare_bool ("Proof_Context.dummies", \<^here>) (K false);
27259
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   767
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   768
fun check_dummies ctxt t =
39508
dfacdb01e1ec simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents: 39507
diff changeset
   769
  if Config.get ctxt dummies then t
27259
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   770
  else Term.no_dummy_patterns t handle TERM _ => error "Illegal dummy pattern(s) in term";
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   771
24767
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
   772
fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1);
6762
a9a515a43ae0 read_term/prop_pat: do not freeze;
wenzelm
parents: 6721
diff changeset
   773
27259
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   774
in
6550
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   775
39508
dfacdb01e1ec simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents: 39507
diff changeset
   776
val allow_dummies = Config.put dummies true;
24505
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   777
24684
80da599dea37 added read_term_pattern/schematic/abbrev;
wenzelm
parents: 24675
diff changeset
   778
fun prepare_patterns ctxt =
24518
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   779
  let val Mode {pattern, ...} = get_mode ctxt in
62958
b41c1cb5e251 Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents: 62773
diff changeset
   780
    Type_Infer.fixate ctxt pattern #>
24767
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
   781
    pattern ? Variable.polymorphic ctxt #>
79437
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   782
    (Same.commit o Same.map o Term.map_types_same)
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   783
      (Same.function_eq (op =) (prepare_patternT ctxt)) #>
27259
6e71abb8c994 export transfer_syntax;
wenzelm
parents: 27195
diff changeset
   784
    (if pattern then prepare_dummies else map (check_dummies ctxt))
24505
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   785
  end;
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   786
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   787
end;
9e6d91f8bb73 reject_vars: accept type-inference params;
wenzelm
parents: 24501
diff changeset
   788
6550
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   789
42250
cc5ac538f991 eliminated odd object-oriented type_context/term_context;
wenzelm
parents: 42242
diff changeset
   790
(* sort constraints *)
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
   791
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   792
local
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   793
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   794
fun prepare_sorts_env ctxt tys =
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
   795
  let
35669
a91c7ed801b8 added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents: 35616
diff changeset
   796
    val tsig = tsig_of ctxt;
45453
304437f43869 more scalable Proof_Context.prepare_sorts;
wenzelm
parents: 45429
diff changeset
   797
    val defaultS = Type.defaultS tsig;
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
   798
53673
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   799
    val dummy_var = ("'_dummy_", ~1);
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   800
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   801
    fun constraint (xi, raw_S) env =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   802
      let val (ps, S) = Term_Position.decode_positionS raw_S in
53673
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   803
        if xi = dummy_var orelse S = dummyS then env
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   804
        else
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   805
          Vartab.insert (op =) (xi, Type.minimize_sort tsig S) env
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   806
            handle Vartab.DUP _ =>
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   807
              error ("Inconsistent sort constraints for type variable " ^
81558
b57996a0688c clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents: 81543
diff changeset
   808
                quote (Term.string_of_vname' xi) ^ Position.here_list (map #pos ps))
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   809
      end;
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   810
45453
304437f43869 more scalable Proof_Context.prepare_sorts;
wenzelm
parents: 45429
diff changeset
   811
    val env =
74232
1091880266e5 clarified signature;
wenzelm
parents: 73690
diff changeset
   812
      Vartab.build (tys |> (fold o fold_atyps)
45453
304437f43869 more scalable Proof_Context.prepare_sorts;
wenzelm
parents: 45429
diff changeset
   813
        (fn TFree (x, S) => constraint ((x, ~1), S)
304437f43869 more scalable Proof_Context.prepare_sorts;
wenzelm
parents: 45429
diff changeset
   814
          | TVar v => constraint v
74232
1091880266e5 clarified signature;
wenzelm
parents: 73690
diff changeset
   815
          | _ => I));
36152
34d1ce2d746d get_sort: suppress dummyS from input;
wenzelm
parents: 35680
diff changeset
   816
53673
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   817
    fun get_sort xi raw_S =
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   818
      if xi = dummy_var then
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   819
        Type.minimize_sort tsig (#2 (Term_Position.decode_positionS raw_S))
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   820
      else
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   821
        (case (Vartab.lookup env xi, Variable.def_sort ctxt xi) of
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   822
          (NONE, NONE) => defaultS
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   823
        | (NONE, SOME S) => S
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   824
        | (SOME S, NONE) => S
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   825
        | (SOME S, SOME S') =>
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   826
            if Type.eq_sort tsig (S, S') then S'
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   827
            else
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   828
              error ("Sort constraint " ^ Syntax.string_of_sort ctxt S ^
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   829
                " inconsistent with default " ^ Syntax.string_of_sort ctxt S' ^
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   830
                " for type variable " ^ quote (Term.string_of_vname' xi)));
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   831
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   832
    fun add_report S pos reports =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   833
      if Position.is_reported pos andalso not (AList.defined (op =) reports pos) then
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 50126
diff changeset
   834
        (pos, Position.reported_text pos Markup.sorting (Syntax.string_of_sort ctxt S)) :: reports
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   835
      else reports;
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   836
49691
74ad6ecf2af2 more error positions;
wenzelm
parents: 49685
diff changeset
   837
    fun get_sort_reports xi raw_S =
74ad6ecf2af2 more error positions;
wenzelm
parents: 49685
diff changeset
   838
      let
81558
b57996a0688c clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents: 81543
diff changeset
   839
        val ps = map #pos (#1 (Term_Position.decode_positionS raw_S));
53673
bcfd16f65014 treat all dummy type variables separately (in contrast to fca432074fb2);
wenzelm
parents: 53380
diff changeset
   840
        val S = get_sort xi raw_S handle ERROR msg => error (msg ^ Position.here_list ps);
49691
74ad6ecf2af2 more error positions;
wenzelm
parents: 49685
diff changeset
   841
      in fold (add_report S) ps end;
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   842
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   843
    val reports =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   844
      (fold o fold_atyps)
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   845
        (fn T =>
81232
9867b5cf3f7a more concise representation of term positions;
wenzelm
parents: 81222
diff changeset
   846
          if Term_Position.detect_positionT T then I
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   847
          else
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   848
            (case T of
49691
74ad6ecf2af2 more error positions;
wenzelm
parents: 49685
diff changeset
   849
              TFree (x, raw_S) => get_sort_reports (x, ~1) raw_S
74ad6ecf2af2 more error positions;
wenzelm
parents: 49685
diff changeset
   850
            | TVar (xi, raw_S) => get_sort_reports xi raw_S
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   851
            | _ => I)) tys [];
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   852
56333
38f1422ef473 support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents: 56294
diff changeset
   853
  in (map #2 reports, get_sort) end;
27286
2ea20e5fdf16 renamed is_abbrev_mode to abbrev_mode;
wenzelm
parents: 27264
diff changeset
   854
79437
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   855
fun replace_sortsT_same get_sort =
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   856
  Term.map_atyps_same
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   857
    (fn T =>
81232
9867b5cf3f7a more concise representation of term positions;
wenzelm
parents: 81222
diff changeset
   858
      if Term_Position.detect_positionT T then raise Same.SAME
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   859
      else
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   860
        (case T of
79437
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   861
          TFree (x, raw_S) => TFree (x, Same.function_eq (op =) (get_sort (x, ~1)) raw_S)
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   862
        | TVar (xi, raw_S) => TVar (xi, Same.function_eq (op =) (get_sort xi) raw_S)
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   863
        | _ => raise Same.SAME));
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   864
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   865
in
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   866
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   867
fun prepare_sortsT ctxt tys =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   868
  let val (sorting_report, get_sort) = prepare_sorts_env ctxt tys
79437
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   869
  in (sorting_report, (Same.commit o Same.map) (replace_sortsT_same get_sort) tys) end;
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   870
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   871
fun prepare_sorts ctxt tms =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   872
  let
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   873
    val tys = rev ((fold o fold_types) cons tms []);
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   874
    val (sorting_report, get_sort) = prepare_sorts_env ctxt tys;
79437
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   875
    val tms' = (Same.commit o Same.map o Term.map_types_same) (replace_sortsT_same get_sort) tms;
848073637388 minor performance tuning;
wenzelm
parents: 79373
diff changeset
   876
  in (sorting_report, tms') end;
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   877
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   878
fun check_tfree ctxt v =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   879
  let
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   880
    val (sorting_report, [TFree a]) = prepare_sortsT ctxt [TFree v];
71675
55cb4271858b less redundant markup reports;
wenzelm
parents: 71674
diff changeset
   881
    val _ = if Context_Position.reports_enabled ctxt then Output.report sorting_report else ();
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   882
  in a end;
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   883
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 48992
diff changeset
   884
end;
36152
34d1ce2d746d get_sort: suppress dummyS from input;
wenzelm
parents: 35680
diff changeset
   885
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   886
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   887
(* certify terms *)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   888
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   889
local
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   890
79451
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   891
fun cert_flags flags ctxt t =
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   892
  let val t' = expand_abbrevs ctxt t in
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   893
    #1 (Sign.certify_flags flags (Context.Proof ctxt) (consts_of ctxt) (theory_of ctxt) t')
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   894
      handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg
ef867bf3e6c9 clarified signature;
wenzelm
parents: 79449
diff changeset
   895
  end;
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   896
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   897
in
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   898
79471
593fdddc6d98 clarified signature;
wenzelm
parents: 79470
diff changeset
   899
val cert_term = cert_flags {prop = false, normalize = false};
593fdddc6d98 clarified signature;
wenzelm
parents: 79470
diff changeset
   900
val cert_prop = cert_flags {prop = true, normalize = false};
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   901
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   902
end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   903
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   904
42405
13ecdb3057d8 split Type_Infer into early and late part, after Proof_Context;
wenzelm
parents: 42402
diff changeset
   905
(* check/uncheck *)
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   906
35616
b342390d296f provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents: 35429
diff changeset
   907
fun def_type ctxt =
b342390d296f provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents: 35429
diff changeset
   908
  let val Mode {pattern, ...} = get_mode ctxt
b342390d296f provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents: 35429
diff changeset
   909
  in Variable.def_type ctxt pattern end;
b342390d296f provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents: 35429
diff changeset
   910
24518
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   911
fun standard_typ_check ctxt =
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   912
  map (cert_typ_mode (Type.get_mode ctxt) ctxt #> prepare_patternT ctxt);
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
   913
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   914
val standard_term_check_finish = prepare_patterns;
24518
4dd086997bab removed unused join_mode;
wenzelm
parents: 24511
diff changeset
   915
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   916
fun standard_term_uncheck ctxt = map (contract_abbrevs ctxt);
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   917
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   918
9553
c2e3773475b6 norm_hhf results;
wenzelm
parents: 9540
diff changeset
   919
21610
52c0d3280798 removed obsolete (export_)standard;
wenzelm
parents: 21600
diff changeset
   920
(** export results **)
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   921
78086
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   922
fun export_ goal inner outer =
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   923
  map (Assumption.export_ goal inner outer) #>
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   924
  Variable.export inner outer;
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   925
78086
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   926
val export = export_{goal = false};
5edd5b12017d tuned signature;
wenzelm
parents: 78084
diff changeset
   927
val export_goal = export_{goal = true};
12704
7bffaadc581e export_single;
wenzelm
parents: 12576
diff changeset
   928
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   929
fun export_morphism inner outer =
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   930
  Assumption.export_morphism inner outer $>
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   931
  Variable.export_morphism inner outer;
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   932
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   933
15758
07e382399a96 binds/thms: do not store options, but delete from table;
wenzelm
parents: 15750
diff changeset
   934
30757
2d2076300185 replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents: 30723
diff changeset
   935
(** term bindings **)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   936
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   937
(* auto bindings *)
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   938
60408
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
   939
fun auto_bind f props ctxt = fold Variable.maybe_bind_term (f ctxt props) ctxt;
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   940
33386
ff29d1549aca modernized structure AutoBind;
wenzelm
parents: 33383
diff changeset
   941
val auto_bind_goal = auto_bind Auto_Bind.goal;
ff29d1549aca modernized structure AutoBind;
wenzelm
parents: 33383
diff changeset
   942
val auto_bind_facts = auto_bind Auto_Bind.facts;
7925
8c50b68b890b warn_extra_tfrees (after declare_term);
wenzelm
parents: 7679
diff changeset
   943
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   944
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   945
(* match bindings *)
60408
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
   946
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   947
fun simult_matches ctxt (t, pats) =
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   948
  (case Seq.pull (Unify.matchers (Context.Proof ctxt) (map (rpair t) pats)) of
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   949
    NONE => error "Pattern match failed!"
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   950
  | SOME (env, _) => Vartab.fold (fn (v, (_, t)) => cons (v, t)) (Envir.term_env env) []);
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   951
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   952
fun maybe_bind_term (xi, t) ctxt =
60408
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
   953
  ctxt
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
   954
  |> Variable.maybe_bind_term (xi, Option.map (cert_term (set_mode mode_default ctxt)) t);
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
   955
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
   956
val bind_term = maybe_bind_term o apsnd SOME;
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   957
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   958
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   959
(* propositions with patterns *)
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   960
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   961
local
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   962
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   963
fun prep_propp prep_props ctxt raw_args =
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   964
  let
60386
16b5b1b9dd02 tuned signature;
wenzelm
parents: 60383
diff changeset
   965
    val props = prep_props ctxt (maps (map fst) raw_args);
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   966
    val props_ctxt = fold Variable.declare_term props ctxt;
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   967
    val patss = maps (map (prep_props (set_mode mode_pattern props_ctxt) o snd)) raw_args;
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   968
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60386
diff changeset
   969
    val propps = unflat raw_args (props ~~ patss);
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   970
    val binds = (maps o maps) (simult_matches props_ctxt) propps;
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   971
  in (map (map fst) propps, binds) end;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   972
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   973
in
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   974
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   975
val cert_propp = prep_propp (map o cert_prop);
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
   976
val read_propp = prep_propp Syntax.read_props;
6789
0e5a965de17a auto_bind_goal, auto_bind_facts;
wenzelm
parents: 6762
diff changeset
   977
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   978
end;
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   979
6789
0e5a965de17a auto_bind_goal, auto_bind_facts;
wenzelm
parents: 6762
diff changeset
   980
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   981
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   982
(** theorems **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   983
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   984
(* fact_tac *)
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   985
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
   986
local
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
   987
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   988
fun comp_hhf_tac ctxt th i st =
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   989
  PRIMSEQ (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true}
60367
e201bd8973db clarified context;
wenzelm
parents: 60242
diff changeset
   990
    (false, Drule.lift_all ctxt (Thm.cprem_of st i) th, 0) i) st;
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
   991
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   992
fun comp_incr_tac _ [] _ = no_tac
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   993
  | comp_incr_tac ctxt (th :: ths) i =
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   994
      (fn st => comp_hhf_tac ctxt (Drule.incr_indexes st th) i st) APPEND
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58668
diff changeset
   995
      comp_incr_tac ctxt ths i;
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
   996
60489
bfd9b7302a82 vacuous fact `TERM x`;
wenzelm
parents: 60469
diff changeset
   997
val vacuous_facts = [Drule.termI];
bfd9b7302a82 vacuous fact `TERM x`;
wenzelm
parents: 60469
diff changeset
   998
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
   999
in
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
  1000
27867
6e6a159671d4 retrieve_thms: transfer fact position to result;
wenzelm
parents: 27828
diff changeset
  1001
fun potential_facts ctxt prop =
60489
bfd9b7302a82 vacuous fact `TERM x`;
wenzelm
parents: 60469
diff changeset
  1002
  let
bfd9b7302a82 vacuous fact `TERM x`;
wenzelm
parents: 60469
diff changeset
  1003
    val body = Term.strip_all_body prop;
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1004
    val vacuous =
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1005
      filter (fn th => Term.could_unify (body, Thm.concl_of th)) vacuous_facts
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1006
      |> map (rpair Position.none);
60489
bfd9b7302a82 vacuous fact `TERM x`;
wenzelm
parents: 60469
diff changeset
  1007
  in Facts.could_unify (facts_of ctxt) body @ vacuous end;
27867
6e6a159671d4 retrieve_thms: transfer fact position to result;
wenzelm
parents: 27828
diff changeset
  1008
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1009
fun fact_tac ctxt facts = Goal.norm_hhf_tac ctxt THEN' comp_incr_tac ctxt facts;
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1010
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1011
fun some_fact_tac ctxt =
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1012
  SUBGOAL (fn (goal, i) => fact_tac ctxt (map #1 (potential_facts ctxt goal)) i);
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
  1013
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
  1014
end;
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 52156
diff changeset
  1015
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
  1016
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1017
(* lookup facts *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1018
61902
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1019
fun lookup_fact ctxt name =
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1020
  let
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1021
    val context = Context.Proof ctxt;
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1022
    val thy = Proof_Context.theory_of ctxt;
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1023
  in
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1024
    (case Facts.lookup context (facts_of ctxt) name of
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1025
      NONE => Facts.lookup context (Global_Theory.facts_of thy) name
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1026
    | some => some)
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1027
  end;
4d162c237e48 more accurate lookup of dynamic facts;
wenzelm
parents: 61811
diff changeset
  1028
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1029
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1030
(* retrieve facts *)
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1031
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
  1032
val dynamic_facts_dummy = Config.declare_bool ("dynamic_facts_dummy_", \<^here>) (K false);
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1033
26361
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1034
local
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1035
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1036
fun retrieve_global context =
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1037
  Facts.retrieve context (Global_Theory.facts_of (Context.theory_of context));
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1038
56156
47015872e795 clarified retrieve_generic: local error takes precedence, which is relevant for completion;
wenzelm
parents: 56155
diff changeset
  1039
fun retrieve_generic (context as Context.Proof ctxt) arg =
47015872e795 clarified retrieve_generic: local error takes precedence, which is relevant for completion;
wenzelm
parents: 56155
diff changeset
  1040
      (Facts.retrieve context (facts_of ctxt) arg handle ERROR local_msg =>
47015872e795 clarified retrieve_generic: local error takes precedence, which is relevant for completion;
wenzelm
parents: 56155
diff changeset
  1041
        (retrieve_global context arg handle ERROR _ => error local_msg))
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1042
  | retrieve_generic context arg = retrieve_global context arg;
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1043
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1044
fun retrieve pick context (Facts.Fact s) =
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
  1045
      let
56140
ed92ce2ac88e just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents: 56139
diff changeset
  1046
        val ctxt = Context.the_proof context;
59795
d453c69596cc clarified input source;
wenzelm
parents: 59621
diff changeset
  1047
        val pos = Syntax.read_input_pos s;
42502
13b41fb77649 literal facts `prop` may contain dummy patterns;
wenzelm
parents: 42501
diff changeset
  1048
        val prop =
13b41fb77649 literal facts `prop` may contain dummy patterns;
wenzelm
parents: 42501
diff changeset
  1049
          Syntax.read_prop (ctxt |> set_mode mode_default |> allow_dummies) s
24511
69d270cc7e4f removed obsolete read/cert variations (cf. Syntax.read/check);
wenzelm
parents: 24505
diff changeset
  1050
          |> singleton (Variable.polymorphic ctxt);
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1051
        fun err ps msg =
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1052
          error (msg ^ Position.here_list (pos :: ps) ^ ":\n" ^ Syntax.string_of_term ctxt prop);
27867
6e6a159671d4 retrieve_thms: transfer fact position to result;
wenzelm
parents: 27828
diff changeset
  1053
42502
13b41fb77649 literal facts `prop` may contain dummy patterns;
wenzelm
parents: 42501
diff changeset
  1054
        val (prop', _) = Term.replace_dummy_patterns prop (Variable.maxidx_of ctxt + 1);
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1055
        fun prove th = Goal.prove ctxt [] [] prop' (K (ALLGOALS (fact_tac ctxt [th])));
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1056
        val results = map_filter (try (apfst prove)) (potential_facts ctxt prop');
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1057
        val (thm, thm_pos) =
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1058
          (case distinct (eq_fst Thm.eq_thm_prop) results of
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1059
            [res] => res
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1060
          | [] => err [] "Failed to retrieve literal fact"
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1061
          | dups => err (distinct (op =) (map #2 dups)) "Ambiguous specification of literal fact");
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1062
71910
f8b0271cc744 tuned signature;
wenzelm
parents: 71675
diff changeset
  1063
        val markup = Position.entity_markup Markup.literal_factN ("", thm_pos);
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1064
        val _ = Context_Position.report_generic context pos markup;
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1065
      in pick true ("", thm_pos) [thm] end
57942
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
  1066
  | retrieve pick context (Facts.Named ((xname, pos), sel)) =
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
  1067
      let
56140
ed92ce2ac88e just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents: 56139
diff changeset
  1068
        val thy = Context.theory_of context;
63609
be0a4a0bf7f5 prefer hardwired "nothing";
wenzelm
parents: 63542
diff changeset
  1069
        fun immediate thms = {name = xname, dynamic = false, thms = map (Thm.transfer thy) thms};
63255
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
  1070
        val {name, dynamic, thms} =
56140
ed92ce2ac88e just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents: 56139
diff changeset
  1071
          (case xname of
63609
be0a4a0bf7f5 prefer hardwired "nothing";
wenzelm
parents: 63542
diff changeset
  1072
            "" => immediate [Drule.dummy_thm]
be0a4a0bf7f5 prefer hardwired "nothing";
wenzelm
parents: 63542
diff changeset
  1073
          | "_" => immediate [Drule.asm_rl]
be0a4a0bf7f5 prefer hardwired "nothing";
wenzelm
parents: 63542
diff changeset
  1074
          | "nothing" => immediate []
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
  1075
          | _ => retrieve_generic context (xname, pos));
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1076
        val thms' =
63255
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
  1077
          if dynamic andalso Config.get_generic context dynamic_facts_dummy
62078
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1078
          then [Drule.free_dummy_thm]
76399b8fde4d more systematic treatment of dynamic facts, when forming closure;
wenzelm
parents: 61902
diff changeset
  1079
          else Facts.select (Facts.Named ((name, pos), sel)) thms;
63255
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
  1080
      in pick (dynamic andalso is_none sel) (name, pos) thms' end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1081
26361
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1082
in
26346
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
  1083
57942
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
  1084
val get_fact_generic =
63255
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
  1085
  retrieve (fn dynamic => fn (name, _) => fn thms =>
3f594efa9504 clarified signature;
wenzelm
parents: 63083
diff changeset
  1086
    (if dynamic then SOME name else NONE, thms));
57942
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
  1087
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
  1088
val get_fact = retrieve (K (K I)) o Context.Proof;
e5bec882fdd0 more informative Token.Fact: retain name of dynamic fact (without selection);
wenzelm
parents: 57928
diff changeset
  1089
val get_fact_single = retrieve (K Facts.the_single) o Context.Proof;
26346
17debd2fff8e simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
  1090
26361
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1091
fun get_thms ctxt = get_fact ctxt o Facts.named;
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1092
fun get_thm ctxt = get_fact_single ctxt o Facts.named;
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1093
7946f459c6c8 Facts.Named: include position;
wenzelm
parents: 26346
diff changeset
  1094
end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1095
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1096
63083
wenzelm
parents: 63080
diff changeset
  1097
(* inner statement mode *)
wenzelm
parents: 63080
diff changeset
  1098
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
  1099
val inner_stmt = Config.declare_bool ("inner_stmt", \<^here>) (K false);
63083
wenzelm
parents: 63080
diff changeset
  1100
fun is_stmt ctxt = Config.get ctxt inner_stmt;
wenzelm
parents: 63080
diff changeset
  1101
val set_stmt = Config.put inner_stmt;
wenzelm
parents: 63080
diff changeset
  1102
val restore_stmt = set_stmt o is_stmt;
wenzelm
parents: 63080
diff changeset
  1103
wenzelm
parents: 63080
diff changeset
  1104
26284
533dcb120a8e replaced obsolete FactIndex.T by Facts.T;
wenzelm
parents: 26268
diff changeset
  1105
(* facts *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1106
61811
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
  1107
fun add_thms_dynamic arg ctxt =
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
  1108
  ctxt |> map_facts_result (Facts.add_dynamic (Context.Proof ctxt) arg);
1530a0f19539 added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
wenzelm
parents: 61268
diff changeset
  1109
28082
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
  1110
local
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
  1111
67670
wenzelm
parents: 67662
diff changeset
  1112
fun add_facts {index} arg ctxt = ctxt
wenzelm
parents: 67662
diff changeset
  1113
  |> map_facts_result (Facts.add_static (Context.Proof ctxt) {strict = false, index = index} arg);
67662
50db601cba27 store facts as lazy values;
wenzelm
parents: 67522
diff changeset
  1114
67670
wenzelm
parents: 67662
diff changeset
  1115
fun update_facts flags (b, SOME ths) ctxt = ctxt |> add_facts flags (b, Lazy.value ths) |> #2
wenzelm
parents: 67662
diff changeset
  1116
  | update_facts _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b));
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1117
79368
a2d8ecb13d3f clarified signature;
wenzelm
parents: 79349
diff changeset
  1118
fun full_name_pos ctxt b = (full_name ctxt b, Binding.default_pos_of b);
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70427
diff changeset
  1119
30761
ac7570d80c3d renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents: 30758
diff changeset
  1120
in
ac7570d80c3d renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents: 30758
diff changeset
  1121
67671
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
  1122
fun add_thms_lazy kind (b, ths) ctxt =
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
  1123
  let
79368
a2d8ecb13d3f clarified signature;
wenzelm
parents: 79349
diff changeset
  1124
    val name_pos = full_name_pos ctxt b;
68661
5820f0f379ae added system option "strict_facts";
wenzelm
parents: 67779
diff changeset
  1125
    val ths' =
5820f0f379ae added system option "strict_facts";
wenzelm
parents: 67779
diff changeset
  1126
      Global_Theory.check_thms_lazy ths
70427
973bf3e42e54 clarified signature;
wenzelm
parents: 70424
diff changeset
  1127
      |> Lazy.map_finished
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70427
diff changeset
  1128
          (Global_Theory.name_thms Global_Theory.unofficial1 name_pos #> map (Thm.kind_rule kind));
67671
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
  1129
    val (_, ctxt') = add_facts {index = is_stmt ctxt} (b, ths') ctxt;
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
  1130
  in ctxt' end;
857da80611ab support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
wenzelm
parents: 67670
diff changeset
  1131
67713
041898537c19 tuned signature;
wenzelm
parents: 67679
diff changeset
  1132
fun note_thms kind ((b, more_atts), facts) ctxt =
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1133
  let
79368
a2d8ecb13d3f clarified signature;
wenzelm
parents: 79349
diff changeset
  1134
    val name_pos = full_name_pos ctxt b;
79349
wenzelm
parents: 79342
diff changeset
  1135
    fun app_fact (thms, atts) =
79373
589d8d9018d8 more accurate Global_Theory.name_facts: burrow into expression of attributed theorems;
wenzelm
parents: 79368
diff changeset
  1136
      fold_map (Thm.proof_attributes (surround (Thm.kind kind) (atts @ more_atts))) thms;
589d8d9018d8 more accurate Global_Theory.name_facts: burrow into expression of attributed theorems;
wenzelm
parents: 79368
diff changeset
  1137
    val (thms', ctxt') =
589d8d9018d8 more accurate Global_Theory.name_facts: burrow into expression of attributed theorems;
wenzelm
parents: 79368
diff changeset
  1138
      fold_maps app_fact (Global_Theory.name_facts Global_Theory.unofficial1 name_pos facts) ctxt;
79349
wenzelm
parents: 79342
diff changeset
  1139
    val thms'' = Global_Theory.name_thms Global_Theory.unofficial2 name_pos thms';
wenzelm
parents: 79342
diff changeset
  1140
    val (_, ctxt'') = ctxt' |> add_facts {index = is_stmt ctxt} (b, Lazy.value thms'');
wenzelm
parents: 79342
diff changeset
  1141
  in ((#1 name_pos, thms''), ctxt'') end;
67713
041898537c19 tuned signature;
wenzelm
parents: 67679
diff changeset
  1142
041898537c19 tuned signature;
wenzelm
parents: 67679
diff changeset
  1143
val note_thmss = fold_map o note_thms;
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
  1144
49747
2cf86639b77e more explicit flags for facts table;
wenzelm
parents: 49691
diff changeset
  1145
fun put_thms index thms ctxt = ctxt
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
  1146
  |> map_naming (K Name_Space.local_naming)
33383
12d79ece3f7e modernized structure Context_Position;
wenzelm
parents: 33368
diff changeset
  1147
  |> Context_Position.set_visible false
67670
wenzelm
parents: 67662
diff changeset
  1148
  |> update_facts {index = index} (apfst Binding.name thms)
33383
12d79ece3f7e modernized structure Context_Position;
wenzelm
parents: 33368
diff changeset
  1149
  |> Context_Position.restore_visible ctxt
28417
74e580c6f2b5 put_thms: ContextPosition.set_visible false;
wenzelm
parents: 28411
diff changeset
  1150
  |> restore_naming ctxt;
28082
37350f301128 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
  1151
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
  1152
end;
9196
1f6f66fe777a facts: handle multiple lists of arguments;
wenzelm
parents: 9133
diff changeset
  1153
77979
a12c48fbf10f back to more elementary concept of aliases as adhoc change of accesses, but now with "suppress" information (see also 31ea5c1f874d);
wenzelm
parents: 77970
diff changeset
  1154
fun alias_fact b c ctxt = map_facts (Facts.alias (naming_of ctxt) b c) ctxt;
66246
c2c18b6b48da tuned signature;
wenzelm
parents: 66245
diff changeset
  1155
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1156
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1157
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
  1158
(** basic logical entities **)
17360
fa1f262dbc4e added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents: 17221
diff changeset
  1159
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1160
(* variables *)
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1161
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
  1162
fun declare_var (x, opt_T, mx) ctxt =
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
  1163
  let val T = (case opt_T of SOME T => T | NONE => Mixfix.default_constraint ctxt mx)
62768
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1164
  in (T, ctxt |> Variable.declare_constraints (Free (x, T))) end;
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
  1165
62767
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1166
fun add_syntax vars ctxt =
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1167
  map_syntax_idents (Local_Syntax.add_syntax ctxt (map (pair Local_Syntax.Fixed) vars)) ctxt;
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1168
57486
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1169
fun check_var internal b =
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1170
  let
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1171
    val x = Variable.check_name b;
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1172
    val check = if internal then Name.reject_skolem else Name.reject_internal;
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1173
    val _ =
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1174
      if can check (x, []) andalso Symbol_Pos.is_identifier x then ()
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1175
      else error ("Bad name: " ^ Binding.print b);
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1176
  in x end;
2131b6633529 check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
wenzelm
parents: 57415
diff changeset
  1177
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1178
local
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1179
62768
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1180
fun check_mixfix ctxt (b, T, mx) =
62767
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1181
  let
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1182
    val ([x], ctxt') = Variable.add_fixes_binding [Binding.reset_pos b] ctxt;
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1183
    val mx' = Mixfix.reset_pos mx;
71675
55cb4271858b less redundant markup reports;
wenzelm
parents: 71674
diff changeset
  1184
    val _ = add_syntax [(x, T, if Context_Position.reports_enabled ctxt then mx else mx')] ctxt';
62767
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1185
  in mx' end;
d6b0d35b3aed relevant check_mixfix happens further at the bottom, to avoid duplicate reports via Specification.prepare;
wenzelm
parents: 62078
diff changeset
  1186
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1187
fun prep_var prep_typ internal (b, raw_T, mx) ctxt =
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1188
  let
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1189
    val x = check_var internal b;
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1190
    fun cond_tvars T =
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1191
      if internal then T
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1192
      else Type.no_tvars T handle TYPE (msg, _, _) => error msg;
60468
wenzelm
parents: 60413
diff changeset
  1193
    val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T;
62768
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1194
    val (T, ctxt') = ctxt |> declare_var (x, opt_T, mx);
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1195
    val mx' = if Mixfix.is_empty mx then mx else check_mixfix ctxt' (b, T, mx);
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1196
  in ((b, SOME T, mx'), ctxt') end;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1197
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1198
in
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1199
60379
51d9dcd71ad7 tuned signature;
wenzelm
parents: 60377
diff changeset
  1200
val read_var = prep_var Syntax.read_typ false;
60468
wenzelm
parents: 60413
diff changeset
  1201
val cert_var = prep_var cert_typ true;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1202
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1203
end;
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
  1204
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1205
74331
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1206
(* syntax *)
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1207
80749
232a839ef8e6 clarified signature: more operations;
wenzelm
parents: 80744
diff changeset
  1208
val is_syntax_const = Syntax.is_const o syntax_of;
232a839ef8e6 clarified signature: more operations;
wenzelm
parents: 80744
diff changeset
  1209
74331
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1210
fun check_syntax_const ctxt (c, pos) =
80749
232a839ef8e6 clarified signature: more operations;
wenzelm
parents: 80744
diff changeset
  1211
  if is_syntax_const ctxt c then c
74331
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1212
  else error ("Unknown syntax const: " ^ quote c ^ Position.here pos);
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1213
74333
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
  1214
fun syntax add mode args ctxt =
81116
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1215
  let
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1216
    val add' = Syntax.effective_polarity ctxt add;
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1217
    val args' = map (pair Local_Syntax.Const) args;
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1218
  in ctxt |> map_syntax (#2 o Local_Syntax.update_modesyntax ctxt add' mode args') end;
74333
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
  1219
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
  1220
fun generic_syntax add mode args =
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
  1221
  Context.mapping (Sign.syntax_global add mode args) (syntax add mode args);
74333
a9b20bc32fa6 localized command 'syntax' and 'no_syntax';
wenzelm
parents: 74332
diff changeset
  1222
81594
7e1b66416b7f commands 'syntax_types' and 'syntax_consts' now work in a local theory context;
wenzelm
parents: 81591
diff changeset
  1223
fun generic_syntax_deps args =
7e1b66416b7f commands 'syntax_types' and 'syntax_consts' now work in a local theory context;
wenzelm
parents: 81591
diff changeset
  1224
  Context.mapping (Sign.syntax_deps args)
7e1b66416b7f commands 'syntax_types' and 'syntax_consts' now work in a local theory context;
wenzelm
parents: 81591
diff changeset
  1225
    (fn ctxt => map_syntax (Local_Syntax.syntax_deps ctxt args) ctxt);
7e1b66416b7f commands 'syntax_types' and 'syntax_consts' now work in a local theory context;
wenzelm
parents: 81591
diff changeset
  1226
74331
b9a3d2fb53e0 clarified signature;
wenzelm
parents: 74232
diff changeset
  1227
81591
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1228
(* translations *)
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1229
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1230
fun translations add args ctxt =
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1231
  let
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1232
    val add' = Syntax.effective_polarity ctxt add;
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1233
    val _ = Sign.check_translations ctxt args;
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1234
  in ctxt |> map_syntax (Local_Syntax.translations ctxt add' args) end;
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1235
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1236
fun generic_translations add args =
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1237
  Context.mapping (Sign.translations_global add args) (translations add args);
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1238
d570d215e380 syntax translations now work in a local theory context;
wenzelm
parents: 81565
diff changeset
  1239
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
  1240
(* notation *)
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
  1241
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1242
local
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1243
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35412
diff changeset
  1244
fun type_syntax (Type (c, args), mx) =
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42287
diff changeset
  1245
      SOME (Local_Syntax.Type, (Lexicon.mark_type c, Mixfix.make_type (length args), mx))
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1246
  | type_syntax _ = NONE;
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1247
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1248
fun const_syntax _ (Free (x, T), mx) = SOME (Local_Syntax.Fixed, (x, T, mx))
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1249
  | const_syntax ctxt (Const (c, _), mx) =
35255
2cb27605301f authentic syntax for *all* term constants;
wenzelm
parents: 35211
diff changeset
  1250
      (case try (Consts.type_scheme (consts_of ctxt)) c of
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42287
diff changeset
  1251
        SOME T => SOME (Local_Syntax.Const, (Lexicon.mark_const c, T, mx))
35255
2cb27605301f authentic syntax for *all* term constants;
wenzelm
parents: 35211
diff changeset
  1252
      | NONE => NONE)
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1253
  | const_syntax _ _ = NONE;
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1254
74332
wenzelm
parents: 74331
diff changeset
  1255
fun gen_notation make_syntax add mode args ctxt =
81116
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1256
  let
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1257
    val add' = Syntax.effective_polarity ctxt add;
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1258
    val args' = map_filter (make_syntax ctxt) args;
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 80918
diff changeset
  1259
  in ctxt |> map_syntax_idents (Local_Syntax.update_modesyntax ctxt add' mode args') end;
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1260
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1261
in
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
  1262
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1263
val type_notation = gen_notation (K type_syntax);
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1264
val notation = gen_notation const_syntax;
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1265
47247
23654b331d5c tuned signature;
wenzelm
parents: 47005
diff changeset
  1266
fun generic_type_notation add mode args phi =
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1267
  let
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1268
    val args' = args |> map_filter (fn (T, mx) =>
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1269
      let
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1270
        val T' = Morphism.typ phi T;
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1271
        val similar = (case (T, T') of (Type (c, _), Type (c', _)) => c = c' | _ => false);
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1272
      in if similar then SOME (T', mx) else NONE end);
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
  1273
  in Context.mapping (Sign.type_notation_global add mode args') (type_notation add mode args') end;
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1274
47247
23654b331d5c tuned signature;
wenzelm
parents: 47005
diff changeset
  1275
fun generic_notation add mode args phi =
33537
06c87d2c5b5a locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents: 33519
diff changeset
  1276
  let
06c87d2c5b5a locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents: 33519
diff changeset
  1277
    val args' = args |> map_filter (fn (t, mx) =>
06c87d2c5b5a locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents: 33519
diff changeset
  1278
      let val t' = Morphism.term phi t
06c87d2c5b5a locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents: 33519
diff changeset
  1279
      in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end);
80897
5328d67ec647 more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
wenzelm
parents: 80749
diff changeset
  1280
  in Context.mapping (Sign.notation_global add mode args') (notation add mode args') end;
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24922
diff changeset
  1281
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
  1282
end;
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 35399
diff changeset
  1283
35680
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
  1284
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
  1285
(* aliases *)
897740382442 aliases for class/type/const;
wenzelm
parents: 35669
diff changeset
  1286
36450
62eaaffe6e47 more systematic naming of tsig operations;
wenzelm
parents: 36448
diff changeset
  1287
fun type_alias b c ctxt = (map_tsig o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt;
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1288
fun const_alias b c ctxt = map_local_consts (Consts.alias (naming_of ctxt) b c) ctxt;
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
  1289
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
  1290
24767
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1291
(* local constants *)
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1292
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1293
fun add_const_constraint (c, opt_T) ctxt =
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1294
  let
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1295
    fun prepT raw_T =
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1296
      let val T = cert_typ ctxt raw_T
b8fb261ce6df removed redundant const_constraint;
wenzelm
parents: 24752
diff changeset
  1297
      in cert_term ctxt (Const (c, T)); T end;
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1298
  in ctxt |> map_local_consts (Consts.constrain (c, Option.map prepT opt_T)) end;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1299
33173
b8ca12f6681a eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents: 33165
diff changeset
  1300
fun add_abbrev mode (b, raw_t) ctxt =
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1301
  let
24675
2be1253a20d3 removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
wenzelm
parents: 24612
diff changeset
  1302
    val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42379
diff changeset
  1303
      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ Binding.print b);
20008
8d9d770e1f06 add_abbrevs/polymorphic: Variable.exportT_terms avoids over-generalization;
wenzelm
parents: 19916
diff changeset
  1304
    val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];
21807
a59f083632a7 add_abbrev: removed Assumption.add_assms (danger of inconsistent naming);
wenzelm
parents: 21803
diff changeset
  1305
    val ((lhs, rhs), consts') = consts_of ctxt
47005
421760a1efe7 maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
wenzelm
parents: 47001
diff changeset
  1306
      |> Consts.abbreviate (Context.Proof ctxt) (tsig_of ctxt) mode (b, t);
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1307
  in
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1308
    ctxt
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1309
    |> map_local_consts (K consts')
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1310
    |> Variable.declare_term rhs
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1311
    |> pair (lhs, rhs)
21704
f4fe6e5a3ee6 simplified add_abbrev -- single argument;
wenzelm
parents: 21696
diff changeset
  1312
  end;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1313
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1314
fun revert_abbrev mode c = map_local_consts (Consts.revert_abbrev mode c);
25052
a014d544f54d added revert_abbrev;
wenzelm
parents: 25039
diff changeset
  1315
47275
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1316
fun generic_add_abbrev mode arg =
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1317
  Context.mapping_result (Sign.add_abbrev mode arg) (add_abbrev mode arg);
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1318
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1319
fun generic_revert_abbrev mode arg =
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1320
  Context.mapping (Sign.revert_abbrev mode arg) (revert_abbrev mode arg);
fc95b8b6c260 tuned signature;
wenzelm
parents: 47247
diff changeset
  1321
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
  1322
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1323
(* fixes *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1324
60469
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1325
local
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1326
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1327
fun gen_fixes prep_var raw_vars ctxt =
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1328
  let
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1329
    val (vars, _) = fold_map prep_var raw_vars ctxt;
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1330
    val (xs, ctxt') = Variable.add_fixes_binding (map #1 vars) ctxt;
62987
dc8a8a7559e7 highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents: 62959
diff changeset
  1331
    val _ =
dc8a8a7559e7 highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents: 62959
diff changeset
  1332
      Context_Position.reports ctxt'
dc8a8a7559e7 highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents: 62959
diff changeset
  1333
        (flat (map2 (fn x => fn pos =>
dc8a8a7559e7 highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents: 62959
diff changeset
  1334
          [(pos, Variable.markup ctxt' x), (pos, Variable.markup_entity_def ctxt' x)])
dc8a8a7559e7 highlighting of entity def/ref positions wrt. cursor;
wenzelm
parents: 62959
diff changeset
  1335
            xs (map (Binding.pos_of o #1) vars)));
62768
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1336
    val vars' = map2 (fn x => fn (_, opt_T, mx) => (x, opt_T, mx)) xs vars;
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1337
    val (Ts, ctxt'') = fold_map declare_var vars' ctxt';
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1338
    val vars'' = map2 (fn T => fn (x, _, mx) => (x, T, mx)) Ts vars';
5f5f11ee4d37 more explicit type;
wenzelm
parents: 62767
diff changeset
  1339
  in (xs, add_syntax vars'' ctxt'') end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1340
60469
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1341
in
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1342
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1343
val add_fixes = gen_fixes cert_var;
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1344
val add_fixes_cmd = gen_fixes read_var;
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1345
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1346
end;
d1ea37df7358 tuned signature;
wenzelm
parents: 60468
diff changeset
  1347
9291
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1348
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1349
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1350
(** assumptions **)
18187
ec44df8ffd21 added revert_skolem, mk_def, add_def;
wenzelm
parents: 18152
diff changeset
  1351
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1352
local
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1353
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1354
fun gen_assms prep_propp exp args ctxt =
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1355
  let
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1356
    val (propss, binds) = prep_propp ctxt (map snd args);
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1357
    val props = flat propss;
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1358
  in
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1359
    ctxt
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1360
    |> fold Variable.declare_term props
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1361
    |> tap (Variable.warn_extra_tfrees ctxt)
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1362
    |> fold_burrow (Assumption.add_assms exp o map (Thm.cterm_of ctxt)) propss
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1363
    |-> (fn premss =>
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1364
      auto_bind_facts props
60408
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
  1365
      #> fold Variable.bind_term binds
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1366
      #> note_thmss "" (map fst args ~~ map (map (fn th => ([th], []))) premss))
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1367
  end;
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1368
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1369
in
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1370
60388
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1371
val add_assms = gen_assms cert_propp;
0c9d2a4f589d clarified Proof_Context.cert_propp/read_propp;
wenzelm
parents: 60387
diff changeset
  1372
val add_assms_cmd = gen_assms read_propp;
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1373
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1374
end;
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1375
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1376
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1377
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1378
(** cases **)
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1379
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1380
fun dest_cases prev_ctxt ctxt =
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1381
  let
63515
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1382
    val serial_of = #serial oo (Name_Space.the_entry o Name_Space.space_of_table);
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1383
    val ignored =
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1384
      (case prev_ctxt of
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 74333
diff changeset
  1385
        NONE => Intset.empty
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1386
      | SOME ctxt0 =>
63515
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1387
          let val cases0 = cases_of ctxt0 in
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 74333
diff changeset
  1388
            Intset.build (cases0 |> Name_Space.fold_table (fn (a, _) =>
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 74333
diff changeset
  1389
              Intset.insert (serial_of cases0 a)))
63515
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1390
          end);
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1391
     val cases = cases_of ctxt;
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1392
  in
63515
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1393
    Name_Space.fold_table (fn (a, c) =>
6c46a1e786da re-use name space serial as unique (!) id;
wenzelm
parents: 63514
diff changeset
  1394
      let val i = serial_of cases a
77723
b761c91c2447 performance tuning: prefer functor Set() over Table();
wenzelm
parents: 74333
diff changeset
  1395
      in not (Intset.member ignored i) ? cons (i, (a, c)) end) cases []
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1396
    |> sort (int_ord o apply2 #1) |> map #2
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1397
  end;
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1398
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1399
local
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1400
60408
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
  1401
fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
  1402
  | drop_schematic b = b;
1fd46ced2fa8 more uniform treatment of auto bindings vs. explicit user bindings;
wenzelm
parents: 60407
diff changeset
  1403
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1404
fun update_case _ ("", _) cases = cases
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1405
  | update_case _ (name, NONE) cases = Name_Space.del_table name cases
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1406
  | update_case context (name, SOME c) cases =
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1407
      #2 (Name_Space.define context false (Binding.name name, c) cases);
60573
e549969355b2 tuned signature;
wenzelm
parents: 60565
diff changeset
  1408
42501
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
  1409
fun fix (b, T) ctxt =
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
  1410
  let val ([x], ctxt') = add_fixes [(b, SOME T, NoSyn)] ctxt
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
  1411
  in (Free (x, T), ctxt') end;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1412
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1413
in
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1414
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1415
fun update_cases args ctxt =
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1416
  let val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming);
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1417
  in map_cases (fold (update_case context) args) ctxt end;
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1418
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1419
fun case_result c ctxt =
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1420
  let
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33224
diff changeset
  1421
    val Rule_Cases.Case {fixes, ...} = c;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1422
    val (ts, ctxt') = ctxt |> fold_map fix fixes;
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33224
diff changeset
  1423
    val Rule_Cases.Case {assumes, binds, cases, ...} = Rule_Cases.apply ts c;
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1424
  in
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1425
    ctxt'
70728
d5559011b9e6 clarified modules;
wenzelm
parents: 70727
diff changeset
  1426
    |> fold (maybe_bind_term o drop_schematic) binds
60573
e549969355b2 tuned signature;
wenzelm
parents: 60565
diff changeset
  1427
    |> update_cases (map (apsnd SOME) cases)
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1428
    |> pair (assumes, (binds, cases))
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1429
  end;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1430
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1431
val apply_case = apfst fst oo case_result;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1432
60629
d4e97fcdf83e allow to specify suffix of goal parameters;
wenzelm
parents: 60581
diff changeset
  1433
fun check_case ctxt internal (name, pos) param_specs =
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1434
  let
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1435
    val (_, Rule_Cases.Case {fixes, assumes, binds, cases}) =
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1436
      Name_Space.check (Context.Proof ctxt) (cases_of ctxt) (name, pos);
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1437
60629
d4e97fcdf83e allow to specify suffix of goal parameters;
wenzelm
parents: 60581
diff changeset
  1438
    val _ = List.app (fn NONE => () | SOME b => ignore (check_var internal b)) param_specs;
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1439
    fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1440
      | replace [] ys = ys
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1441
      | replace (_ :: _) [] =
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1442
          error ("Too many parameters for case " ^ quote name ^ Position.here pos);
60629
d4e97fcdf83e allow to specify suffix of goal parameters;
wenzelm
parents: 60581
diff changeset
  1443
    val fixes' = replace param_specs fixes;
53378
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1444
    val binds' = map drop_schematic binds;
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1445
  in
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1446
    if null (fold (Term.add_tvarsT o snd) fixes []) andalso
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1447
      null (fold (fold Term.add_vars o snd) assumes []) then
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1448
        Rule_Cases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases}
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1449
    else error ("Illegal schematic variable(s) in case " ^ quote name ^ Position.here pos)
07990ba8c0ea cases: more position information and PIDE markup;
wenzelm
parents: 52223
diff changeset
  1450
  end;
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1451
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1452
end;
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1453
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1454
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1455
(* structured statements *)
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1456
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1457
type stmt =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1458
 {vars: ((binding * typ option * mixfix) * (string * term)) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1459
  propss: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1460
  binds: (indexname * term) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1461
  result_binds: (indexname * term) list};
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1462
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1463
type statement =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1464
 {fixes: (string * term) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1465
  assumes: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1466
  shows: term list list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1467
  result_binds: (indexname * term option) list,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1468
  text: term,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1469
  result_text: term};
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1470
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1471
local
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1472
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1473
fun export_binds ctxt' ctxt params binds =
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1474
  let
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1475
    val rhss =
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1476
      map (the_list o Option.map (Logic.close_term params) o snd) binds
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1477
      |> burrow (Variable.export_terms ctxt' ctxt)
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1478
      |> map (try the_single);
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1479
  in map fst binds ~~ rhss end;
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1480
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1481
fun prep_stmt prep_var prep_propp raw_vars raw_propps ctxt =
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1482
  let
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1483
    val (vars, vars_ctxt) = fold_map prep_var raw_vars ctxt;
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1484
    val xs = map (Variable.check_name o #1) vars;
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1485
    val (xs', fixes_ctxt) = add_fixes vars vars_ctxt;
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1486
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1487
    val (propss, binds) = prep_propp fixes_ctxt raw_propps;
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1488
    val (ps, params_ctxt) = fixes_ctxt
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1489
      |> (fold o fold) Variable.declare_term propss
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1490
      |> fold_map inferred_param xs';
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1491
    val params = xs ~~ map Free ps;
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1492
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1493
    val vars' = map2 (fn (b, _, mx) => fn (_, T) => (b, SOME T, mx)) vars ps;
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1494
    val binds' = binds
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1495
      |> map (apsnd SOME)
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1496
      |> export_binds params_ctxt ctxt params
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1497
      |> map (apsnd the);
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1498
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1499
    val _ = Variable.warn_extra_tfrees fixes_ctxt params_ctxt;
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1500
    val result : stmt =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1501
      {vars = vars' ~~ params, propss = propss, binds = binds, result_binds = binds'};
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1502
  in (result, params_ctxt) end;
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1503
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1504
fun prep_statement prep_var prep_propp raw_fixes raw_assumes raw_shows ctxt =
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1505
  let
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1506
    val ((fixes, (assumes, shows), binds), ctxt') = ctxt
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1507
      |> prep_stmt prep_var prep_propp raw_fixes (raw_assumes @ raw_shows)
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1508
      |-> (fn {vars, propss, binds, ...} =>
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1509
        fold Variable.bind_term binds #>
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1510
        pair (map #2 vars, chop (length raw_assumes) propss, binds));
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1511
    val binds' =
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1512
      (Auto_Bind.facts ctxt' (flat shows) @
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1513
        (case try List.last (flat shows) of
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1514
          NONE => []
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1515
        | SOME prop => map (apsnd (SOME o Auto_Bind.abs_params prop)) binds))
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1516
      |> export_binds ctxt' ctxt fixes;
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1517
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1518
    val text = Logic.close_prop fixes (flat assumes) (Logic.mk_conjunction_list (flat shows));
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1519
    val text' = singleton (Variable.export_terms ctxt' ctxt) text;
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1520
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1521
    val result : statement =
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1522
     {fixes = fixes,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1523
      assumes = assumes,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1524
      shows = shows,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1525
      result_binds = binds',
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1526
      text = text,
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1527
      result_text = text'};
31364e70ff3e clarified signature;
wenzelm
parents: 70733
diff changeset
  1528
  in (result, ctxt') end;
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1529
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1530
in
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1531
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1532
val cert_stmt = prep_stmt cert_var cert_propp;
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
  1533
val read_stmt = prep_stmt read_var read_propp;
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1534
val cert_statement = prep_statement cert_var cert_propp;
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
  1535
val read_statement = prep_statement read_var read_propp;
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1536
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1537
end;
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1538
b8b672f70d76 clarified modules;
wenzelm
parents: 62987
diff changeset
  1539
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1540
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1541
(** print context information **)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1542
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1543
(* local syntax *)
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1544
80074
951c371c1cd9 clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents: 79471
diff changeset
  1545
val print_syntax = Syntax.print_syntax o syntax_of;
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1546
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1547
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1548
(* abbreviations *)
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1549
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1550
fun pretty_abbrevs verbose show_globals ctxt =
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1551
  let
56025
d74fed45fa8b abstract type Name_Space.table;
wenzelm
parents: 56007
diff changeset
  1552
    val space = const_space ctxt;
81543
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1553
    val (_, consts, global_consts) = #consts (rep_data ctxt);
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1554
    val constants = #constants (Consts.dest consts);
fa37ee54644c clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents: 81539
diff changeset
  1555
    val globals = Symtab.make (#constants (Consts.dest global_consts));
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1556
    fun add_abbr (_, (_, NONE)) = I
25406
1aa7927a6759 simplified Consts.dest;
wenzelm
parents: 25353
diff changeset
  1557
      | add_abbr (c, (T, SOME t)) =
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1558
          if not show_globals andalso Symtab.defined globals c then I
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1559
          else cons (c, Logic.mk_equals (Const (c, T), t));
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1560
    val abbrevs = Name_Space.markup_entries verbose ctxt space (fold add_abbr constants []);
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1561
  in
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1562
    if null abbrevs then []
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1563
    else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1564
  end;
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1565
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1566
fun print_abbrevs verbose = Pretty.writeln_chunks o pretty_abbrevs verbose true;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1567
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1568
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1569
(* term bindings *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1570
57415
e721124f1b1e command 'print_term_bindings' supersedes 'print_binds';
wenzelm
parents: 56867
diff changeset
  1571
fun pretty_term_bindings ctxt =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1572
  let
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
  1573
    val binds = Variable.binds_of ctxt;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1574
    fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1575
  in
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1576
    if Vartab.is_empty binds then []
15758
07e382399a96 binds/thms: do not store options, but delete from table;
wenzelm
parents: 15750
diff changeset
  1577
    else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1578
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1579
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1580
56155
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1581
(* local facts *)
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1582
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1583
fun pretty_local_facts verbose ctxt =
20012
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1584
  let
56155
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1585
    val facts = facts_of ctxt;
63337
ae9330fdbc16 position information for literal facts;
wenzelm
parents: 63255
diff changeset
  1586
    val props = map #1 (Facts.props facts);
56155
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1587
    val local_facts =
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1588
      (if null props then [] else [("<unnamed>", props)]) @
56158
c2c6d560e7b2 more explicit treatment of verbose mode, which includes concealed entries;
wenzelm
parents: 56156
diff changeset
  1589
      Facts.dest_static verbose [Global_Theory.facts_of (theory_of ctxt)] facts;
20012
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1590
  in
56155
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1591
    if null local_facts then []
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1592
    else
1b9c089ed12d clarified print_local_facts;
wenzelm
parents: 56145
diff changeset
  1593
      [Pretty.big_list "local facts:"
60924
610794dff23c tuned signature, in accordance to sortBy in Scala;
wenzelm
parents: 60799
diff changeset
  1594
        (map #1 (sort_by (#1 o #2) (map (`(pretty_fact ctxt)) local_facts)))]
20012
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1595
  end;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1596
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1597
fun print_local_facts verbose ctxt =
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1598
  Pretty.writeln_chunks (pretty_local_facts verbose ctxt);
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1599
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1600
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1601
(* named local contexts *)
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1602
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1603
local
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1604
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1605
fun pretty_case (name, (fixes, ((asms, (lets, cs)), ctxt))) =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1606
  let
80307
718daea1cf99 clarified modules;
wenzelm
parents: 80299
diff changeset
  1607
    val prt_name = Thy_Header.pretty_name' ctxt;
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
  1608
    val prt_term = Syntax.pretty_term ctxt;
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1609
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1610
    fun prt_let (xi, t) = Pretty.block
10818
37fa05a12459 tuned output;
wenzelm
parents: 10810
diff changeset
  1611
      [Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1,
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1612
        Pretty.quote (prt_term t)];
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1613
60565
b7ee41f72add clarified 'case' command;
wenzelm
parents: 60489
diff changeset
  1614
    fun prt_asm (a, ts) =
60242
3a8501876dba tuned output -- avoid empty quites and extra breaks;
wenzelm
parents: 59990
diff changeset
  1615
      Pretty.block (Pretty.breaks
64398
5076725247fa more robust printing of names in the context of outer syntax;
wenzelm
parents: 63640
diff changeset
  1616
        ((if a = "" then [] else [prt_name a, Pretty.str ":"]) @
60242
3a8501876dba tuned output -- avoid empty quites and extra breaks;
wenzelm
parents: 59990
diff changeset
  1617
          map (Pretty.quote o prt_term) ts));
13425
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
  1618
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1619
    fun prt_sect _ _ _ [] = []
56041
1403a22e5538 tuned message;
wenzelm
parents: 56040
diff changeset
  1620
      | prt_sect head sep prt xs =
1403a22e5538 tuned message;
wenzelm
parents: 56040
diff changeset
  1621
          [Pretty.block (Pretty.breaks (head ::
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1622
            flat (separate sep (map (single o prt) xs))))];
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1623
  in
70255
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1624
    Pretty.block
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1625
      (prt_name name :: Pretty.str ":" :: Pretty.fbrk ::
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1626
        Pretty.fbreaks
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1627
          (prt_sect (Pretty.keyword1 "fix") [] (prt_name o Binding.name_of o fst) fixes @
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1628
           prt_sect (Pretty.keyword1 "let") [Pretty.keyword2 "and"] prt_let
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1629
             (map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1630
           (if forall (null o #2) asms then []
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1631
            else prt_sect (Pretty.keyword1 "assume") [Pretty.keyword2 "and"] prt_asm asms) @
81c6a9a9a791 proper formatting (amending 5076725247fa);
wenzelm
parents: 69575
diff changeset
  1632
           prt_sect (Pretty.str "subcases:") [] (prt_name o fst) cs))
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1633
  end;
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1634
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1635
in
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1636
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1637
fun pretty_cases ctxt =
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1638
  let
69045
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1639
    val cases =
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1640
      dest_cases NONE ctxt |> map (fn (name, c as Rule_Cases.Case {fixes, ...}) =>
8c240fdeffcb discontinued old-style goal cases;
wenzelm
parents: 68661
diff changeset
  1641
        (name, (fixes, case_result c ctxt)));
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1642
  in
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1643
    if null cases then []
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1644
    else [Pretty.big_list "cases:" (map pretty_case cases)]
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1645
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1646
26722
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1647
end;
a239220108a5 print_cases: proper context for revert_skolem;
wenzelm
parents: 26717
diff changeset
  1648
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1649
fun print_cases_proof ctxt0 ctxt =
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1650
  let
63514
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1651
    fun trim_name x = if Name.is_internal x then Name.clean x else "_";
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 66246
diff changeset
  1652
    val trim_names = map trim_name #> drop_suffix (equal "_");
63514
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1653
80307
718daea1cf99 clarified modules;
wenzelm
parents: 80299
diff changeset
  1654
    val print_name = Thy_Header.print_name' ctxt;
718daea1cf99 clarified modules;
wenzelm
parents: 80299
diff changeset
  1655
63514
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1656
    fun print_case name xs =
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1657
      (case trim_names xs of
80307
718daea1cf99 clarified modules;
wenzelm
parents: 80299
diff changeset
  1658
        [] => print_name name
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 80897
diff changeset
  1659
      | xs' => enclose "(" ")" (implode_space (map print_name (name :: xs'))));
63514
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1660
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1661
    fun is_case x t =
d4d3df24f536 clarified;
wenzelm
parents: 63513
diff changeset
  1662
      x = Rule_Cases.case_conclN andalso not (Term.exists_subterm Term.is_Var t);
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1663
73616
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1664
    fun indentation depth = prefix (Symbol.spaces (2 * depth));
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1665
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1666
    fun print_proof depth (name, Rule_Cases.Case {fixes, binds, cases, ...}) =
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1667
          let
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1668
            val indent = indentation depth;
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1669
            val head = indent ("case " ^ print_case name (map (Binding.name_of o #1) fixes));
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1670
            val tail =
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1671
              if null cases then
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1672
                let val concl =
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1673
                  if exists (fn ((x, _), SOME t) => is_case x t | _ => false) binds
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1674
                  then Rule_Cases.case_conclN else Auto_Bind.thesisN
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1675
                in indent ("then show ?" ^ concl ^ " sorry") end
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1676
              else print_proofs depth cases;
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1677
          in head ^ "\n" ^ tail end
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1678
    and print_proofs 0 [] = ""
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1679
      | print_proofs depth cases =
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1680
          let
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1681
            val indent = indentation depth;
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1682
            val body = map (print_proof (depth + 1)) cases |> separate (indent "next")
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1683
          in
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1684
            if depth = 0 then body @ [indent "qed"]
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1685
            else if length cases = 1 then body
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1686
            else indent "{" :: body @ [indent "}"]
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1687
          end |> cat_lines;
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1688
  in
73616
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1689
    (case print_proofs 0 (dest_cases (SOME ctxt0) ctxt) of
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1690
      "" => ""
b0ea03e837b1 support nested cases;
wenzelm
parents: 71910
diff changeset
  1691
    | s => "Proof outline with cases:\n" ^ Active.sendback_markup_command s)
63513
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1692
  end;
9f8d06f23c09 information about proof outline with cases (sendback);
wenzelm
parents: 63509
diff changeset
  1693
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1694
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1695
(* core context *)
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1696
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
  1697
val debug = Config.declare_bool ("Proof_Context.debug", \<^here>) (K false);
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69349
diff changeset
  1698
val verbose = Config.declare_bool ("Proof_Context.verbose", \<^here>) (K false);
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1699
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1700
fun pretty_ctxt ctxt =
42717
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42502
diff changeset
  1701
  if not (Config.get ctxt debug) then []
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1702
  else
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1703
    let
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
  1704
      val prt_term = Syntax.pretty_term ctxt;
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1705
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1706
      (*structures*)
59152
66e6c539a36d more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents: 59066
diff changeset
  1707
      val {structs, ...} = Syntax_Trans.get_idents ctxt;
35139
e1a226a191b6 misc tuning and simplification;
wenzelm
parents: 35129
diff changeset
  1708
      val prt_structs =
e1a226a191b6 misc tuning and simplification;
wenzelm
parents: 35129
diff changeset
  1709
        if null structs then []
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1710
        else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1711
          Pretty.commas (map Pretty.str structs))];
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
  1712
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1713
      (*fixes*)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1714
      fun prt_fix (x, x') =
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1715
        if x = x' then Pretty.str x
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1716
        else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')];
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1717
      val fixes =
55948
bb21b380f65d tuned signature;
wenzelm
parents: 55923
diff changeset
  1718
        filter_out ((Name.is_internal orf member (op =) structs) o #1)
42488
4638622bcaa1 reorganized fixes as specialized (global) name space;
wenzelm
parents: 42469
diff changeset
  1719
          (Variable.dest_fixes ctxt);
35139
e1a226a191b6 misc tuning and simplification;
wenzelm
parents: 35129
diff changeset
  1720
      val prt_fixes =
e1a226a191b6 misc tuning and simplification;
wenzelm
parents: 35129
diff changeset
  1721
        if null fixes then []
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1722
        else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1723
          Pretty.commas (map prt_fix fixes))];
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1724
60413
0824fd1e9c40 tuned message;
wenzelm
parents: 60408
diff changeset
  1725
      (*assumptions*)
0824fd1e9c40 tuned message;
wenzelm
parents: 60408
diff changeset
  1726
      val prt_assms =
51584
98029ceda8ce more item markup;
wenzelm
parents: 51583
diff changeset
  1727
        (case Assumption.all_prems_of ctxt of
98029ceda8ce more item markup;
wenzelm
parents: 51583
diff changeset
  1728
          [] => []
60413
0824fd1e9c40 tuned message;
wenzelm
parents: 60408
diff changeset
  1729
        | prems => [Pretty.big_list "assumptions:" [pretty_fact ctxt ("", prems)]]);
0824fd1e9c40 tuned message;
wenzelm
parents: 60408
diff changeset
  1730
    in prt_structs @ prt_fixes @ prt_assms end;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1731
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1732
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1733
(* main context *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1734
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1735
fun pretty_context ctxt =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1736
  let
42717
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42502
diff changeset
  1737
    val verbose = Config.get ctxt verbose;
0bbb56867091 proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents: 42502
diff changeset
  1738
    fun verb f x = if verbose then f (x ()) else [];
35141
182f27a8716c simplified meaning of ProofContext.verbose;
wenzelm
parents: 35139
diff changeset
  1739
24922
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
  1740
    val prt_term = Syntax.pretty_term ctxt;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
  1741
    val prt_typ = Syntax.pretty_typ ctxt;
577ec55380d8 generic Syntax.pretty/string_of operations;
wenzelm
parents: 24812
diff changeset
  1742
    val prt_sort = Syntax.pretty_sort ctxt;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1743
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1744
    (*theory*)
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1745
    val pretty_thy = Pretty.block
17384
wenzelm
parents: 17360
diff changeset
  1746
      [Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)];
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1747
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1748
    (*defaults*)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1749
    fun prt_atom prt prtT (x, X) = Pretty.block
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1750
      [prt x, Pretty.str " ::", Pretty.brk 1, prtT X];
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1751
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1752
    fun prt_var (x, ~1) = prt_term (Syntax.free x)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1753
      | prt_var xi = prt_term (Syntax.var xi);
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1754
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1755
    fun prt_varT (x, ~1) = prt_typ (TFree (x, []))
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1756
      | prt_varT xi = prt_typ (TVar (xi, []));
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1757
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1758
    val prt_defT = prt_atom prt_var prt_typ;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1759
    val prt_defS = prt_atom prt_varT prt_sort;
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1760
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
  1761
    val (types, sorts) = Variable.constraints_of ctxt;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1762
  in
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1763
    verb single (K pretty_thy) @
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1764
    pretty_ctxt ctxt @
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1765
    verb (pretty_abbrevs true false) (K ctxt) @
57415
e721124f1b1e command 'print_term_bindings' supersedes 'print_binds';
wenzelm
parents: 56867
diff changeset
  1766
    verb pretty_term_bindings (K ctxt) @
59917
9830c944670f more uniform "verbose" option to print name space;
wenzelm
parents: 59896
diff changeset
  1767
    verb (pretty_local_facts true) (K ctxt) @
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1768
    verb pretty_cases (K ctxt) @
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1769
    verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
  1770
    verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1771
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1772
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1773
end;
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35412
diff changeset
  1774
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
  1775
val show_abbrevs = Proof_Context.show_abbrevs;