| author | paulson | 
| Thu, 07 Jan 2016 17:42:01 +0000 | |
| changeset 62088 | 8463e386eaec | 
| parent 62078 | 76399b8fde4d | 
| child 62767 | d6b0d35b3aed | 
| permissions | -rw-r--r-- | 
| 5819 | 1  | 
(* Title: Pure/Isar/proof_context.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 19001 | 4  | 
The key concept of Isar proof contexts: elevates primitive local  | 
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 | 7  | 
*)  | 
8  | 
||
9  | 
signature PROOF_CONTEXT =  | 
|
10  | 
sig  | 
|
| 20310 | 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  | 
| 51686 | 13  | 
val get_global: 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 | 15  | 
val mode_default: mode  | 
16  | 
val mode_stmt: mode  | 
|
17  | 
val mode_pattern: mode  | 
|
18  | 
val mode_schematic: mode  | 
|
19  | 
val mode_abbrev: mode  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
20  | 
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
 | 
21  | 
val get_mode: Proof.context -> mode  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
22  | 
val restore_mode: Proof.context -> Proof.context -> Proof.context  | 
| 27286 | 23  | 
val abbrev_mode: Proof.context -> bool  | 
| 
21667
 
ce813b82c88b
add_notation: permissive about undeclared consts;
 
wenzelm 
parents: 
21648 
diff
changeset
 | 
24  | 
val set_stmt: bool -> Proof.context -> Proof.context  | 
| 35111 | 25  | 
val syn_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
 | 
26  | 
val tsig_of: Proof.context -> Type.tsig  | 
| 36451 | 27  | 
val set_defsort: sort -> Proof.context -> Proof.context  | 
| 35680 | 28  | 
val default_sort: Proof.context -> indexname -> sort  | 
| 59840 | 29  | 
val arity_sorts: Proof.context -> string -> sort -> sort list  | 
| 20310 | 30  | 
val consts_of: Proof.context -> Consts.T  | 
| 20784 | 31  | 
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context  | 
| 20310 | 32  | 
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
 | 
33  | 
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
 | 
34  | 
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
 | 
35  | 
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
 | 
36  | 
val full_name: Proof.context -> binding -> string  | 
| 59886 | 37  | 
val get_scope: Proof.context -> Binding.scope option  | 
38  | 
val new_scope: Proof.context -> Binding.scope * Proof.context  | 
|
| 
59923
 
b21c82422d65
support private scope for individual local theory commands;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
39  | 
val private_scope: Binding.scope -> Proof.context -> Proof.context  | 
| 
 
b21c82422d65
support private scope for individual local theory commands;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
40  | 
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
 | 
41  | 
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
 | 
42  | 
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
 | 
43  | 
val concealed: Proof.context -> Proof.context  | 
| 42359 | 44  | 
val class_space: Proof.context -> Name_Space.T  | 
45  | 
val type_space: Proof.context -> Name_Space.T  | 
|
46  | 
val const_space: Proof.context -> Name_Space.T  | 
|
| 
61261
 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 
wenzelm 
parents: 
61168 
diff
changeset
 | 
47  | 
val defs_context: Proof.context -> Defs.context  | 
| 42359 | 48  | 
val intern_class: Proof.context -> xstring -> string  | 
49  | 
val intern_type: Proof.context -> xstring -> string  | 
|
50  | 
val intern_const: Proof.context -> xstring -> string  | 
|
51  | 
val extern_class: Proof.context -> string -> xstring  | 
|
| 55304 | 52  | 
val markup_class: Proof.context -> string -> string  | 
53  | 
val pretty_class: Proof.context -> string -> Pretty.T  | 
|
| 42359 | 54  | 
val extern_type: Proof.context -> string -> xstring  | 
| 55304 | 55  | 
val markup_type: Proof.context -> string -> string  | 
56  | 
val pretty_type: Proof.context -> string -> Pretty.T  | 
|
| 42359 | 57  | 
val extern_const: Proof.context -> string -> xstring  | 
| 55304 | 58  | 
val markup_const: Proof.context -> string -> string  | 
59  | 
val pretty_const: Proof.context -> string -> Pretty.T  | 
|
| 20310 | 60  | 
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
 | 
61  | 
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
 | 
62  | 
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
 | 
63  | 
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
 | 
64  | 
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
 | 
65  | 
val facts_of_fact: Proof.context -> string -> Facts.T  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
66  | 
val markup_extern_fact: Proof.context -> string -> Markup.T * xstring  | 
| 21728 | 67  | 
val pretty_term_abbrev: Proof.context -> term -> Pretty.T  | 
| 20310 | 68  | 
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
 | 
69  | 
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
 | 
70  | 
val read_class: Proof.context -> string -> class  | 
| 20310 | 71  | 
val read_typ: Proof.context -> string -> typ  | 
72  | 
val read_typ_syntax: Proof.context -> string -> typ  | 
|
73  | 
val read_typ_abbrev: Proof.context -> string -> typ  | 
|
74  | 
val cert_typ: Proof.context -> typ -> typ  | 
|
75  | 
val cert_typ_syntax: Proof.context -> typ -> typ  | 
|
76  | 
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
 | 
77  | 
val infer_type: Proof.context -> string * typ -> typ  | 
| 60407 | 78  | 
val inferred_param: string -> Proof.context -> (string * typ) * Proof.context  | 
| 25328 | 79  | 
val inferred_fixes: Proof.context -> (string * typ) list * Proof.context  | 
| 56002 | 80  | 
  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
 | 
81  | 
xstring * Position.T -> typ * Position.report list  | 
| 56002 | 82  | 
  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
 | 
83  | 
val consts_completion_message: Proof.context -> xstring * Position.T list -> string  | 
| 56002 | 84  | 
  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
 | 
85  | 
xstring * Position.T list -> term * Position.report list  | 
| 56002 | 86  | 
  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
 | 
87  | 
val read_arity: Proof.context -> xstring * string list * string -> arity  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
88  | 
val cert_arity: Proof.context -> arity -> arity  | 
| 27259 | 89  | 
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
 | 
90  | 
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
 | 
91  | 
val prepare_sorts: Proof.context -> term list -> string list * term list  | 
| 36152 | 92  | 
val check_tfree: Proof.context -> string * sort -> string * sort  | 
| 24684 | 93  | 
val read_term_pattern: Proof.context -> string -> term  | 
94  | 
val read_term_schematic: Proof.context -> string -> term  | 
|
95  | 
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
 | 
96  | 
val show_abbrevs_raw: Config.raw  | 
| 
 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
97  | 
val show_abbrevs: bool Config.T  | 
| 
25345
 
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
 
wenzelm 
parents: 
25332 
diff
changeset
 | 
