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