98  | 
val expand_abbrevs: Proof.context -> term -> term  | 
| 20310 | 99  | 
val cert_term: Proof.context -> term -> term  | 
100  | 
val cert_prop: Proof.context -> term -> term  | 
|
| 
35616
 
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
101  | 
val def_type: Proof.context -> indexname -> typ option  | 
| 45429 | 102  | 
val standard_typ_check: Proof.context -> typ list -> typ list  | 
103  | 
val standard_term_check_finish: Proof.context -> term list -> term list  | 
|
104  | 
val standard_term_uncheck: Proof.context -> term list -> term list  | 
|
| 20310 | 105  | 
val goal_export: Proof.context -> Proof.context -> thm list -> thm list  | 
106  | 
val export: Proof.context -> Proof.context -> thm list -> thm list  | 
|
| 21531 | 107  | 
val export_morphism: Proof.context -> Proof.context -> morphism  | 
| 28396 | 108  | 
val norm_export_morphism: Proof.context -> Proof.context -> morphism  | 
| 20310 | 109  | 
val auto_bind_goal: term list -> Proof.context -> Proof.context  | 
110  | 
val auto_bind_facts: term list -> Proof.context -> Proof.context  | 
|
| 60383 | 111  | 
val match_bind: bool -> (term list * term) list -> Proof.context ->  | 
112  | 
term list * Proof.context  | 
|
113  | 
val match_bind_cmd: bool -> (string list * string) list -> Proof.context ->  | 
|
114  | 
term list * Proof.context  | 
|
| 60388 | 115  | 
val cert_propp: Proof.context -> (term * term list) list list ->  | 
116  | 
(term list list * (indexname * term) list)  | 
|
117  | 
val read_propp: Proof.context -> (string * string list) list list ->  | 
|
118  | 
(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
 | 
119  | 
val fact_tac: Proof.context -> thm list -> int -> tactic  | 
| 20310 | 120  | 
val some_fact_tac: Proof.context -> int -> tactic  | 
| 61902 | 121  | 
val lookup_fact: Proof.context -> string -> (bool * thm list) option  | 
| 
62078
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
122  | 
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
 | 
123  | 
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
 | 
124  | 
val get_fact: Proof.context -> Facts.ref -> thm list  | 
| 
 
17debd2fff8e
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
125  | 
val get_fact_single: Proof.context -> Facts.ref -> thm  | 
| 
 
17debd2fff8e
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
126  | 
val get_thms: Proof.context -> xstring -> thm list  | 
| 
 
17debd2fff8e
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
127  | 
val get_thm: Proof.context -> xstring -> thm  | 
| 
61811
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
128  | 
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
 | 
129  | 
Proof.context -> string * Proof.context  | 
| 
30761
 
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
 
wenzelm 
parents: 
30758 
diff
changeset
 | 
130  | 
val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list ->  | 
| 30211 | 131  | 
Proof.context -> (string * thm list) list * Proof.context  | 
| 
28082
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
132  | 
val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context  | 
| 60379 | 133  | 
val read_var: binding * string option * mixfix -> Proof.context ->  | 
134  | 
(binding * typ option * mixfix) * Proof.context  | 
|
135  | 
val cert_var: binding * typ option * mixfix -> Proof.context ->  | 
|
136  | 
(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
 | 
137  | 
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
 | 
138  | 
string list * Proof.context  | 
| 60469 | 139  | 
val add_fixes_cmd: (binding * string option * mixfix) list -> Proof.context ->  | 
140  | 
string list * Proof.context  | 
|
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
141  | 
val add_assms: Assumption.export ->  | 
| 60377 | 142  | 
(Thm.binding * (term * term list) list) list ->  | 
| 
28082
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
143  | 
Proof.context -> (string * thm list) list * Proof.context  | 
| 60377 | 144  | 
val add_assms_cmd: Assumption.export ->  | 
145  | 
(Thm.binding * (string * string list) list) list ->  | 
|
| 
28082
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
146  | 
Proof.context -> (string * thm list) list * Proof.context  | 
| 60799 | 147  | 
  val dest_cases: Proof.context -> (string * (Rule_Cases.T * {legacy: bool})) list
 | 
| 60573 | 148  | 
val update_cases: (string * Rule_Cases.T option) list -> Proof.context -> Proof.context  | 
149  | 
val update_cases_legacy: (string * Rule_Cases.T option) list -> Proof.context -> Proof.context  | 
|
| 60565 | 150  | 
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
 | 
151  | 
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
 | 
152  | 
string * Position.T -> binding option list -> Rule_Cases.T  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
153  | 
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context  | 
| 24949 | 154  | 
val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context  | 
| 47247 | 155  | 
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
 | 
156  | 
Context.generic -> Context.generic  | 
| 47247 | 157  | 
val generic_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->  | 
| 21744 | 158  | 
Context.generic -> Context.generic  | 
| 35680 | 159  | 
val class_alias: binding -> class -> Proof.context -> Proof.context  | 
160  | 
val type_alias: binding -> string -> Proof.context -> Proof.context  | 
|
161  | 
val const_alias: binding -> string -> Proof.context -> Proof.context  | 
|
| 57887 | 162  | 
val fact_alias: binding -> string -> Proof.context -> Proof.context  | 
| 24767 | 163  | 
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
 | 
164  | 
val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context  | 
| 25052 | 165  | 
val revert_abbrev: string -> string -> Proof.context -> Proof.context  | 
| 47275 | 166  | 
val generic_add_abbrev: string -> binding * term -> Context.generic ->  | 
167  | 
(term * term) * Context.generic  | 
|
168  | 
val generic_revert_abbrev: string -> string -> Context.generic -> Context.generic  | 
|
| 20310 | 169  | 
val print_syntax: Proof.context -> unit  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
170  | 
val print_abbrevs: bool -> Proof.context -> unit  | 
| 
57415
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
171  | 
val pretty_term_bindings: Proof.context -> Pretty.T list  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
172  | 
val pretty_local_facts: bool -> Proof.context -> Pretty.T list  | 
| 
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
173  | 
val print_local_facts: bool -> Proof.context -> unit  | 
| 56867 | 174  | 
val pretty_cases: Proof.context -> Pretty.T list  | 
| 
42717
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42502 
diff
changeset
 | 
175  | 
val debug: bool Config.T  | 
| 
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42502 
diff
changeset
 | 
176  | 
val verbose: bool Config.T  | 
| 20310 | 177  | 
val pretty_ctxt: Proof.context -> Pretty.T list  | 
178  | 
val pretty_context: Proof.context -> Pretty.T list  | 
|
| 5819 | 179  | 
end;  | 
180  | 
||
| 42360 | 181  | 
structure Proof_Context: PROOF_CONTEXT =  | 
| 5819 | 182  | 
struct  | 
183  | 
||
| 42363 | 184  | 
val theory_of = Proof_Context.theory_of;  | 
185  | 
val init_global = Proof_Context.init_global;  | 
|
| 51686 | 186  | 
val get_global = Proof_Context.get_global;  | 
| 42363 | 187  | 
|
| 12057 | 188  | 
|
| 7270 | 189  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
190  | 
(** inner syntax mode **)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
191  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
192  | 
datatype mode =  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
193  | 
Mode of  | 
| 24486 | 194  | 
   {stmt: bool,                (*inner statement mode*)
 | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
195  | 
pattern: bool, (*pattern binding schematic variables*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
196  | 
schematic: bool, (*term referencing loose schematic variables*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
197  | 
abbrev: bool}; (*abbrev mode -- no normalization*)  | 
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
198  | 
|
| 24486 | 199  | 
fun make_mode (stmt, pattern, schematic, abbrev) =  | 
200  | 
  Mode {stmt = stmt, pattern = pattern, schematic = schematic, abbrev = abbrev};
 | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
201  | 
|
| 24486 | 202  | 
val mode_default = make_mode (false, false, false, false);  | 
203  | 
val mode_stmt = make_mode (true, false, false, false);  | 
|
204  | 
val mode_pattern = make_mode (false, true, false, false);  | 
|
205  | 
val mode_schematic = make_mode (false, false, true, false);  | 
|
206  | 
val mode_abbrev = make_mode (false, false, false, true);  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
207  | 
|
| 
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
208  | 
|
| 5819 | 209  | 
|
| 16540 | 210  | 
(** Isar proof context information **)  | 
| 5819 | 211  | 
|
| 60573 | 212  | 
type cases = ((Rule_Cases.T * {legacy: bool}) * int) Name_Space.table;
 | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
213  | 
val empty_cases: cases = Name_Space.empty_table Markup.caseN;  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
214  | 
|
| 45650 | 215  | 
datatype data =  | 
216  | 
Data of  | 
|
| 36451 | 217  | 
   {mode: mode,                  (*inner syntax mode*)
 | 
218  | 
syntax: Local_Syntax.T, (*local syntax*)  | 
|
219  | 
tsig: Type.tsig * Type.tsig, (*local/global type signature -- local name space / defsort only*)  | 
|
220  | 
consts: Consts.T * Consts.T, (*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
 | 
221  | 
facts: Facts.T, (*local facts, based on initial global facts*)  | 
| 60573 | 222  | 
cases: cases}; (*named case contexts: case, legacy, running index*)  | 
| 5819 | 223  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
224  | 
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
 | 
225  | 
  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
 | 
226  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37145 
diff
changeset
 | 
227  | 
structure Data = Proof_Data  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
228  | 
(  | 
| 45650 | 229  | 
type T = data;  | 
| 16540 | 230  | 
fun init thy =  | 
| 
56139
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
231  | 
make_data (mode_default,  | 
| 
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
232  | 
Local_Syntax.init thy,  | 
| 
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
233  | 
(Type.change_ignore (Sign.tsig_of thy), Sign.tsig_of thy),  | 
| 
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
234  | 
(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
 | 
235  | 
Global_Theory.facts_of thy,  | 
| 
56139
 
b7add947a6ef
more frugal recording of changes: join merely requires information from one side;
 
wenzelm 
parents: 
56062 
diff
changeset
 | 
236  | 
empty_cases);  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
237  | 
);  | 
| 5819 | 238  | 
|
| 45650 | 239  | 
fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep);  | 
| 5819 | 240  | 
|
| 
61811
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
241  | 
fun map_data_result f ctxt =  | 
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
242  | 
let  | 
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
243  | 
    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
 | 
244  | 
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
 | 
245  | 
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
 | 
246  | 
|
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
247  | 
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
 | 
248  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
249  | 
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
 | 
250  | 
(mode, syntax, tsig, consts, facts, cases));  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
251  | 
|
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
252  | 
fun map_mode f =  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
253  | 
  map_data (fn (Mode {stmt, pattern, schematic, abbrev}, 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
 | 
254  | 
(make_mode (f (stmt, pattern, schematic, abbrev)), syntax, tsig, consts, facts, cases));  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
255  | 
|
| 19001 | 256  | 
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
 | 
257  | 
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
 | 
258  | 
(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
 | 
259  | 
|
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
260  | 
fun map_syntax_idents f ctxt =  | 
| 60382 | 261  | 
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
 | 
262  | 
ctxt  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
263  | 
|> map_syntax (K syntax')  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
264  | 
|> (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
 | 
265  | 
end;  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
266  | 
|
| 36450 | 267  | 
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
 | 
268  | 
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
 | 
269  | 
(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
 | 
270  | 
|
| 19001 | 271  | 
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
 | 
272  | 
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
 | 
273  | 
(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
 | 
274  | 
|
| 
61811
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
275  | 
fun map_facts_result f =  | 
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
276  | 
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
 | 
277  | 
let val (res, facts') = f facts  | 
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
278  | 
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
 | 
279  | 
|
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
280  | 
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
 | 
281  | 
|
| 19001 | 282  | 
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
 | 
283  | 
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
 | 
284  | 
(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
 | 
285  | 
|
| 45650 | 286  | 
val get_mode = #mode o rep_data;  | 
| 28407 | 287  | 
val restore_mode = set_mode o get_mode;  | 
| 27286 | 288  | 
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
 | 
289  | 
|
| 24486 | 290  | 
fun set_stmt stmt =  | 
291  | 
map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev));  | 
|
| 19001 | 292  | 
|
| 45650 | 293  | 
val syntax_of = #syntax o rep_data;  | 
| 33387 | 294  | 
val syn_of = Local_Syntax.syn_of o syntax_of;  | 
295  | 
val set_syntax_mode = map_syntax o Local_Syntax.set_mode;  | 
|
296  | 
val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of;  | 
|
| 19001 | 297  | 
|
| 45650 | 298  | 
val tsig_of = #1 o #tsig o rep_data;  | 
| 36451 | 299  | 
val set_defsort = map_tsig o apfst o Type.set_defsort;  | 
| 35680 | 300  | 
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
 | 
301  | 
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
 | 
302  | 
|
| 45650 | 303  | 
val consts_of = #1 o #consts o rep_data;  | 
304  | 
val cases_of = #cases o rep_data;  | 
|
| 5819 | 305  | 
|
306  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
307  | 
(* naming *)  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
308  | 
|
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
309  | 
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
 | 
310  | 
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
 | 
311  | 
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
 | 
312  | 
|
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
313  | 
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
 | 
314  | 
|
| 59886 | 315  | 
val get_scope = Name_Space.get_scope o naming_of;  | 
316  | 
||
317  | 
fun new_scope ctxt =  | 
|
318  | 
let  | 
|
319  | 
val (scope, naming') = Name_Space.new_scope (naming_of ctxt);  | 
|
320  | 
val ctxt' = map_naming (K naming') ctxt;  | 
|
321  | 
in (scope, ctxt') end;  | 
|
322  | 
||
| 
59923
 
b21c82422d65
support private scope for individual local theory commands;
 
wenzelm 
parents: 
59917 
diff
changeset
 | 
323  | 
val private_scope = map_naming o Name_Space.private_scope;  | 
| 59896 | 324  | 
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
 | 
325  | 
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
 | 
326  | 
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
 | 
327  | 
|
| 
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
 | 
328  | 
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
 | 
329  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
47001 
diff
changeset
 | 
330  | 
|
| 42359 | 331  | 
(* name spaces *)  | 
332  | 
||
333  | 
val class_space = Type.class_space o tsig_of;  | 
|
334  | 
val type_space = Type.type_space o tsig_of;  | 
|
335  | 
val const_space = Consts.space_of o consts_of;  | 
|
336  | 
||
| 61265 | 337  | 
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
 | 
338  | 
|
| 42359 | 339  | 
val intern_class = Name_Space.intern o class_space;  | 
340  | 
val intern_type = Name_Space.intern o type_space;  | 
|
341  | 
val intern_const = Name_Space.intern o const_space;  | 
|
342  | 
||
343  | 
fun extern_class ctxt = Name_Space.extern ctxt (class_space ctxt);  | 
|
344  | 
fun extern_type ctxt = Name_Space.extern ctxt (type_space ctxt);  | 
|
345  | 
fun extern_const ctxt = Name_Space.extern ctxt (const_space ctxt);  | 
|
346  | 
||
| 55304 | 347  | 
fun markup_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |-> Markup.markup;  | 
348  | 
fun markup_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |-> Markup.markup;  | 
|
349  | 
fun markup_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |-> Markup.markup;  | 
|
350  | 
||
351  | 
fun pretty_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |> Pretty.mark_str;  | 
|
352  | 
fun pretty_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |> Pretty.mark_str;  | 
|
353  | 
fun pretty_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |> Pretty.mark_str;  | 
|
354  | 
||
| 42359 | 355  | 
|
| 20367 | 356  | 
(* theory transfer *)  | 
| 12093 | 357  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
358  | 
fun transfer_syntax thy ctxt = ctxt |>  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
359  | 
map_syntax (Local_Syntax.rebuild thy) |>  | 
| 36450 | 360  | 
map_tsig (fn tsig as (local_tsig, global_tsig) =>  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
361  | 
let val thy_tsig = Sign.tsig_of thy in  | 
| 36450 | 362  | 
if Type.eq_tsig (thy_tsig, global_tsig) then tsig  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
363  | 
else (Type.merge_tsig (Context.Proof ctxt) (local_tsig, thy_tsig), thy_tsig) (*historic merge order*)  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
364  | 
end) |>  | 
| 
30424
 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 
wenzelm 
parents: 
30420 
diff
changeset
 | 
365  | 
map_consts (fn consts as (local_consts, global_consts) =>  | 
| 
 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 
wenzelm 
parents: 
30420 
diff
changeset
 | 
366  | 
let val thy_consts = Sign.consts_of thy in  | 
| 
 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 
wenzelm 
parents: 
30420 
diff
changeset
 | 
367  | 
if Consts.eq_consts (thy_consts, global_consts) then consts  | 
| 
57928
 
f4ff42c5b05b
transfer result of Global_Theory.add_thms_dynamic to context stack;
 
wenzelm 
parents: 
57887 
diff
changeset
 | 
368  | 
else (Consts.merge (local_consts, thy_consts), thy_consts) (*historic merge order*)  | 
| 
30424
 
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
 
wenzelm 
parents: 
30420 
diff
changeset
 | 
369  | 
end);  | 
| 17072 | 370  | 
|
| 
33031
 
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
 
wenzelm 
parents: 
32966 
diff
changeset
 | 
371  | 
fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;  | 
| 17072 | 372  | 
|
| 
57928
 
f4ff42c5b05b
transfer result of Global_Theory.add_thms_dynamic to context stack;
 
wenzelm 
parents: 
57887 
diff
changeset
 | 
373  | 
fun transfer_facts thy =  | 
| 
 
f4ff42c5b05b
transfer result of Global_Theory.add_thms_dynamic to context stack;
 
wenzelm 
parents: 
57887 
diff
changeset
 | 
374  | 
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
 | 
375  | 
|
| 
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
 | 
376  | 
fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt;  | 
| 20367 | 377  | 
|
| 
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
 | 
378  | 
fun background_theory_result f ctxt =  | 
| 20367 | 379  | 
let val (res, thy') = f (theory_of ctxt)  | 
380  | 
in (res, ctxt |> transfer thy') end;  | 
|
| 19019 | 381  | 
|
| 12093 | 382  | 
|
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
383  | 
(* hybrid facts *)  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
384  | 
|
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
385  | 
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
 | 
386  | 
|
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
387  | 
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
 | 
388  | 
let  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
389  | 
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
 | 
390  | 
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
 | 
391  | 
in  | 
| 
56143
 
ed2b660a52a1
more accurate resolution of hybrid facts, which actually changes the sort order of results;
 
wenzelm 
parents: 
56141 
diff
changeset
 | 
392  | 
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
 | 
393  | 
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
 | 
394  | 
end;  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
395  | 
|
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
396  | 
fun markup_extern_fact ctxt name =  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
397  | 
Facts.markup_extern ctxt (facts_of_fact ctxt name) name;  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
398  | 
|
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
399  | 
|
| 12093 | 400  | 
|
| 14828 | 401  | 
(** pretty printing **)  | 
402  | 
||
| 24922 | 403  | 
fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt);  | 
| 14828 | 404  | 
|
| 42378 | 405  | 
fun pretty_fact_name ctxt a =  | 
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
406  | 
Pretty.block [Pretty.mark_str (markup_extern_fact ctxt a), Pretty.str ":"];  | 
| 
28209
 
02f3222a392d
pretty_fact: extern fact name wrt. the given context, assuming that is the proper one for presentation;
 
wenzelm 
parents: 
28087 
diff
changeset
 | 
407  | 
|
| 51583 | 408  | 
fun pretty_fact ctxt =  | 
409  | 
let  | 
|
| 61268 | 410  | 
val pretty_thm = Thm.pretty_thm ctxt;  | 
411  | 
val pretty_thms = map (Thm.pretty_thm_item ctxt);  | 
|
| 51583 | 412  | 
in  | 
413  | 
    fn ("", [th]) => pretty_thm th
 | 
|
414  | 
     | ("", ths) => Pretty.blk (0, Pretty.fbreaks (pretty_thms ths))
 | 
|
415  | 
| (a, [th]) => Pretty.block [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm th]  | 
|
416  | 
| (a, ths) => Pretty.block (Pretty.fbreaks (pretty_fact_name ctxt a :: pretty_thms ths))  | 
|
417  | 
end;  | 
|
| 14828 | 418  | 
|
419  | 
||
420  | 
||
| 5819 | 421  | 
(** prepare types **)  | 
422  | 
||
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
423  | 
(* classes *)  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
424  | 
|
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
425  | 
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
 | 
426  | 
let  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
|
| 56007 | 430  | 
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
 | 
431  | 
handle TYPE (msg, _, _) =>  | 
| 
55840
 
2982d233d798
consider completion report as part of error message -- less stateful, may get handled;
 
wenzelm 
parents: 
55839 
diff
changeset
 | 
432  | 
error (msg ^ Position.here pos ^  | 
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
433  | 
Markup.markup_report (Completion.reported_text  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
434  | 
(Name_Space.completion (Context.Proof ctxt) class_space (xname, pos))));  | 
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
435  | 
val reports =  | 
| 55923 | 436  | 
if Context_Position.is_reported ctxt pos  | 
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
437  | 
then [(pos, Name_Space.markup class_space name)] else [];  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
438  | 
in (name, reports) end;  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
439  | 
|
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
440  | 
fun read_class ctxt text =  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
441  | 
let  | 
| 59795 | 442  | 
val source = Syntax.read_input text;  | 
| 59066 | 443  | 
val (c, reports) = check_class ctxt (Input.source_content source, Input.pos_of source);  | 
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
444  | 
val _ = Position.reports reports;  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
445  | 
in c end;  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
446  | 
|
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
447  | 
|
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
448  | 
(* types *)  | 
| 24277 | 449  | 
|
450  | 
fun read_typ_mode mode ctxt s =  | 
|
| 24486 | 451  | 
Syntax.read_typ (Type.set_mode mode ctxt) s;  | 
| 24277 | 452  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
453  | 
val read_typ = read_typ_mode Type.mode_default;  | 
| 24277 | 454  | 
val read_typ_syntax = read_typ_mode Type.mode_syntax;  | 
455  | 
val read_typ_abbrev = read_typ_mode Type.mode_abbrev;  | 
|
456  | 
||
457  | 
||
458  | 
fun cert_typ_mode mode ctxt T =  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
459  | 
Type.cert_typ_mode mode (tsig_of ctxt) T  | 
| 24277 | 460  | 
handle TYPE (msg, _, _) => error msg;  | 
461  | 
||
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
462  | 
val cert_typ = cert_typ_mode Type.mode_default;  | 
| 24277 | 463  | 
val cert_typ_syntax = cert_typ_mode Type.mode_syntax;  | 
464  | 
val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev;  | 
|
465  | 
||
466  | 
||
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
467  | 
|
| 5819 | 468  | 
(** prepare terms and propositions **)  | 
469  | 
||
| 25328 | 470  | 
(* inferred types of parameters *)  | 
471  | 
||
472  | 
fun infer_type ctxt x =  | 
|
| 
36505
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36503 
diff
changeset
 | 
473  | 
Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt)) (Free x));  | 
| 25328 | 474  | 
|
475  | 
fun inferred_param x ctxt =  | 
|
| 60407 | 476  | 
let val p = (x, infer_type ctxt (x, dummyT))  | 
477  | 
in (p, ctxt |> Variable.declare_term (Free p)) end;  | 
|
| 25328 | 478  | 
|
479  | 
fun inferred_fixes ctxt =  | 
|
| 60407 | 480  | 
fold_map inferred_param (map #2 (Variable.dest_fixes ctxt)) ctxt;  | 
| 25328 | 481  | 
|
482  | 
||
| 55842 | 483  | 
(* type names *)  | 
| 25328 | 484  | 
|
| 56002 | 485  | 
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
 | 
486  | 
if Lexicon.is_tid c then  | 
| 56040 | 487  | 
    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
 | 
488  | 
else  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
489  | 
let  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
490  | 
val reports =  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
491  | 
if Context_Position.is_reported ctxt pos  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
492  | 
then [(pos, Markup.tfree)] else [];  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
493  | 
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
 | 
494  | 
else  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
495  | 
let  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
496  | 
val ((d, reports), decl) = Type.check_decl (Context.Proof ctxt) (tsig_of ctxt) (c, pos);  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
497  | 
      fun err () = error ("Bad type name: " ^ quote d ^ Position.here pos);
 | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
498  | 
val args =  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
499  | 
(case decl of  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
500  | 
Type.LogicalType n => n  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
501  | 
| Type.Abbreviation (vs, _, _) => if strict then err () else length vs  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
502  | 
| Type.Nonterminal => if strict then err () else 0);  | 
| 
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55843 
diff
changeset
 | 
503  | 
in (Type (d, replicate args dummyT), reports) end;  | 
| 25328 | 504  | 
|
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
505  | 
fun read_type_name ctxt flags text =  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
506  | 
let  | 
| 59795 | 507  | 
val source = Syntax.read_input text;  | 
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
508  | 
val (T, reports) =  | 
| 59066 | 509  | 
check_type_name ctxt flags (Input.source_content source, Input.pos_of source);  | 
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
510  | 
val _ = Position.reports reports;  | 
| 
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55950 
diff
changeset
 | 
511  | 
in T end;  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
512  | 
|
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
513  | 
|
| 55842 | 514  | 
(* constant names *)  | 
515  | 
||
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
516  | 
fun consts_completion_message ctxt (c, ps) =  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
517  | 
ps |> map (fn pos =>  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
518  | 
Name_Space.completion (Context.Proof ctxt) (Consts.space_of (consts_of ctxt)) (c, pos)  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
519  | 
|> Completion.reported_text)  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
520  | 
|> implode  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
521  | 
|> Markup.markup_report;  | 
| 
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
522  | 
|
| 56002 | 523  | 
fun check_const {proper, strict} ctxt (c, ps) =
 | 
| 55842 | 524  | 
let  | 
| 
55956
 
94d384d621b0
reject internal term names outright, and complete consts instead;
 
wenzelm 
parents: 
55955 
diff
changeset
 | 
525  | 
val _ =  | 
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
526  | 
Name.reject_internal (c, ps) handle ERROR msg =>  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
527  | 
error (msg ^ consts_completion_message ctxt (c, ps));  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
528  | 
fun err msg = error (msg ^ Position.here_list ps);  | 
| 55842 | 529  | 
val consts = consts_of ctxt;  | 
| 55954 | 530  | 
val fixed = if proper then NONE else Variable.lookup_fixed ctxt c;  | 
531  | 
val (t, reports) =  | 
|
532  | 
(case (fixed, Variable.lookup_const ctxt c) of  | 
|
533  | 
(SOME x, NONE) =>  | 
|
| 
55843
 
3fa61f39d1f2
prefer Name_Space.check with its builtin reports (including completion);
 
wenzelm 
parents: 
55842 
diff
changeset
 | 
534  | 
let  | 
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
535  | 
val reports = ps  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
536  | 
|> filter (Context_Position.is_reported ctxt)  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
537  | 
|> map (fn pos =>  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
538  | 
(pos, Markup.name x (if Name.is_skolem x then Markup.skolem else Markup.free)));  | 
| 
55955
 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
539  | 
in (Free (x, infer_type ctxt (x, dummyT)), reports) end  | 
| 55954 | 540  | 
| (_, SOME d) =>  | 
541  | 
let  | 
|
542  | 
val T = Consts.type_scheme consts d handle TYPE (msg, _, _) => err msg;  | 
|
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
543  | 
val reports = ps  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
544  | 
|> filter (Context_Position.is_reported ctxt)  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
545  | 
|> map (fn pos => (pos, Name_Space.markup (Consts.space_of consts) d));  | 
| 
55950
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55949 
diff
changeset
 | 
546  | 
in (Const (d, T), reports) end  | 
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55956 
diff
changeset
 | 
547  | 
| _ => Consts.check_const (Context.Proof ctxt) consts (c, ps));  | 
| 55842 | 548  | 
val _ =  | 
| 55954 | 549  | 
(case (strict, t) of  | 
550  | 
(true, Const (d, _)) =>  | 
|
551  | 
(ignore (Consts.the_const consts d) handle TYPE (msg, _, _) => err msg)  | 
|
552  | 
| _ => ());  | 
|
| 
55950
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55949 
diff
changeset
 | 
553  | 
in (t, reports) end;  | 
| 55842 | 554  | 
|
| 
55955
 
e8f1bf005661
eliminated odd type constraint for read_const (see also 79c1d2bbe5a9);
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
555  | 
fun read_const ctxt flags text =  | 
| 
55950
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55949 
diff
changeset
 | 
556  | 
let  | 
| 59795 | 557  | 
val source = Syntax.read_input text;  | 
| 59066 | 558  | 
val (t, reports) = check_const ctxt flags (Input.source_content source, [Input.pos_of source]);  | 
| 
55950
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55949 
diff
changeset
 | 
559  | 
val _ = Position.reports reports;  | 
| 
 
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
 
wenzelm 
parents: 
55949 
diff
changeset
 | 
560  | 
in t end;  | 
| 25328 | 561  | 
|
562  | 
||
| 
46922
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
563  | 
(* type arities *)  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
564  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
565  | 
local  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
566  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
567  | 
fun prep_arity prep_tycon prep_sort ctxt (t, Ss, S) =  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
568  | 
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
 | 
569  | 
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
 | 
570  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
571  | 
in  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
572  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
573  | 
val read_arity =  | 
| 56002 | 574  | 
  prep_arity ((#1 o dest_Type) 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
 | 
575  | 
|
| 
46922
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
576  | 
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
 | 
577  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
578  | 
end;  | 
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
579  | 
|
| 
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
46849 
diff
changeset
 | 
580  | 
|
| 24684 | 581  | 
(* read_term *)  | 
582  | 
||
583  | 
fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt);  | 
|
584  | 
||
585  | 
val read_term_pattern = read_term_mode mode_pattern;  | 
|
586  | 
val read_term_schematic = read_term_mode mode_schematic;  | 
|
587  | 
val read_term_abbrev = read_term_mode mode_abbrev;  | 
|
588  | 
||
589  | 
||
| 19001 | 590  | 
(* local abbreviations *)  | 
| 5819 | 591  | 
|
| 24501 | 592  | 
local  | 
593  | 
||
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
594  | 
fun certify_consts ctxt =  | 
| 
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
595  | 
Consts.certify (Context.Proof ctxt) (tsig_of ctxt)  | 
| 
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
596  | 
(not (abbrev_mode ctxt)) (consts_of ctxt);  | 
| 19001 | 597  | 
|
| 
38979
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
598  | 
fun expand_binds ctxt =  | 
| 
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
599  | 
let  | 
| 
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
600  | 
    val Mode {pattern, schematic, ...} = get_mode ctxt;
 | 
| 5819 | 601  | 
|
| 
38979
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
602  | 
fun reject_schematic (t as Var _) =  | 
| 
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
603  | 
          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
 | 
604  | 
| 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
 | 
605  | 
| 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
 | 
606  | 
| reject_schematic _ = ();  | 
| 
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
607  | 
in  | 
| 24495 | 608  | 
if pattern then I  | 
609  | 
else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic)  | 
|
610  | 
end;  | 
|
| 5819 | 611  | 
|
| 24501 | 612  | 
in  | 
613  | 
||
614  | 
fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt;  | 
|
615  | 
||
616  | 
end;  | 
|
617  | 
||
| 56438 | 618  | 
val show_abbrevs_raw = Config.declare ("show_abbrevs", @{here}) (fn _ => Config.Bool true);
 | 
| 
40879
 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
619  | 
val show_abbrevs = Config.bool show_abbrevs_raw;  | 
| 5819 | 620  | 
|
| 24922 | 621  | 
fun contract_abbrevs ctxt t =  | 
622  | 
let  | 
|
623  | 
val thy = theory_of ctxt;  | 
|
624  | 
val consts = consts_of ctxt;  | 
|
625  | 
    val Mode {abbrev, ...} = get_mode ctxt;
 | 
|
| 
30566
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
626  | 
val retrieve = Consts.retrieve_abbrevs consts (print_mode_value () @ [""]);  | 
| 
 
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
627  | 
fun match_abbrev u = Option.map #1 (get_first (Pattern.match_rew thy u) (retrieve u));  | 
| 24922 | 628  | 
in  | 
| 
40879
 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
629  | 
if abbrev orelse not (Config.get ctxt show_abbrevs) orelse not (can Term.type_of t) then t  | 
| 
35211
 
5d2fe4e09354
Use top-down rewriting to contract abbreviations.
 
berghofe 
parents: 
35141 
diff
changeset
 | 
630  | 
else Pattern.rewrite_term_top thy [] [match_abbrev] t  | 
| 24922 | 631  | 
end;  | 
632  | 
||
633  | 
||
| 24518 | 634  | 
(* patterns *)  | 
635  | 
||
| 32003 | 636  | 
fun prepare_patternT ctxt T =  | 
637  | 
let  | 
|
638  | 
    val Mode {pattern, schematic, ...} = get_mode ctxt;
 | 
|
639  | 
val _ =  | 
|
640  | 
pattern orelse schematic orelse  | 
|
641  | 
T |> Term.exists_subtype  | 
|
| 
38979
 
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
642  | 
(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
 | 
643  | 
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
 | 
644  | 
              error ("Illegal schematic type variable: " ^ Syntax.string_of_typ ctxt T)
 | 
| 32003 | 645  | 
| _ => false)  | 
646  | 
in T end;  | 
|
| 24518 | 647  | 
|
| 22712 | 648  | 
|
| 24505 | 649  | 
local  | 
| 6550 | 650  | 
|
| 56438 | 651  | 
val dummies =  | 
652  | 
  Config.bool (Config.declare ("Proof_Context.dummies", @{here}) (K (Config.Bool false)));
 | 
|
| 27259 | 653  | 
|
654  | 
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
 | 
655  | 
if Config.get ctxt dummies then t  | 
| 27259 | 656  | 
else Term.no_dummy_patterns t handle TERM _ => error "Illegal dummy pattern(s) in term";  | 
657  | 
||
| 24767 | 658  | 
fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1);  | 
| 6762 | 659  | 
|
| 27259 | 660  | 
in  | 
| 6550 | 661  | 
|
| 
39508
 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 
wenzelm 
parents: 
39507 
diff
changeset
 | 
662  | 
val allow_dummies = Config.put dummies true;  | 
| 24505 | 663  | 
|
| 24684 | 664  | 
fun prepare_patterns ctxt =  | 
| 24518 | 665  | 
  let val Mode {pattern, ...} = get_mode ctxt in
 | 
| 39296 | 666  | 
Type_Infer.fixate ctxt #>  | 
| 24767 | 667  | 
pattern ? Variable.polymorphic ctxt #>  | 
| 24684 | 668  | 
(map o Term.map_types) (prepare_patternT ctxt) #>  | 
| 27259 | 669  | 
(if pattern then prepare_dummies else map (check_dummies ctxt))  | 
| 24505 | 670  | 
end;  | 
671  | 
||
672  | 
end;  | 
|
673  | 
||
| 6550 | 674  | 
|
| 
42250
 
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
 
wenzelm 
parents: 
42242 
diff
changeset
 | 
675  | 
(* sort constraints *)  | 
| 27286 | 676  | 
|
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
677  | 
local  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
678  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
679  | 
fun prepare_sorts_env ctxt tys =  | 
| 27286 | 680  | 
let  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35616 
diff
changeset
 | 
681  | 
val tsig = tsig_of ctxt;  | 
| 45453 | 682  | 
val defaultS = Type.defaultS tsig;  | 
| 27286 | 683  | 
|
| 
53673
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
684  | 
    val dummy_var = ("'_dummy_", ~1);
 | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
685  | 
|
| 
49685
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
686  | 
fun constraint (xi, raw_S) env =  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
687  | 
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
 | 
688  | 
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
 | 
689  | 
else  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
690  | 
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
 | 
691  | 
handle Vartab.DUP _ =>  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
692  | 
              error ("Inconsistent sort constraints for type variable " ^
 | 
| 49691 | 693  | 
quote (Term.string_of_vname' xi) ^ Position.here_list ps)  | 
| 
49685
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
694  | 
end;  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
695  | 
|
| 45453 | 696  | 
val env =  | 
697  | 
(fold o fold_atyps)  | 
|
698  | 
(fn TFree (x, S) => constraint ((x, ~1), S)  | 
|
699  | 
| TVar v => constraint v  | 
|
700  | 
| _ => I) tys Vartab.empty;  | 
|
| 36152 | 701  | 
|
| 
53673
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
702  | 
fun get_sort xi raw_S =  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
703  | 
if xi = dummy_var then  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
704  | 
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
 | 
705  | 
else  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
706  | 
(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
 | 
707  | 
(NONE, NONE) => defaultS  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
708  | 
| (NONE, SOME S) => S  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
709  | 
| (SOME S, NONE) => S  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
710  | 
| (SOME S, SOME S') =>  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
711  | 
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
 | 
712  | 
else  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
713  | 
              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
 | 
714  | 
" 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
 | 
715  | 
" for type variable " ^ quote (Term.string_of_vname' xi)));  | 
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
716  | 
|
| 
49685
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
717  | 
fun add_report S pos reports =  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
718  | 
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
 | 
719  | 
(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
 | 
720  | 
else reports;  | 
| 
 
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
 
wenzelm 
parents: 
49674 
diff
changeset
 | 
721  | 
|
| 49691 | 722  | 
fun get_sort_reports xi raw_S =  | 
723  | 
let  | 
|
724  | 
val ps = #1 (Term_Position.decode_positionS raw_S);  | 
|
| 
53673
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
725  | 
val S = get_sort xi raw_S handle ERROR msg => error (msg ^ Position.here_list ps);  | 
| 49691 | 726  | 
in fold (add_report S) ps end;  | 
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
727  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
728  | 
val reports =  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
729  | 
(fold o fold_atyps)  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
730  | 
(fn T =>  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
731  | 
if Term_Position.is_positionT T then I  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
732  | 
else  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
733  | 
(case T of  | 
| 49691 | 734  | 
TFree (x, raw_S) => get_sort_reports (x, ~1) raw_S  | 
735  | 
| TVar (xi, raw_S) => get_sort_reports xi raw_S  | 
|
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
736  | 
| _ => I)) tys [];  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
737  | 
|
| 
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
 | 
738  | 
in (map #2 reports, get_sort) end;  | 
| 27286 | 739  | 
|
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
740  | 
fun replace_sortsT get_sort =  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
741  | 
map_atyps  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
742  | 
(fn T =>  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
743  | 
if Term_Position.is_positionT T then T  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
744  | 
else  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
745  | 
(case T of  | 
| 
53673
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
746  | 
TFree (x, raw_S) => TFree (x, get_sort (x, ~1) raw_S)  | 
| 
 
bcfd16f65014
treat all dummy type variables separately (in contrast to fca432074fb2);
 
wenzelm 
parents: 
53380 
diff
changeset
 | 
747  | 
| TVar (xi, raw_S) => TVar (xi, get_sort xi raw_S)  | 
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
748  | 
| _ => T));  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
749  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
750  | 
in  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
751  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
752  | 
fun prepare_sortsT ctxt tys =  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
753  | 
let val (sorting_report, get_sort) = prepare_sorts_env ctxt tys  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
754  | 
in (sorting_report, map (replace_sortsT get_sort) tys) end;  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
755  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
756  | 
fun prepare_sorts ctxt tms =  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
757  | 
let  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
758  | 
val tys = rev ((fold o fold_types) cons tms []);  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
759  | 
val (sorting_report, get_sort) = prepare_sorts_env ctxt tys;  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
760  | 
in (sorting_report, (map o map_types) (replace_sortsT get_sort) tms) end;  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
761  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
762  | 
fun check_tfree ctxt v =  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
763  | 
let  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
764  | 
val (sorting_report, [TFree a]) = prepare_sortsT ctxt [TFree v];  | 
| 
56294
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
56202 
diff
changeset
 | 
765  | 
val _ = if Context_Position.is_visible ctxt then Output.report sorting_report else ();  | 
| 
49674
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
766  | 
in a end;  | 
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
767  | 
|
| 
 
dbadb4d03cbc
report sort assignment of visible type variables;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
768  | 
end;  | 
| 36152 | 769  | 
|
| 5819 | 770  | 
|
771  | 
(* certify terms *)  | 
|
772  | 
||
| 10554 | 773  | 
local  | 
774  | 
||
| 24684 | 775  | 
fun gen_cert prop ctxt t =  | 
776  | 
t  | 
|
777  | 
|> expand_abbrevs ctxt  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
778  | 
|> (fn t' =>  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
779  | 
#1 (Sign.certify' prop (Context.Proof ctxt) false (consts_of ctxt) (theory_of ctxt) t')  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
780  | 
handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg);  | 
| 16501 | 781  | 
|
| 10554 | 782  | 
in  | 
| 8096 | 783  | 
|
| 24684 | 784  | 
val cert_term = gen_cert false;  | 
785  | 
val cert_prop = gen_cert true;  | 
|
| 10554 | 786  | 
|
787  | 
end;  | 
|
| 5819 | 788  | 
|
789  | 
||
| 
42405
 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 
wenzelm 
parents: 
42402 
diff
changeset
 | 
790  | 
(* check/uncheck *)  | 
| 22701 | 791  | 
|
| 
35616
 
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
792  | 
fun def_type ctxt =  | 
| 
 
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
793  | 
  let val Mode {pattern, ...} = get_mode ctxt
 | 
| 
 
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
794  | 
in Variable.def_type ctxt pattern end;  | 
| 
 
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
795  | 
|
| 24518 | 796  | 
fun standard_typ_check ctxt =  | 
| 45429 | 797  | 
map (cert_typ_mode (Type.get_mode ctxt) ctxt #> prepare_patternT ctxt);  | 
| 24922 | 798  | 
|
| 45429 | 799  | 
val standard_term_check_finish = prepare_patterns;  | 
| 24518 | 800  | 
|
| 45429 | 801  | 
fun standard_term_uncheck ctxt = map (contract_abbrevs ctxt);  | 
| 22701 | 802  | 
|
803  | 
||
| 9553 | 804  | 
|
| 21610 | 805  | 
(** export results **)  | 
| 21531 | 806  | 
|
| 20310 | 807  | 
fun common_export is_goal inner outer =  | 
808  | 
map (Assumption.export is_goal inner outer) #>  | 
|
809  | 
Variable.export inner outer;  | 
|
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
810  | 
|
| 20310 | 811  | 
val goal_export = common_export true;  | 
812  | 
val export = common_export false;  | 
|
| 12704 | 813  | 
|
| 21531 | 814  | 
fun export_morphism inner outer =  | 
815  | 
Assumption.export_morphism inner outer $>  | 
|
816  | 
Variable.export_morphism inner outer;  | 
|
817  | 
||
| 28396 | 818  | 
fun norm_export_morphism inner outer =  | 
819  | 
export_morphism inner outer $>  | 
|
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
820  | 
Morphism.thm_morphism "Proof_Context.norm_export" (Goal.norm_result outer);  | 
| 28396 | 821  | 
|
| 21531 | 822  | 
|
| 
15758
 
07e382399a96
binds/thms: do not store options, but delete from table;
 
wenzelm 
parents: 
15750 
diff
changeset
 | 
823  | 
|
| 
30757
 
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
 
wenzelm 
parents: 
30723 
diff
changeset
 | 
824  | 
(** term bindings **)  | 
| 5819 | 825  | 
|
| 8096 | 826  | 
(* simult_matches *)  | 
827  | 
||
| 19867 | 828  | 
fun simult_matches ctxt (t, pats) =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58668 
diff
changeset
 | 
829  | 
(case Seq.pull (Unify.matchers (Context.Proof ctxt) (map (rpair t) pats)) of  | 
| 19867 | 830  | 
NONE => error "Pattern match failed!"  | 
| 32032 | 831  | 
| SOME (env, _) => Vartab.fold (fn (v, (_, t)) => cons (v, t)) (Envir.term_env env) []);  | 
| 8096 | 832  | 
|
833  | 
||
| 
30757
 
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
 
wenzelm 
parents: 
30723 
diff
changeset
 | 
834  | 
(* auto_bind *)  | 
| 10810 | 835  | 
|
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
836  | 
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
 | 
837  | 
|
| 33386 | 838  | 
val auto_bind_goal = auto_bind Auto_Bind.goal;  | 
839  | 
val auto_bind_facts = auto_bind Auto_Bind.facts;  | 
|
| 7925 | 840  | 
|
| 5819 | 841  | 
|
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
842  | 
(* bind terms (non-schematic) *)  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
843  | 
|
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
844  | 
fun cert_maybe_bind_term (xi, t) ctxt =  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
845  | 
ctxt  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
846  | 
|> 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
 | 
847  | 
|
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
848  | 
val cert_bind_term = cert_maybe_bind_term o apsnd SOME;  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
849  | 
|
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
850  | 
|
| 60377 | 851  | 
(* match_bind *)  | 
| 5819 | 852  | 
|
| 8096 | 853  | 
local  | 
854  | 
||
| 24684 | 855  | 
fun gen_bind prep_terms gen raw_binds ctxt =  | 
| 5819 | 856  | 
let  | 
| 24684 | 857  | 
fun prep_bind (raw_pats, t) ctxt1 =  | 
858  | 
let  | 
|
859  | 
val T = Term.fastype_of t;  | 
|
860  | 
val ctxt2 = Variable.declare_term t ctxt1;  | 
|
861  | 
val pats = prep_terms (set_mode mode_pattern ctxt2) T raw_pats;  | 
|
862  | 
val binds = simult_matches ctxt2 (t, pats);  | 
|
863  | 
in (binds, ctxt2) end;  | 
|
| 7670 | 864  | 
|
| 24686 | 865  | 
val ts = prep_terms ctxt dummyT (map snd raw_binds);  | 
866  | 
val (binds, ctxt') = apfst flat (fold_map prep_bind (map fst raw_binds ~~ ts) ctxt);  | 
|
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
867  | 
val binds' =  | 
| 
19916
 
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
 
wenzelm 
parents: 
19897 
diff
changeset
 | 
868  | 
if gen then map #1 binds ~~ Variable.exportT_terms ctxt' ctxt (map #2 binds)  | 
| 
8616
 
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
 
wenzelm 
parents: 
8462 
diff
changeset
 | 
869  | 
else binds;  | 
| 18310 | 870  | 
val ctxt'' =  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
871  | 
tap (Variable.warn_extra_tfrees ctxt)  | 
| 18310 | 872  | 
(if gen then  | 
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
873  | 
ctxt (*sic!*)  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
874  | 
|> fold Variable.declare_term (map #2 binds')  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
875  | 
|> fold cert_bind_term binds'  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
876  | 
else ctxt' |> fold cert_bind_term binds');  | 
| 18310 | 877  | 
in (ts, ctxt'') end;  | 
| 8096 | 878  | 
|
879  | 
in  | 
|
| 5935 | 880  | 
|
| 24684 | 881  | 
fun read_terms ctxt T =  | 
| 39288 | 882  | 
map (Syntax.parse_term ctxt #> Type.constraint T) #> Syntax.check_terms ctxt;  | 
| 24684 | 883  | 
|
| 60377 | 884  | 
val match_bind = gen_bind (fn ctxt => fn _ => map (cert_term ctxt));  | 
885  | 
val match_bind_cmd = gen_bind read_terms;  | 
|
| 8096 | 886  | 
|
887  | 
end;  | 
|
| 5935 | 888  | 
|
889  | 
||
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
890  | 
(* propositions with patterns *)  | 
| 5935 | 891  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
892  | 
local  | 
| 8096 | 893  | 
|
| 60388 | 894  | 
fun prep_propp prep_props ctxt raw_args =  | 
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
895  | 
let  | 
| 60386 | 896  | 
val props = prep_props ctxt (maps (map fst) raw_args);  | 
| 60388 | 897  | 
val props_ctxt = fold Variable.declare_term props ctxt;  | 
898  | 
val patss = maps (map (prep_props (set_mode mode_pattern props_ctxt) o snd)) raw_args;  | 
|
| 5935 | 899  | 
|
| 
60387
 
76359ff1090f
more careful treatment of term bindings in 'obtain' proof body;
 
wenzelm 
parents: 
60386 
diff
changeset
 | 
900  | 
val propps = unflat raw_args (props ~~ patss);  | 
| 60388 | 901  | 
val binds = (maps o maps) (simult_matches props_ctxt) propps;  | 
902  | 
in (map (map fst) propps, binds) end;  | 
|
| 8096 | 903  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
904  | 
in  | 
| 
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
905  | 
|
| 60388 | 906  | 
val cert_propp = prep_propp (map o cert_prop);  | 
907  | 
val read_propp = prep_propp Syntax.read_props;  | 
|
| 6789 | 908  | 
|
| 
10465
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
909  | 
end;  | 
| 
 
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
 
wenzelm 
parents: 
10381 
diff
changeset
 | 
910  | 
|
| 6789 | 911  | 
|
| 5819 | 912  | 
|
913  | 
(** theorems **)  | 
|
914  | 
||
| 18042 | 915  | 
(* fact_tac *)  | 
916  | 
||
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
917  | 
local  | 
| 
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
918  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58668 
diff
changeset
 | 
919  | 
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
 | 
920  | 
  PRIMSEQ (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true}
 | 
| 60367 | 921  | 
(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
 | 
922  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58668 
diff
changeset
 | 
923  | 
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
 | 
924  | 
| 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
 | 
925  | 
(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
 | 
926  | 
comp_incr_tac ctxt ths i;  | 
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
927  | 
|
| 60489 | 928  | 
val vacuous_facts = [Drule.termI];  | 
929  | 
||
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
930  | 
in  | 
| 18042 | 931  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58668 
diff
changeset
 | 
932  | 
fun fact_tac ctxt facts = Goal.norm_hhf_tac ctxt THEN' comp_incr_tac ctxt facts;  | 
| 18122 | 933  | 
|
| 27867 | 934  | 
fun potential_facts ctxt prop =  | 
| 60489 | 935  | 
let  | 
936  | 
val body = Term.strip_all_body prop;  | 
|
937  | 
val vacuous = filter (fn th => Term.could_unify (body, Thm.concl_of th)) vacuous_facts;  | 
|
938  | 
in Facts.could_unify (facts_of ctxt) body @ vacuous end;  | 
|
| 27867 | 939  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
940  | 
fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => fact_tac ctxt (potential_facts ctxt goal) i);  | 
| 18042 | 941  | 
|
| 
52223
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
942  | 
end;  | 
| 
 
5bb6ae8acb87
tuned signature -- more explicit flags for low-level Thm.bicompose;
 
wenzelm 
parents: 
52156 
diff
changeset
 | 
943  | 
|
| 18042 | 944  | 
|
| 
62078
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
945  | 
(* lookup facts *)  | 
| 5819 | 946  | 
|
| 61902 | 947  | 
fun lookup_fact ctxt name =  | 
948  | 
let  | 
|
949  | 
val context = Context.Proof ctxt;  | 
|
950  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
951  | 
in  | 
|
952  | 
(case Facts.lookup context (facts_of ctxt) name of  | 
|
953  | 
NONE => Facts.lookup context (Global_Theory.facts_of thy) name  | 
|
954  | 
| some => some)  | 
|
955  | 
end;  | 
|
956  | 
||
| 
62078
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
957  | 
|
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
958  | 
(* retrieve facts *)  | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
959  | 
|
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
960  | 
val dynamic_facts_dummy =  | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
961  | 
  Config.bool (Config.declare ("dynamic_facts_dummy_", @{here}) (fn _ => Config.Bool false));
 | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
962  | 
|
| 26361 | 963  | 
local  | 
964  | 
||
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
965  | 
fun retrieve_global context =  | 
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
966  | 
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
 | 
967  | 
|
| 
56156
 
47015872e795
clarified retrieve_generic: local error takes precedence, which is relevant for completion;
 
wenzelm 
parents: 
56155 
diff
changeset
 | 
968  | 
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
 | 
969  | 
(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
 | 
970  | 
(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
 | 
971  | 
| 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
 | 
972  | 
|
| 
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
973  | 
fun retrieve pick context (Facts.Fact s) =  | 
| 16501 | 974  | 
let  | 
| 
56140
 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
975  | 
val ctxt = Context.the_proof context;  | 
| 59795 | 976  | 
val pos = Syntax.read_input_pos s;  | 
| 42502 | 977  | 
val prop =  | 
978  | 
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
 | 
979  | 
|> singleton (Variable.polymorphic ctxt);  | 
| 48992 | 980  | 
fun err msg = error (msg ^ Position.here pos ^ ":\n" ^ Syntax.string_of_term ctxt prop);  | 
| 27867 | 981  | 
|
| 42502 | 982  | 
val (prop', _) = Term.replace_dummy_patterns prop (Variable.maxidx_of ctxt + 1);  | 
| 27867 | 983  | 
fun prove_fact th =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
984  | 
Goal.prove ctxt [] [] prop' (K (ALLGOALS (fact_tac ctxt [th])));  | 
| 42502 | 985  | 
val results = map_filter (try prove_fact) (potential_facts ctxt prop');  | 
| 
56140
 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
986  | 
val thm =  | 
| 42502 | 987  | 
(case distinct Thm.eq_thm_prop results of  | 
| 
56140
 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
988  | 
[thm] => thm  | 
| 42502 | 989  | 
| [] => err "Failed to retrieve literal fact"  | 
990  | 
| _ => err "Ambiguous specification of literal fact");  | 
|
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
991  | 
      in pick true ("", Position.none) [thm] end
 | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
992  | 
| retrieve pick context (Facts.Named ((xname, pos), sel)) =  | 
| 18042 | 993  | 
let  | 
| 
56140
 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
994  | 
val thy = Context.theory_of context;  | 
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
995  | 
        fun immediate thm = {name = xname, static = true, thms = [Thm.transfer thy thm]};
 | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
996  | 
        val {name, static, thms} =
 | 
| 
56140
 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 
wenzelm 
parents: 
56139 
diff
changeset
 | 
997  | 
(case xname of  | 
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
998  | 
"" => immediate Drule.dummy_thm  | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
999  | 
| "_" => immediate Drule.asm_rl  | 
| 
56141
 
c06202417c4a
back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
 
wenzelm 
parents: 
56140 
diff
changeset
 | 
1000  | 
| _ => retrieve_generic context (xname, pos));  | 
| 
62078
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
1001  | 
val thms' =  | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
1002  | 
if not static andalso Config.get_generic context dynamic_facts_dummy  | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
1003  | 
then [Drule.free_dummy_thm]  | 
| 
 
76399b8fde4d
more systematic treatment of dynamic facts, when forming closure;
 
wenzelm 
parents: 
61902 
diff
changeset
 | 
1004  | 
else Facts.select (Facts.Named ((name, pos), sel)) thms;  | 
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1005  | 
in pick (static orelse is_some sel) (name, pos) thms' end;  | 
| 5819 | 1006  | 
|
| 26361 | 1007  | 
in  | 
| 
26346
 
17debd2fff8e
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
1008  | 
|
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1009  | 
val get_fact_generic =  | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1010  | 
retrieve (fn static => fn (name, _) => fn thms =>  | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1011  | 
(if static then NONE else SOME name, thms));  | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1012  | 
|
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
57928 
diff
changeset
 | 
1013  | 
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
 | 
1014  | 
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
 | 
1015  | 
|
| 26361 | 1016  | 
fun get_thms ctxt = get_fact ctxt o Facts.named;  | 
1017  | 
fun get_thm ctxt = get_fact_single ctxt o Facts.named;  | 
|
1018  | 
||
1019  | 
end;  | 
|
| 5819 | 1020  | 
|
1021  | 
||
| 26284 | 1022  | 
(* facts *)  | 
| 5819 | 1023  | 
|
| 
61811
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
1024  | 
fun add_thms_dynamic arg ctxt =  | 
| 
 
1530a0f19539
added Proof_Context.add_thms_dynamic, which is potentially useful for Eisbach;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
1025  | 
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
 | 
1026  | 
|
| 
28082
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
1027  | 
local  | 
| 
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
1028  | 
|
| 28965 | 1029  | 
fun update_thms _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b))  | 
| 49747 | 1030  | 
| update_thms flags (b, SOME ths) ctxt = ctxt |> map_facts  | 
1031  | 
(Facts.add_static (Context.Proof ctxt) flags (b, ths) #> snd);  | 
|
| 5819 | 1032  | 
|
| 
30761
 
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
 
wenzelm 
parents: 
30758 
diff
changeset
 | 
1033  | 
in  | 
| 
 
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
 
wenzelm 
parents: 
30758 
diff
changeset
 | 
1034  | 
|
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
1035  | 
fun note_thmss kind = fold_map (fn ((b, more_atts), raw_facts) => fn ctxt =>  | 
| 5819 | 1036  | 
let  | 
| 28965 | 1037  | 
val name = full_name ctxt b;  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39508 
diff
changeset
 | 
1038  | 
val facts = Global_Theory.name_thmss false name raw_facts;  | 
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
1039  | 
fun app (ths, atts) =  | 
| 
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
1040  | 
fold_map (Thm.proof_attributes (surround (Thm.kind kind) (atts @ more_atts))) ths;  | 
| 
21443
 
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
 
wenzelm 
parents: 
21370 
diff
changeset
 | 
1041  | 
val (res, ctxt') = fold_map app facts ctxt;  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
39508 
diff
changeset
 | 
1042  | 
val thms = Global_Theory.name_thms false false name (flat res);  | 
| 
24388
 
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
 
wenzelm 
parents: 
24371 
diff
changeset
 | 
1043  | 
    val Mode {stmt, ...} = get_mode ctxt;
 | 
| 49747 | 1044  | 
  in ((name, thms), ctxt' |> update_thms {strict = false, index = stmt} (b, SOME thms)) end);
 | 
| 12711 | 1045  | 
|
| 49747 | 1046  | 
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
 | 
1047  | 
|> map_naming (K Name_Space.local_naming)  | 
| 33383 | 1048  | 
|> Context_Position.set_visible false  | 
| 49747 | 1049  | 
  |> update_thms {strict = false, index = index} (apfst Binding.name thms)
 | 
| 33383 | 1050  | 
|> Context_Position.restore_visible ctxt  | 
| 28417 | 1051  | 
|> restore_naming ctxt;  | 
| 
28082
 
37350f301128
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
1052  | 
|
| 12711 | 1053  | 
end;  | 
| 9196 | 1054  | 
|
| 5819 | 1055  | 
|
1056  | 
||
| 35680 | 1057  | 
(** basic logical entities **)  | 
| 
17360
 
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
 
wenzelm 
parents: 
17221 
diff
changeset
 | 
1058  | 
|
| 8096 | 1059  | 
(* variables *)  | 
1060  | 
||
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
1061  | 
fun declare_var (x, opt_T, mx) ctxt =  | 
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42267 
diff
changeset
 | 
1062  | 
let val T = (case opt_T of SOME T => T | NONE => Mixfix.mixfixT mx)  | 
| 20163 | 1063  | 
in ((x, T, mx), ctxt |> Variable.declare_constraints (Free (x, T))) end;  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
1064  | 
|
| 
57486
 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
1065  | 
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
 | 
1066  | 
let  | 
| 
 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
1067  | 
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
 | 
1068  | 
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
 | 
1069  | 
val _ =  | 
| 
 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
1070  | 
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
 | 
1071  | 
      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
 | 
1072  | 
in x end;  | 
| 
 
2131b6633529
check 'case' variable bindings as for 'fix', which means internal names are rejected as usual;
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
1073  | 
|
| 10381 | 1074  | 
local  | 
1075  | 
||
| 60379 | 1076  | 
fun prep_var prep_typ internal (b, raw_T, mx) ctxt =  | 
1077  | 
let  | 
|
1078  | 
val x = check_var internal b;  | 
|
1079  | 
fun cond_tvars T =  | 
|
1080  | 
if internal then T  | 
|
1081  | 
else Type.no_tvars T handle TYPE (msg, _, _) => error msg;  | 
|
| 60468 | 1082  | 
val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T;  | 
| 60379 | 1083  | 
val (_, ctxt') = ctxt |> declare_var (x, opt_T, mx);  | 
1084  | 
in ((b, opt_T, mx), ctxt') end;  | 
|
| 8096 | 1085  | 
|
| 10381 | 1086  | 
in  | 
1087  | 
||
| 60379 | 1088  | 
val read_var = prep_var Syntax.read_typ false;  | 
| 60468 | 1089  | 
val cert_var = prep_var cert_typ true;  | 
| 8096 | 1090  | 
|
| 10381 | 1091  | 
end;  | 
1092  | 
||
| 8096 | 1093  | 
|
| 21744 | 1094  | 
(* notation *)  | 
1095  | 
||
| 24949 | 1096  | 
local  | 
1097  | 
||
| 
35429
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
1098  | 
fun type_syntax (Type (c, args), mx) =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
1099  | 
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
 | 
1100  | 
| type_syntax _ = NONE;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1101  | 
|
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1102  | 
fun const_syntax _ (Free (x, T), mx) = SOME (Local_Syntax.Fixed, (x, T, mx))  | 
| 24949 | 1103  | 
| const_syntax ctxt (Const (c, _), mx) =  | 
| 35255 | 1104  | 
(case try (Consts.type_scheme (consts_of ctxt)) c of  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
1105  | 
SOME T => SOME (Local_Syntax.Const, (Lexicon.mark_const c, T, mx))  | 
| 35255 | 1106  | 
| NONE => NONE)  | 
| 24949 | 1107  | 
| const_syntax _ _ = NONE;  | 
1108  | 
||
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1109  | 
fun gen_notation syntax add mode args ctxt =  | 
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
1110  | 
ctxt |> map_syntax_idents  | 
| 
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
1111  | 
(Local_Syntax.update_modesyntax ctxt add mode (map_filter (syntax ctxt) args));  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1112  | 
|
| 24949 | 1113  | 
in  | 
| 21744 | 1114  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1115  | 
val type_notation = gen_notation (K type_syntax);  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1116  | 
val notation = gen_notation const_syntax;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1117  | 
|
| 47247 | 1118  | 
fun generic_type_notation add mode args phi =  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1119  | 
let  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1120  | 
val args' = args |> map_filter (fn (T, mx) =>  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1121  | 
let  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1122  | 
val T' = Morphism.typ phi T;  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1123  | 
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
 | 
1124  | 
in if similar then SOME (T', mx) else NONE end);  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1125  | 
in Context.mapping (Sign.type_notation add mode args') (type_notation add mode args') end;  | 
| 24949 | 1126  | 
|
| 47247 | 1127  | 
fun generic_notation add mode args phi =  | 
| 
33537
 
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
1128  | 
let  | 
| 
 
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
1129  | 
val args' = args |> map_filter (fn (t, mx) =>  | 
| 
 
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
1130  | 
let val t' = Morphism.term phi t  | 
| 
 
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
1131  | 
in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end);  | 
| 24949 | 1132  | 
in Context.mapping (Sign.notation add mode args') (notation add mode args') end;  | 
1133  | 
||
| 35680 | 1134  | 
end;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35399 
diff
changeset
 | 
1135  | 
|
| 35680 | 1136  | 
|
1137  | 
(* aliases *)  | 
|
1138  | 
||
| 36450 | 1139  | 
fun class_alias b c ctxt = (map_tsig o apfst) (Type.class_alias (naming_of ctxt) b c) ctxt;  | 
1140  | 
fun type_alias b c ctxt = (map_tsig o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt;  | 
|
| 35680 | 1141  | 
fun const_alias b c ctxt = (map_consts o apfst) (Consts.alias (naming_of ctxt) b c) ctxt;  | 
| 57887 | 1142  | 
fun fact_alias b c ctxt = map_facts (Facts.alias (naming_of ctxt) b c) ctxt;  | 
| 21744 | 1143  | 
|
1144  | 
||
| 24767 | 1145  | 
(* local constants *)  | 
1146  | 
||
1147  | 
fun add_const_constraint (c, opt_T) ctxt =  | 
|
1148  | 
let  | 
|
1149  | 
fun prepT raw_T =  | 
|
1150  | 
let val T = cert_typ ctxt raw_T  | 
|
1151  | 
in cert_term ctxt (Const (c, T)); T end;  | 
|
| 25039 | 1152  | 
in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end;  | 
| 19001 | 1153  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33165 
diff
changeset
 | 
1154  | 
fun add_abbrev mode (b, raw_t) ctxt =  | 
| 19001 | 1155  | 
let  | 
| 
24675
 
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
1156  | 
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
 | 
1157  | 
      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
 | 
1158  | 
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
 | 
1159  | 
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
 | 
1160  | 
|> Consts.abbreviate (Context.Proof ctxt) (tsig_of ctxt) mode (b, t);  | 
| 19001 | 1161  | 
in  | 
1162  | 
ctxt  | 
|
| 25039 | 1163  | 
|> (map_consts o apfst) (K consts')  | 
| 21803 | 1164  | 
|> Variable.declare_term rhs  | 
1165  | 
|> pair (lhs, rhs)  | 
|
| 21704 | 1166  | 
end;  | 
| 19001 | 1167  | 
|
| 25052 | 1168  | 
fun revert_abbrev mode c = (map_consts o apfst) (Consts.revert_abbrev mode c);  | 
1169  | 
||
| 47275 | 1170  | 
fun generic_add_abbrev mode arg =  | 
1171  | 
Context.mapping_result (Sign.add_abbrev mode arg) (add_abbrev mode arg);  | 
|
1172  | 
||
1173  | 
fun generic_revert_abbrev mode arg =  | 
|
1174  | 
Context.mapping (Sign.revert_abbrev mode arg) (revert_abbrev mode arg);  | 
|
1175  | 
||
| 19001 | 1176  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1177  | 
(* fixes *)  | 
| 5819 | 1178  | 
|
| 60469 | 1179  | 
local  | 
1180  | 
||
1181  | 
fun gen_fixes prep_var raw_vars ctxt =  | 
|
1182  | 
let  | 
|
1183  | 
val (vars, _) = fold_map prep_var raw_vars ctxt;  | 
|
1184  | 
val (xs, ctxt') = Variable.add_fixes_binding (map #1 vars) ctxt;  | 
|
1185  | 
in  | 
|
1186  | 
ctxt'  | 
|
1187  | 
|> fold_map declare_var (map2 (fn x => fn (_, T, mx) => (x, T, mx)) xs vars)  | 
|
1188  | 
|-> (map_syntax_idents o Local_Syntax.add_syntax ctxt o map (pair Local_Syntax.Fixed))  | 
|
1189  | 
|> pair xs  | 
|
| 
42491
 
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
 
wenzelm 
parents: 
42488 
diff
changeset
 | 
1190  | 
end;  | 
| 5819 | 1191  | 
|
| 60469 | 1192  | 
in  | 
1193  | 
||
1194  | 
val add_fixes = gen_fixes cert_var;  | 
|
1195  | 
val add_fixes_cmd = gen_fixes read_var;  | 
|
1196  | 
||
1197  | 
end;  | 
|
1198  | 
||
| 
9291
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1199  | 
|
| 
 
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
 
wenzelm 
parents: 
9274 
diff
changeset
 | 
1200  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1201  | 
(** assumptions **)  | 
| 18187 | 1202  | 
|
| 20209 | 1203  | 
local  | 
1204  | 
||
| 60388 | 1205  | 
fun gen_assms prep_propp exp args ctxt =  | 
| 20209 | 1206  | 
let  | 
| 60388 | 1207  | 
val (propss, binds) = prep_propp ctxt (map snd args);  | 
1208  | 
val props = flat propss;  | 
|
| 
20234
 
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
 
wenzelm 
parents: 
20209 
diff
changeset
 | 
1209  | 
in  | 
| 60388 | 1210  | 
ctxt  | 
1211  | 
|> fold Variable.declare_term props  | 
|
1212  | 
|> tap (Variable.warn_extra_tfrees ctxt)  | 
|
1213  | 
|> fold_burrow (Assumption.add_assms exp o map (Thm.cterm_of ctxt)) propss  | 
|
1214  | 
|-> (fn premss =>  | 
|
1215  | 
auto_bind_facts props  | 
|
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
1216  | 
#> fold Variable.bind_term binds  | 
| 60388 | 1217  | 
#> 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
 | 
1218  | 
end;  | 
| 20209 | 1219  | 
|
1220  | 
in  | 
|
1221  | 
||
| 60388 | 1222  | 
val add_assms = gen_assms cert_propp;  | 
1223  | 
val add_assms_cmd = gen_assms read_propp;  | 
|
| 20209 | 1224  | 
|
1225  | 
end;  | 
|
1226  | 
||
1227  | 
||
| 5819 | 1228  | 
|
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1229  | 
(** cases **)  | 
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1230  | 
|
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1231  | 
fun dest_cases ctxt =  | 
| 56025 | 1232  | 
Name_Space.fold_table (fn (a, (c, i)) => cons (i, (a, c))) (cases_of ctxt) []  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
1233  | 
|> sort (int_ord o apply2 #1) |> map #2;  | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1234  | 
|
| 16147 | 1235  | 
local  | 
1236  | 
||
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
1237  | 
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
 | 
1238  | 
| drop_schematic b = b;  | 
| 
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
1239  | 
|
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1240  | 
fun update_case _ _ ("", _) res = res
 | 
| 56025 | 1241  | 
| update_case _ _ (name, NONE) (cases, index) =  | 
1242  | 
(Name_Space.del_table name cases, index)  | 
|
| 60573 | 1243  | 
| update_case context legacy (name, SOME c) (cases, index) =  | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1244  | 
let  | 
| 60573 | 1245  | 
val binding = Binding.name name |> legacy ? Binding.concealed;  | 
1246  | 
val (_, cases') = cases  | 
|
1247  | 
          |> Name_Space.define context false (binding, ((c, {legacy = legacy}), index));
 | 
|
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1248  | 
val index' = index + 1;  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1249  | 
in (cases', index') end;  | 
| 18609 | 1250  | 
|
| 60573 | 1251  | 
fun update_cases' legacy args ctxt =  | 
1252  | 
let  | 
|
1253  | 
val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming);  | 
|
1254  | 
val cases = cases_of ctxt;  | 
|
1255  | 
val index = Name_Space.fold_table (fn _ => Integer.add 1) cases 0;  | 
|
1256  | 
val (cases', _) = fold (update_case context legacy) args (cases, index);  | 
|
1257  | 
in map_cases (K cases') ctxt end;  | 
|
1258  | 
||
| 
42501
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
1259  | 
fun fix (b, T) ctxt =  | 
| 
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
1260  | 
let val ([x], ctxt') = add_fixes [(b, SOME T, NoSyn)] ctxt  | 
| 
 
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
1261  | 
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
 | 
1262  | 
|
| 16147 | 1263  | 
in  | 
1264  | 
||
| 60573 | 1265  | 
val update_cases = update_cases' false;  | 
1266  | 
val update_cases_legacy = update_cases' true;  | 
|
| 18609 | 1267  | 
|
1268  | 
fun case_result c ctxt =  | 
|
1269  | 
let  | 
|
| 33368 | 1270  | 
    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
 | 
1271  | 
val (ts, ctxt') = ctxt |> fold_map fix fixes;  | 
| 33368 | 1272  | 
    val Rule_Cases.Case {assumes, binds, cases, ...} = Rule_Cases.apply ts c;
 | 
| 18609 | 1273  | 
in  | 
1274  | 
ctxt'  | 
|
| 
60408
 
1fd46ced2fa8
more uniform treatment of auto bindings vs. explicit user bindings;
 
wenzelm 
parents: 
60407 
diff
changeset
 | 
1275  | 
|> fold (cert_maybe_bind_term o drop_schematic) binds  | 
| 60573 | 1276  | 
|> update_cases (map (apsnd SOME) cases)  | 
| 18609 | 1277  | 
|> pair (assumes, (binds, cases))  | 
1278  | 
end;  | 
|
1279  | 
||
1280  | 
val apply_case = apfst fst oo case_result;  | 
|
1281  | 
||
| 60629 | 1282  | 
fun check_case ctxt internal (name, pos) param_specs =  | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1283  | 
let  | 
| 60573 | 1284  | 
    val (_, ((Rule_Cases.Case {fixes, assumes, binds, cases}, {legacy}), _)) =
 | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1285  | 
Name_Space.check (Context.Proof ctxt) (cases_of ctxt) (name, pos);  | 
| 60581 | 1286  | 
val _ =  | 
1287  | 
if legacy then  | 
|
1288  | 
        legacy_feature ("Bad case " ^ quote name ^ Position.here pos ^
 | 
|
| 61168 | 1289  | 
" -- use proof method \"goal_cases\" instead")  | 
| 60581 | 1290  | 
else ();  | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1291  | 
|
| 60629 | 1292  | 
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
 | 
1293  | 
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
 | 
1294  | 
| replace [] ys = ys  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1295  | 
| replace (_ :: _) [] =  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1296  | 
          error ("Too many parameters for case " ^ quote name ^ Position.here pos);
 | 
| 60629 | 1297  | 
val fixes' = replace param_specs fixes;  | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1298  | 
val binds' = map drop_schematic binds;  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1299  | 
in  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1300  | 
if null (fold (Term.add_tvarsT o snd) fixes []) andalso  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1301  | 
null (fold (fold Term.add_vars o snd) assumes []) then  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1302  | 
        Rule_Cases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases}
 | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1303  | 
    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
 | 
1304  | 
end;  | 
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1305  | 
|
| 16147 | 1306  | 
end;  | 
| 
8373
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1307  | 
|
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1308  | 
|
| 
 
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
 
wenzelm 
parents: 
8186 
diff
changeset
 | 
1309  | 
|
| 10810 | 1310  | 
(** print context information **)  | 
1311  | 
||
| 
12072
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1312  | 
(* local syntax *)  | 
| 
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1313  | 
|
| 12093 | 1314  | 
val print_syntax = Syntax.print_syntax o syn_of;  | 
| 
12072
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1315  | 
|
| 
 
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
 
wenzelm 
parents: 
12066 
diff
changeset
 | 
1316  | 
|
| 21728 | 1317  | 
(* abbreviations *)  | 
| 18971 | 1318  | 
|
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1319  | 
fun pretty_abbrevs verbose show_globals ctxt =  | 
| 18971 | 1320  | 
let  | 
| 56025 | 1321  | 
val space = const_space ctxt;  | 
| 56062 | 1322  | 
val (constants, global_constants) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
1323  | 
apply2 (#constants o Consts.dest) (#consts (rep_data ctxt));  | 
| 56062 | 1324  | 
val globals = Symtab.make global_constants;  | 
| 21803 | 1325  | 
fun add_abbr (_, (_, NONE)) = I  | 
| 25406 | 1326  | 
| add_abbr (c, (T, SOME t)) =  | 
| 21728 | 1327  | 
if not show_globals andalso Symtab.defined globals c then I  | 
1328  | 
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
 | 
1329  | 
val abbrevs = Name_Space.markup_entries verbose ctxt space (fold add_abbr constants []);  | 
| 18971 | 1330  | 
in  | 
| 35141 | 1331  | 
if null abbrevs then []  | 
| 21728 | 1332  | 
else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]  | 
| 18971 | 1333  | 
end;  | 
1334  | 
||
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1335  | 
fun print_abbrevs verbose = Pretty.writeln_chunks o pretty_abbrevs verbose true;  | 
| 21728 | 1336  | 
|
| 18971 | 1337  | 
|
| 10810 | 1338  | 
(* term bindings *)  | 
1339  | 
||
| 
57415
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
1340  | 
fun pretty_term_bindings ctxt =  | 
| 10810 | 1341  | 
let  | 
| 
19897
 
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
 
wenzelm 
parents: 
19882 
diff
changeset
 | 
1342  | 
val binds = Variable.binds_of ctxt;  | 
| 21728 | 1343  | 
fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));  | 
| 10810 | 1344  | 
in  | 
| 35141 | 1345  | 
if Vartab.is_empty binds then []  | 
| 
15758
 
07e382399a96
binds/thms: do not store options, but delete from table;
 
wenzelm 
parents: 
15750 
diff
changeset
 | 
1346  | 
else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]  | 
| 10810 | 1347  | 
end;  | 
1348  | 
||
1349  | 
||
| 56155 | 1350  | 
(* local facts *)  | 
| 10810 | 1351  | 
|
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1352  | 
fun pretty_local_facts verbose ctxt =  | 
| 20012 | 1353  | 
let  | 
| 56155 | 1354  | 
val facts = facts_of ctxt;  | 
1355  | 
val props = Facts.props facts;  | 
|
1356  | 
val local_facts =  | 
|
| 35141 | 1357  | 
      (if null props then [] else [("<unnamed>", props)]) @
 | 
| 
56158
 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 
wenzelm 
parents: 
56156 
diff
changeset
 | 
1358  | 
Facts.dest_static verbose [Global_Theory.facts_of (theory_of ctxt)] facts;  | 
| 20012 | 1359  | 
in  | 
| 56155 | 1360  | 
if null local_facts then []  | 
1361  | 
else  | 
|
1362  | 
[Pretty.big_list "local facts:"  | 
|
| 
60924
 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 
wenzelm 
parents: 
60799 
diff
changeset
 | 
1363  | 
(map #1 (sort_by (#1 o #2) (map (`(pretty_fact ctxt)) local_facts)))]  | 
| 20012 | 1364  | 
end;  | 
| 10810 | 1365  | 
|
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1366  | 
fun print_local_facts verbose ctxt =  | 
| 
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1367  | 
Pretty.writeln_chunks (pretty_local_facts verbose ctxt);  | 
| 10810 | 1368  | 
|
1369  | 
||
1370  | 
(* local contexts *)  | 
|
1371  | 
||
| 26722 | 1372  | 
local  | 
1373  | 
||
1374  | 
fun pretty_case (name, (fixes, ((asms, (lets, cs)), ctxt))) =  | 
|
| 10810 | 1375  | 
let  | 
| 24922 | 1376  | 
val prt_term = Syntax.pretty_term ctxt;  | 
| 12057 | 1377  | 
|
| 10810 | 1378  | 
fun prt_let (xi, t) = Pretty.block  | 
| 10818 | 1379  | 
[Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1,  | 
| 10810 | 1380  | 
Pretty.quote (prt_term t)];  | 
1381  | 
||
| 60565 | 1382  | 
fun prt_asm (a, ts) =  | 
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59990 
diff
changeset
 | 
1383  | 
Pretty.block (Pretty.breaks  | 
| 60565 | 1384  | 
((if a = "" then [] else [Pretty.str a, Pretty.str ":"]) @  | 
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59990 
diff
changeset
 | 
1385  | 
map (Pretty.quote o prt_term) ts));  | 
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1386  | 
|
| 10810 | 1387  | 
fun prt_sect _ _ _ [] = []  | 
| 56041 | 1388  | 
| prt_sect head sep prt xs =  | 
1389  | 
[Pretty.block (Pretty.breaks (head ::  | 
|
| 35141 | 1390  | 
flat (separate sep (map (single o prt) xs))))];  | 
| 26722 | 1391  | 
in  | 
1392  | 
Pretty.block (Pretty.fbreaks  | 
|
| 10810 | 1393  | 
(Pretty.str (name ^ ":") ::  | 
| 56041 | 1394  | 
prt_sect (Pretty.keyword1 "fix") [] (Pretty.str o Binding.name_of o fst) fixes @  | 
1395  | 
prt_sect (Pretty.keyword1 "let") [Pretty.keyword2 "and"] prt_let  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19422 
diff
changeset
 | 
1396  | 
(map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @  | 
| 
13425
 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 
wenzelm 
parents: 
13415 
diff
changeset
 | 
1397  | 
(if forall (null o #2) asms then []  | 
| 56041 | 1398  | 
else prt_sect (Pretty.keyword1 "assume") [Pretty.keyword2 "and"] prt_asm asms) @  | 
1399  | 
prt_sect (Pretty.str "subcases:") [] (Pretty.str o fst) cs))  | 
|
| 26722 | 1400  | 
end;  | 
| 16540 | 1401  | 
|
| 26722 | 1402  | 
in  | 
1403  | 
||
1404  | 
fun pretty_cases ctxt =  | 
|
1405  | 
let  | 
|
| 60573 | 1406  | 
    fun mk_case (_, (_, {legacy = true})) = NONE
 | 
1407  | 
      | mk_case (name, (c as Rule_Cases.Case {fixes, ...}, {legacy = false})) =
 | 
|
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1408  | 
SOME (name, (fixes, case_result c ctxt));  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
52223 
diff
changeset
 | 
1409  | 
val cases = dest_cases ctxt |> map_filter mk_case;  | 
| 10810 | 1410  | 
in  | 
| 35141 | 1411  | 
if null cases then []  | 
| 26722 | 1412  | 
else [Pretty.big_list "cases:" (map pretty_case cases)]  | 
| 10810 | 1413  | 
end;  | 
1414  | 
||
| 26722 | 1415  | 
end;  | 
1416  | 
||
| 10810 | 1417  | 
|
| 12057 | 1418  | 
(* core context *)  | 
| 10810 | 1419  | 
|
| 56438 | 1420  | 
val debug =  | 
1421  | 
  Config.bool (Config.declare ("Proof_Context.debug", @{here}) (K (Config.Bool false)));
 | 
|
1422  | 
||
1423  | 
val verbose =  | 
|
1424  | 
  Config.bool (Config.declare ("Proof_Context.verbose", @{here}) (K (Config.Bool false)));
 | 
|
| 10810 | 1425  | 
|
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1426  | 
fun pretty_ctxt ctxt =  | 
| 
42717
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42502 
diff
changeset
 | 
1427  | 
if not (Config.get ctxt debug) then []  | 
| 20310 | 1428  | 
else  | 
1429  | 
let  | 
|
| 24922 | 1430  | 
val prt_term = Syntax.pretty_term ctxt;  | 
| 12057 | 1431  | 
|
| 20310 | 1432  | 
(*structures*)  | 
| 
59152
 
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
 
wenzelm 
parents: 
59066 
diff
changeset
 | 
1433  | 
      val {structs, ...} = Syntax_Trans.get_idents ctxt;
 | 
| 35139 | 1434  | 
val prt_structs =  | 
1435  | 
if null structs then []  | 
|
| 20310 | 1436  | 
else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::  | 
1437  | 
Pretty.commas (map Pretty.str structs))];  | 
|
| 12093 | 1438  | 
|
| 20310 | 1439  | 
(*fixes*)  | 
1440  | 
fun prt_fix (x, x') =  | 
|
1441  | 
if x = x' then Pretty.str x  | 
|
1442  | 
else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')];  | 
|
1443  | 
val fixes =  | 
|
| 55948 | 1444  | 
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
 | 
1445  | 
(Variable.dest_fixes ctxt);  | 
| 35139 | 1446  | 
val prt_fixes =  | 
1447  | 
if null fixes then []  | 
|
| 20310 | 1448  | 
else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::  | 
1449  | 
Pretty.commas (map prt_fix fixes))];  | 
|
| 12057 | 1450  | 
|
| 60413 | 1451  | 
(*assumptions*)  | 
1452  | 
val prt_assms =  | 
|
| 51584 | 1453  | 
(case Assumption.all_prems_of ctxt of  | 
1454  | 
[] => []  | 
|
| 60413 | 1455  | 
        | prems => [Pretty.big_list "assumptions:" [pretty_fact ctxt ("", prems)]]);
 | 
1456  | 
in prt_structs @ prt_fixes @ prt_assms end;  | 
|
| 10810 | 1457  | 
|
1458  | 
||
1459  | 
(* main context *)  | 
|
1460  | 
||
| 16540 | 1461  | 
fun pretty_context ctxt =  | 
| 10810 | 1462  | 
let  | 
| 
42717
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42502 
diff
changeset
 | 
1463  | 
val verbose = Config.get ctxt verbose;  | 
| 
 
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
 
wenzelm 
parents: 
42502 
diff
changeset
 | 
1464  | 
fun verb f x = if verbose then f (x ()) else [];  | 
| 35141 | 1465  | 
|
| 24922 | 1466  | 
val prt_term = Syntax.pretty_term ctxt;  | 
1467  | 
val prt_typ = Syntax.pretty_typ ctxt;  | 
|
1468  | 
val prt_sort = Syntax.pretty_sort ctxt;  | 
|
| 10810 | 1469  | 
|
1470  | 
(*theory*)  | 
|
| 12057 | 1471  | 
val pretty_thy = Pretty.block  | 
| 17384 | 1472  | 
[Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)];  | 
| 10810 | 1473  | 
|
1474  | 
(*defaults*)  | 
|
1475  | 
fun prt_atom prt prtT (x, X) = Pretty.block  | 
|
1476  | 
[prt x, Pretty.str " ::", Pretty.brk 1, prtT X];  | 
|
1477  | 
||
1478  | 
fun prt_var (x, ~1) = prt_term (Syntax.free x)  | 
|
1479  | 
| prt_var xi = prt_term (Syntax.var xi);  | 
|
1480  | 
||
1481  | 
fun prt_varT (x, ~1) = prt_typ (TFree (x, []))  | 
|
1482  | 
| prt_varT xi = prt_typ (TVar (xi, []));  | 
|
1483  | 
||
1484  | 
val prt_defT = prt_atom prt_var prt_typ;  | 
|
1485  | 
val prt_defS = prt_atom prt_varT prt_sort;  | 
|
| 16540 | 1486  | 
|
| 20163 | 1487  | 
val (types, sorts) = Variable.constraints_of ctxt;  | 
| 10810 | 1488  | 
in  | 
| 18609 | 1489  | 
verb single (K pretty_thy) @  | 
| 
18672
 
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
 
wenzelm 
parents: 
18619 
diff
changeset
 | 
1490  | 
pretty_ctxt ctxt @  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1491  | 
verb (pretty_abbrevs true false) (K ctxt) @  | 
| 
57415
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
1492  | 
verb pretty_term_bindings (K ctxt) @  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59896 
diff
changeset
 | 
1493  | 
verb (pretty_local_facts true) (K ctxt) @  | 
| 10810 | 1494  | 
verb pretty_cases (K ctxt) @  | 
| 18609 | 1495  | 
verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @  | 
| 20163 | 1496  | 
verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))  | 
| 10810 | 1497  | 
end;  | 
1498  | 
||
| 5819 | 1499  | 
end;  | 
| 
35429
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
1500  | 
|
| 42360 | 1501  | 
val show_abbrevs = Proof_Context.show_abbrevs;  |