author | wenzelm |
Sat, 25 Jun 2011 19:38:35 +0200 | |
changeset 43552 | 156c822f181a |
parent 43548 | f231a7594e54 |
child 43794 | 49cbbe2768a8 |
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 |
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
13 |
type mode |
24501 | 14 |
val mode_default: mode |
15 |
val mode_stmt: mode |
|
16 |
val mode_pattern: mode |
|
17 |
val mode_schematic: mode |
|
18 |
val mode_abbrev: mode |
|
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 |
21667
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
23 |
val set_stmt: bool -> Proof.context -> Proof.context |
33165 | 24 |
val local_naming: Name_Space.naming |
25 |
val map_naming: (Name_Space.naming -> Name_Space.naming) -> Proof.context -> Proof.context |
|
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33031
diff
changeset
|
26 |
val naming_of: Proof.context -> Name_Space.naming |
33165 | 27 |
val restore_naming: Proof.context -> Proof.context -> Proof.context |
29581 | 28 |
val full_name: Proof.context -> binding -> string |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
42224
diff
changeset
|
29 |
val syntax_of: Proof.context -> Local_Syntax.T |
35111 | 30 |
val syn_of: Proof.context -> Syntax.syntax |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
31 |
val tsig_of: Proof.context -> Type.tsig |
36451 | 32 |
val set_defsort: sort -> Proof.context -> Proof.context |
35680 | 33 |
val default_sort: Proof.context -> indexname -> sort |
20310 | 34 |
val consts_of: Proof.context -> Consts.T |
24752 | 35 |
val the_const_constraint: Proof.context -> string -> typ |
20784 | 36 |
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context |
20310 | 37 |
val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context |
26284 | 38 |
val facts_of: Proof.context -> Facts.T |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
33957
diff
changeset
|
39 |
val cases_of: Proof.context -> (string * (Rule_Cases.T * bool)) list |
42359 | 40 |
val class_space: Proof.context -> Name_Space.T |
41 |
val type_space: Proof.context -> Name_Space.T |
|
42 |
val const_space: Proof.context -> Name_Space.T |
|
43 |
val intern_class: Proof.context -> xstring -> string |
|
44 |
val intern_type: Proof.context -> xstring -> string |
|
45 |
val intern_const: Proof.context -> xstring -> string |
|
46 |
val extern_class: Proof.context -> string -> xstring |
|
47 |
val extern_type: Proof.context -> string -> xstring |
|
48 |
val extern_const: Proof.context -> string -> xstring |
|
27259 | 49 |
val transfer_syntax: theory -> Proof.context -> Proof.context |
20310 | 50 |
val transfer: 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
|
51 |
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
|
52 |
val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context |
28212 | 53 |
val extern_fact: Proof.context -> string -> xstring |
21728 | 54 |
val pretty_term_abbrev: Proof.context -> term -> Pretty.T |
30723 | 55 |
val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T |
20310 | 56 |
val pretty_fact: Proof.context -> string * thm list -> Pretty.T |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
57 |
val read_class: Proof.context -> xstring -> class |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
58 |
val read_arity: Proof.context -> xstring * string list * string -> arity |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
59 |
val cert_arity: Proof.context -> arity -> arity |
20310 | 60 |
val read_typ: Proof.context -> string -> typ |
61 |
val read_typ_syntax: Proof.context -> string -> typ |
|
62 |
val read_typ_abbrev: Proof.context -> string -> typ |
|
63 |
val cert_typ: Proof.context -> typ -> typ |
|
64 |
val cert_typ_syntax: Proof.context -> typ -> typ |
|
65 |
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
|
66 |
val infer_type: Proof.context -> string * typ -> typ |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
67 |
val inferred_param: string -> Proof.context -> typ * Proof.context |
25328 | 68 |
val inferred_fixes: Proof.context -> (string * typ) list * Proof.context |
35360
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
69 |
val read_type_name: Proof.context -> bool -> string -> typ |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
70 |
val read_type_name_proper: Proof.context -> bool -> string -> typ |
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
71 |
val read_const_proper: Proof.context -> bool -> string -> term |
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36503
diff
changeset
|
72 |
val read_const: Proof.context -> bool -> typ -> string -> term |
27259 | 73 |
val allow_dummies: Proof.context -> Proof.context |
42250
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
74 |
val get_sort: Proof.context -> (indexname * sort) list -> indexname -> sort |
36152 | 75 |
val check_tvar: Proof.context -> indexname * sort -> indexname * sort |
76 |
val check_tfree: Proof.context -> string * sort -> string * sort |
|
42250
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
77 |
val intern_skolem: Proof.context -> string -> string option |
24684 | 78 |
val read_term_pattern: Proof.context -> string -> term |
79 |
val read_term_schematic: Proof.context -> string -> term |
|
80 |
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
|
81 |
val show_abbrevs_raw: Config.raw |
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
39557
diff
changeset
|
82 |
val show_abbrevs: bool Config.T |
25345
dd5b851f8ef0
renamed ProofContext.read_const' to ProofContext.read_const_proper;
wenzelm
parents:
25332
diff
changeset
|
83 |
val expand_abbrevs: Proof.context -> term -> term |
20310 | 84 |
val cert_term: Proof.context -> term -> term |
85 |
val cert_prop: Proof.context -> term -> term |
|
35616
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents:
35429
diff
changeset
|
86 |
val def_type: Proof.context -> indexname -> typ option |
20310 | 87 |
val goal_export: Proof.context -> Proof.context -> thm list -> thm list |
88 |
val export: Proof.context -> Proof.context -> thm list -> thm list |
|
21531 | 89 |
val export_morphism: Proof.context -> Proof.context -> morphism |
28396 | 90 |
val norm_export_morphism: Proof.context -> Proof.context -> morphism |
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
91 |
val bind_terms: (indexname * term option) list -> Proof.context -> Proof.context |
20310 | 92 |
val auto_bind_goal: term list -> Proof.context -> Proof.context |
93 |
val auto_bind_facts: term list -> Proof.context -> Proof.context |
|
94 |
val match_bind: bool -> (string list * string) list -> Proof.context -> term list * Proof.context |
|
95 |
val match_bind_i: bool -> (term list * term) list -> Proof.context -> term list * Proof.context |
|
96 |
val read_propp: Proof.context * (string * string list) list list |
|
97 |
-> Proof.context * (term * term list) list list |
|
98 |
val cert_propp: Proof.context * (term * term list) list list |
|
99 |
-> Proof.context * (term * term list) list list |
|
100 |
val read_propp_schematic: Proof.context * (string * string list) list list |
|
101 |
-> Proof.context * (term * term list) list list |
|
102 |
val cert_propp_schematic: Proof.context * (term * term list) list list |
|
103 |
-> Proof.context * (term * term list) list list |
|
104 |
val bind_propp: Proof.context * (string * string list) list list |
|
105 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
106 |
val bind_propp_i: Proof.context * (term * term list) list list |
|
107 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
108 |
val bind_propp_schematic: Proof.context * (string * string list) list list |
|
109 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
110 |
val bind_propp_schematic_i: Proof.context * (term * term list) list list |
|
111 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
18042 | 112 |
val fact_tac: thm list -> int -> tactic |
20310 | 113 |
val some_fact_tac: Proof.context -> int -> tactic |
26346
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
114 |
val get_fact: Proof.context -> Facts.ref -> thm list |
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
115 |
val get_fact_single: Proof.context -> Facts.ref -> thm |
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
116 |
val get_thms: Proof.context -> xstring -> thm list |
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
117 |
val get_thm: Proof.context -> xstring -> thm |
30761
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
118 |
val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list -> |
30211 | 119 |
Proof.context -> (string * thm list) list * Proof.context |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
120 |
val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context |
29581 | 121 |
val read_vars: (binding * string option * mixfix) list -> Proof.context -> |
122 |
(binding * typ option * mixfix) list * Proof.context |
|
123 |
val cert_vars: (binding * typ option * mixfix) list -> Proof.context -> |
|
124 |
(binding * typ option * mixfix) list * Proof.context |
|
30763
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents:
30761
diff
changeset
|
125 |
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
|
126 |
string list * Proof.context |
20310 | 127 |
val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a) |
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
128 |
val add_assms: Assumption.export -> |
30211 | 129 |
(Thm.binding * (string * string list) list) list -> |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
130 |
Proof.context -> (string * thm list) list * Proof.context |
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
131 |
val add_assms_i: Assumption.export -> |
30211 | 132 |
(Thm.binding * (term * term list) list) list -> |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
133 |
Proof.context -> (string * thm list) list * Proof.context |
33368 | 134 |
val add_cases: bool -> (string * Rule_Cases.T option) list -> Proof.context -> Proof.context |
135 |
val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context |
|
42496 | 136 |
val get_case: Proof.context -> string -> binding option list -> Rule_Cases.T |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
137 |
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context |
24949 | 138 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
139 |
val target_type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> morphism -> |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
140 |
Context.generic -> Context.generic |
24949 | 141 |
val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism -> |
21744 | 142 |
Context.generic -> Context.generic |
35680 | 143 |
val class_alias: binding -> class -> Proof.context -> Proof.context |
144 |
val type_alias: binding -> string -> Proof.context -> Proof.context |
|
145 |
val const_alias: binding -> string -> Proof.context -> Proof.context |
|
24767 | 146 |
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
|
147 |
val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context |
25052 | 148 |
val revert_abbrev: string -> string -> Proof.context -> Proof.context |
20310 | 149 |
val print_syntax: Proof.context -> unit |
21728 | 150 |
val print_abbrevs: Proof.context -> unit |
20310 | 151 |
val print_binds: Proof.context -> unit |
152 |
val print_lthms: Proof.context -> unit |
|
153 |
val print_cases: Proof.context -> unit |
|
42717
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
154 |
val debug: bool Config.T |
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
155 |
val verbose: bool Config.T |
20310 | 156 |
val pretty_ctxt: Proof.context -> Pretty.T list |
157 |
val pretty_context: Proof.context -> Pretty.T list |
|
5819 | 158 |
end; |
159 |
||
42360 | 160 |
structure Proof_Context: PROOF_CONTEXT = |
5819 | 161 |
struct |
162 |
||
42363 | 163 |
val theory_of = Proof_Context.theory_of; |
164 |
val init_global = Proof_Context.init_global; |
|
165 |
||
12057 | 166 |
|
7270 | 167 |
|
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
168 |
(** inner syntax mode **) |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
169 |
|
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
170 |
datatype mode = |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
171 |
Mode of |
24486 | 172 |
{stmt: bool, (*inner statement mode*) |
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
173 |
pattern: bool, (*pattern binding schematic variables*) |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
174 |
schematic: bool, (*term referencing loose schematic variables*) |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
175 |
abbrev: bool}; (*abbrev mode -- no normalization*) |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
176 |
|
24486 | 177 |
fun make_mode (stmt, pattern, schematic, abbrev) = |
178 |
Mode {stmt = stmt, pattern = pattern, schematic = schematic, abbrev = abbrev}; |
|
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
179 |
|
24486 | 180 |
val mode_default = make_mode (false, false, false, false); |
181 |
val mode_stmt = make_mode (true, false, false, false); |
|
182 |
val mode_pattern = make_mode (false, true, false, false); |
|
183 |
val mode_schematic = make_mode (false, false, true, false); |
|
184 |
val mode_abbrev = make_mode (false, false, false, true); |
|
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
185 |
|
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
186 |
|
5819 | 187 |
|
16540 | 188 |
(** Isar proof context information **) |
5819 | 189 |
|
16540 | 190 |
datatype ctxt = |
191 |
Ctxt of |
|
36451 | 192 |
{mode: mode, (*inner syntax mode*) |
193 |
naming: Name_Space.naming, (*local naming conventions*) |
|
194 |
syntax: Local_Syntax.T, (*local syntax*) |
|
195 |
tsig: Type.tsig * Type.tsig, (*local/global type signature -- local name space / defsort only*) |
|
196 |
consts: Consts.T * Consts.T, (*local/global consts -- local name space / abbrevs only*) |
|
197 |
facts: Facts.T, (*local facts*) |
|
33368 | 198 |
cases: (string * (Rule_Cases.T * bool)) list}; (*named case contexts*) |
5819 | 199 |
|
36450 | 200 |
fun make_ctxt (mode, naming, syntax, tsig, consts, facts, cases) = |
26240 | 201 |
Ctxt {mode = mode, naming = naming, syntax = syntax, |
36450 | 202 |
tsig = tsig, consts = consts, facts = facts, cases = cases}; |
5819 | 203 |
|
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33031
diff
changeset
|
204 |
val local_naming = Name_Space.default_naming |> Name_Space.add_path "local"; |
19079
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
wenzelm
parents:
19062
diff
changeset
|
205 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
206 |
structure Data = Proof_Data |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
207 |
( |
16540 | 208 |
type T = ctxt; |
209 |
fun init thy = |
|
33387 | 210 |
make_ctxt (mode_default, local_naming, Local_Syntax.init thy, |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
211 |
(Sign.tsig_of thy, Sign.tsig_of thy), |
26284 | 212 |
(Sign.consts_of thy, Sign.consts_of thy), Facts.empty, []); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
213 |
); |
5819 | 214 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
215 |
fun rep_context ctxt = Data.get ctxt |> (fn Ctxt args => args); |
5819 | 216 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
217 |
fun map_context f = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
37145
diff
changeset
|
218 |
Data.map (fn Ctxt {mode, naming, syntax, tsig, consts, facts, cases} => |
36450 | 219 |
make_ctxt (f (mode, naming, syntax, tsig, consts, facts, cases))); |
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
220 |
|
36450 | 221 |
fun set_mode mode = map_context (fn (_, naming, syntax, tsig, consts, facts, cases) => |
222 |
(mode, naming, syntax, tsig, consts, facts, cases)); |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
223 |
|
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
224 |
fun map_mode f = |
36450 | 225 |
map_context (fn (Mode {stmt, pattern, schematic, abbrev}, naming, syntax, tsig, consts, facts, cases) => |
226 |
(make_mode (f (stmt, pattern, schematic, abbrev)), naming, syntax, tsig, consts, facts, cases)); |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
227 |
|
19001 | 228 |
fun map_naming f = |
36450 | 229 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
230 |
(mode, f naming, syntax, tsig, consts, facts, cases)); |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
231 |
|
19001 | 232 |
fun map_syntax f = |
36450 | 233 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
234 |
(mode, naming, 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
|
235 |
|
36450 | 236 |
fun map_tsig f = |
237 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
|
238 |
(mode, naming, 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
|
239 |
|
19001 | 240 |
fun map_consts f = |
36450 | 241 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
242 |
(mode, naming, 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
|
243 |
|
26284 | 244 |
fun map_facts f = |
36450 | 245 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
246 |
(mode, naming, syntax, tsig, consts, f facts, cases)); |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
247 |
|
19001 | 248 |
fun map_cases f = |
36450 | 249 |
map_context (fn (mode, naming, syntax, tsig, consts, facts, cases) => |
250 |
(mode, naming, syntax, tsig, consts, facts, f cases)); |
|
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
251 |
|
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
252 |
val get_mode = #mode o rep_context; |
28407 | 253 |
val restore_mode = set_mode o get_mode; |
27286 | 254 |
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
|
255 |
|
24486 | 256 |
fun set_stmt stmt = |
257 |
map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev)); |
|
19001 | 258 |
|
259 |
val naming_of = #naming o rep_context; |
|
33165 | 260 |
val restore_naming = map_naming o K o naming_of |
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33031
diff
changeset
|
261 |
val full_name = Name_Space.full_name o naming_of; |
5819 | 262 |
|
16540 | 263 |
val syntax_of = #syntax o rep_context; |
33387 | 264 |
val syn_of = Local_Syntax.syn_of o syntax_of; |
265 |
val set_syntax_mode = map_syntax o Local_Syntax.set_mode; |
|
266 |
val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of; |
|
19001 | 267 |
|
36450 | 268 |
val tsig_of = #1 o #tsig o rep_context; |
36451 | 269 |
val set_defsort = map_tsig o apfst o Type.set_defsort; |
35680 | 270 |
fun default_sort ctxt = the_default (Type.defaultS (tsig_of ctxt)) o Variable.def_sort ctxt; |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
271 |
|
25052 | 272 |
val consts_of = #1 o #consts o rep_context; |
24752 | 273 |
val the_const_constraint = Consts.the_constraint o consts_of; |
5819 | 274 |
|
26284 | 275 |
val facts_of = #facts o rep_context; |
16540 | 276 |
val cases_of = #cases o rep_context; |
5819 | 277 |
|
278 |
||
42359 | 279 |
(* name spaces *) |
280 |
||
281 |
val class_space = Type.class_space o tsig_of; |
|
282 |
val type_space = Type.type_space o tsig_of; |
|
283 |
val const_space = Consts.space_of o consts_of; |
|
284 |
||
285 |
val intern_class = Name_Space.intern o class_space; |
|
286 |
val intern_type = Name_Space.intern o type_space; |
|
287 |
val intern_const = Name_Space.intern o const_space; |
|
288 |
||
289 |
fun extern_class ctxt = Name_Space.extern ctxt (class_space ctxt); |
|
290 |
fun extern_type ctxt = Name_Space.extern ctxt (type_space ctxt); |
|
291 |
fun extern_const ctxt = Name_Space.extern ctxt (const_space ctxt); |
|
292 |
||
293 |
||
20367 | 294 |
(* theory transfer *) |
12093 | 295 |
|
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
296 |
fun transfer_syntax thy ctxt = ctxt |> |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
297 |
map_syntax (Local_Syntax.rebuild thy) |> |
36450 | 298 |
map_tsig (fn tsig as (local_tsig, global_tsig) => |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
299 |
let val thy_tsig = Sign.tsig_of thy in |
36450 | 300 |
if Type.eq_tsig (thy_tsig, global_tsig) then tsig |
42387 | 301 |
else (Type.merge_tsig ctxt (local_tsig, thy_tsig), thy_tsig) |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
302 |
end) |> |
30424
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
wenzelm
parents:
30420
diff
changeset
|
303 |
map_consts (fn consts as (local_consts, global_consts) => |
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
wenzelm
parents:
30420
diff
changeset
|
304 |
let val thy_consts = Sign.consts_of thy in |
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
wenzelm
parents:
30420
diff
changeset
|
305 |
if Consts.eq_consts (thy_consts, global_consts) then consts |
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
wenzelm
parents:
30420
diff
changeset
|
306 |
else (Consts.merge (local_consts, thy_consts), thy_consts) |
692279df7cc2
Consts.eq_const is back again (cf. 907da436f8a9) -- required in ProofContext.transfer_syntax to prevent expensive merges of local_consts/global_consts;
wenzelm
parents:
30420
diff
changeset
|
307 |
end); |
17072 | 308 |
|
33031
b75c35574e04
backpatching of structure Proof and ProofContext -- avoid odd aliases;
wenzelm
parents:
32966
diff
changeset
|
309 |
fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy; |
17072 | 310 |
|
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
|
311 |
fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt; |
20367 | 312 |
|
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
|
313 |
fun background_theory_result f ctxt = |
20367 | 314 |
let val (res, thy') = f (theory_of ctxt) |
315 |
in (res, ctxt |> transfer thy') end; |
|
19019 | 316 |
|
12093 | 317 |
|
318 |
||
14828 | 319 |
(** pretty printing **) |
320 |
||
28212 | 321 |
(* extern *) |
322 |
||
42378 | 323 |
fun which_facts ctxt name = |
28212 | 324 |
let |
325 |
val local_facts = facts_of ctxt; |
|
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39508
diff
changeset
|
326 |
val global_facts = Global_Theory.facts_of (theory_of ctxt); |
28212 | 327 |
in |
328 |
if is_some (Facts.lookup (Context.Proof ctxt) local_facts name) |
|
42378 | 329 |
then local_facts else global_facts |
28212 | 330 |
end; |
331 |
||
42379 | 332 |
fun markup_fact ctxt name = Name_Space.markup (Facts.space_of (which_facts ctxt name)) name; |
42378 | 333 |
|
334 |
fun extern_fact ctxt name = Facts.extern ctxt (which_facts ctxt name) name; |
|
335 |
||
28212 | 336 |
|
337 |
(* pretty *) |
|
338 |
||
24922 | 339 |
fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt); |
14828 | 340 |
|
42378 | 341 |
fun pretty_fact_name ctxt a = |
342 |
Pretty.block [Pretty.mark_str (markup_fact ctxt a, extern_fact ctxt a), Pretty.str ":"]; |
|
28209
02f3222a392d
pretty_fact: extern fact name wrt. the given context, assuming that is the proper one for presentation;
wenzelm
parents:
28087
diff
changeset
|
343 |
|
32090
39acf19e9f3a
moved ProofContext.pretty_thm to Display.pretty_thm etc.;
wenzelm
parents:
32032
diff
changeset
|
344 |
fun pretty_fact_aux ctxt flag ("", ths) = |
39acf19e9f3a
moved ProofContext.pretty_thm to Display.pretty_thm etc.;
wenzelm
parents:
32032
diff
changeset
|
345 |
Display.pretty_thms_aux ctxt flag ths |
30723 | 346 |
| pretty_fact_aux ctxt flag (a, [th]) = Pretty.block |
32090
39acf19e9f3a
moved ProofContext.pretty_thm to Display.pretty_thm etc.;
wenzelm
parents:
32032
diff
changeset
|
347 |
[pretty_fact_name ctxt a, Pretty.brk 1, Display.pretty_thm_aux ctxt flag th] |
30723 | 348 |
| pretty_fact_aux ctxt flag (a, ths) = Pretty.block |
32090
39acf19e9f3a
moved ProofContext.pretty_thm to Display.pretty_thm etc.;
wenzelm
parents:
32032
diff
changeset
|
349 |
(Pretty.fbreaks (pretty_fact_name ctxt a :: map (Display.pretty_thm_aux ctxt flag) ths)); |
30723 | 350 |
|
351 |
fun pretty_fact ctxt = pretty_fact_aux ctxt true; |
|
14828 | 352 |
|
353 |
||
354 |
||
5819 | 355 |
(** prepare types **) |
356 |
||
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
357 |
(* classes *) |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
358 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
359 |
fun read_class ctxt text = |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
360 |
let |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
361 |
val tsig = tsig_of ctxt; |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
362 |
val (syms, pos) = Syntax.read_token text; |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
363 |
val c = Type.cert_class tsig (Type.intern_class tsig (Symbol_Pos.content syms)) |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
364 |
handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos); |
42467 | 365 |
val _ = Context_Position.report ctxt pos (Name_Space.markup (Type.class_space tsig) c); |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
366 |
in c end; |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
367 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
368 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
369 |
(* type arities *) |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
370 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
371 |
local |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
372 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
373 |
fun prep_arity prep_tycon prep_sort ctxt (t, Ss, S) = |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
374 |
let val arity = (prep_tycon ctxt t, map (prep_sort ctxt) Ss, prep_sort ctxt S) |
42387 | 375 |
in Type.add_arity ctxt arity (tsig_of ctxt); arity end; |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
376 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
377 |
in |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
378 |
|
42359 | 379 |
val read_arity = prep_arity intern_type Syntax.read_sort; |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
380 |
val cert_arity = prep_arity (K I) (Type.cert_sort o tsig_of); |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
381 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
382 |
end; |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
383 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
384 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
385 |
(* types *) |
24277 | 386 |
|
387 |
fun read_typ_mode mode ctxt s = |
|
24486 | 388 |
Syntax.read_typ (Type.set_mode mode ctxt) s; |
24277 | 389 |
|
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
390 |
val read_typ = read_typ_mode Type.mode_default; |
24277 | 391 |
val read_typ_syntax = read_typ_mode Type.mode_syntax; |
392 |
val read_typ_abbrev = read_typ_mode Type.mode_abbrev; |
|
393 |
||
394 |
||
395 |
fun cert_typ_mode mode ctxt T = |
|
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
396 |
Type.cert_typ_mode mode (tsig_of ctxt) T |
24277 | 397 |
handle TYPE (msg, _, _) => error msg; |
398 |
||
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
399 |
val cert_typ = cert_typ_mode Type.mode_default; |
24277 | 400 |
val cert_typ_syntax = cert_typ_mode Type.mode_syntax; |
401 |
val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev; |
|
402 |
||
403 |
||
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
404 |
|
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
405 |
(** prepare variables **) |
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
406 |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
407 |
(* check Skolem constants *) |
7679 | 408 |
|
18678 | 409 |
fun no_skolem internal x = |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
410 |
if can Name.dest_skolem x then |
18678 | 411 |
error ("Illegal reference to internal Skolem constant: " ^ quote x) |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
412 |
else if not internal andalso can Name.dest_internal x then |
18678 | 413 |
error ("Illegal reference to internal variable: " ^ quote x) |
7679 | 414 |
else x; |
415 |
||
416 |
||
18187 | 417 |
|
5819 | 418 |
(** prepare terms and propositions **) |
419 |
||
25328 | 420 |
(* inferred types of parameters *) |
421 |
||
422 |
fun infer_type ctxt x = |
|
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36503
diff
changeset
|
423 |
Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt)) (Free x)); |
25328 | 424 |
|
425 |
fun inferred_param x ctxt = |
|
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36503
diff
changeset
|
426 |
let val T = infer_type ctxt (x, dummyT) |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
427 |
in (T, ctxt |> Variable.declare_term (Free (x, T))) end; |
25328 | 428 |
|
429 |
fun inferred_fixes ctxt = |
|
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
430 |
let |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
431 |
val xs = map #2 (Variable.dest_fixes ctxt); |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
432 |
val (Ts, ctxt') = fold_map inferred_param xs ctxt; |
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
433 |
in (xs ~~ Ts, ctxt') end; |
25328 | 434 |
|
435 |
||
436 |
(* type and constant names *) |
|
437 |
||
27821 | 438 |
local |
439 |
||
30573 | 440 |
val token_content = Syntax.read_token #>> Symbol_Pos.content; |
27821 | 441 |
|
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
442 |
fun prep_const_proper ctxt strict (c, pos) = |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
443 |
let |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
444 |
fun err msg = error (msg ^ Position.str_of pos); |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
445 |
val consts = consts_of ctxt; |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
446 |
val t as Const (d, _) = |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
447 |
(case Variable.lookup_const ctxt c of |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
448 |
SOME d => |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
449 |
Const (d, Consts.type_scheme (consts_of ctxt) d handle TYPE (msg, _, _) => err msg) |
42469 | 450 |
| NONE => Consts.read_const consts (c, pos)); |
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
451 |
val _ = |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
452 |
if strict then ignore (Consts.the_type consts d) handle TYPE (msg, _, _) => err msg |
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
453 |
else (); |
42467 | 454 |
val _ = Context_Position.report ctxt pos (Name_Space.markup (Consts.space_of consts) d); |
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
455 |
in t end; |
27821 | 456 |
|
457 |
in |
|
25328 | 458 |
|
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
459 |
fun read_type_name ctxt strict text = |
27821 | 460 |
let |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
461 |
val tsig = tsig_of ctxt; |
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
462 |
val (c, pos) = token_content text; |
27821 | 463 |
in |
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
42287
diff
changeset
|
464 |
if Lexicon.is_tid c then |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39441
diff
changeset
|
465 |
(Context_Position.report ctxt pos Markup.tfree; |
35680 | 466 |
TFree (c, default_sort ctxt (c, ~1))) |
27821 | 467 |
else |
468 |
let |
|
42359 | 469 |
val d = intern_type ctxt c; |
42468 | 470 |
val decl = Type.the_decl tsig (d, pos); |
42467 | 471 |
fun err () = error ("Bad type name: " ^ quote d ^ Position.str_of pos); |
35360
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
472 |
val args = |
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
473 |
(case decl of |
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
474 |
Type.LogicalType n => n |
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
475 |
| Type.Abbreviation (vs, _, _) => if strict then err () else length vs |
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
476 |
| Type.Nonterminal => if strict then err () else 0); |
42467 | 477 |
val _ = Context_Position.report ctxt pos (Name_Space.markup (Type.type_space tsig) d); |
35360
df2b2168e43a
clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents:
35262
diff
changeset
|
478 |
in Type (d, replicate args dummyT) end |
27821 | 479 |
end; |
25328 | 480 |
|
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
481 |
fun read_type_name_proper ctxt strict text = |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
482 |
(case read_type_name ctxt strict text of |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
483 |
T as Type _ => T |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
484 |
| T => error ("Not a type constructor: " ^ Syntax.string_of_typ ctxt T)); |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
485 |
|
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
486 |
|
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
487 |
fun read_const_proper ctxt strict = prep_const_proper ctxt strict o token_content; |
27821 | 488 |
|
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36503
diff
changeset
|
489 |
fun read_const ctxt strict ty text = |
36542
7cb6b40d19b2
read_const: disallow internal names as usual in visible Isar text;
wenzelm
parents:
36505
diff
changeset
|
490 |
let |
7cb6b40d19b2
read_const: disallow internal names as usual in visible Isar text;
wenzelm
parents:
36505
diff
changeset
|
491 |
val (c, pos) = token_content text; |
7cb6b40d19b2
read_const: disallow internal names as usual in visible Isar text;
wenzelm
parents:
36505
diff
changeset
|
492 |
val _ = no_skolem false c; |
7cb6b40d19b2
read_const: disallow internal names as usual in visible Isar text;
wenzelm
parents:
36505
diff
changeset
|
493 |
in |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
494 |
(case (Variable.lookup_fixed ctxt c, Variable.is_const ctxt c) of |
27821 | 495 |
(SOME x, false) => |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39441
diff
changeset
|
496 |
(Context_Position.report ctxt pos |
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39441
diff
changeset
|
497 |
(Markup.name x (if can Name.dest_skolem x then Markup.skolem else Markup.free)); |
36505
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
wenzelm
parents:
36503
diff
changeset
|
498 |
Free (x, infer_type ctxt (x, ty))) |
35399
3881972fcfca
clarified ProofContext.read_const(_proper)/Args.const(_proper) wrt. strict logical consts;
wenzelm
parents:
35360
diff
changeset
|
499 |
| _ => prep_const_proper ctxt strict (c, pos)) |
27821 | 500 |
end; |
501 |
||
502 |
end; |
|
25328 | 503 |
|
504 |
||
42250
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
505 |
(* skolem variables *) |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
506 |
|
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
507 |
fun intern_skolem ctxt x = |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
508 |
let |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
509 |
val _ = no_skolem false x; |
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
510 |
val sko = Variable.lookup_fixed ctxt x; |
42250
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
511 |
val is_const = can (read_const_proper ctxt false) x orelse Long_Name.is_qualified x; |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
512 |
val is_declared = is_some (Variable.def_type ctxt false (x, ~1)); |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
513 |
in |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
514 |
if Variable.is_const ctxt x then NONE |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
515 |
else if is_some sko then sko |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
516 |
else if not is_const orelse is_declared then SOME x |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
517 |
else NONE |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
518 |
end; |
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
519 |
|
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
520 |
|
24684 | 521 |
(* read_term *) |
522 |
||
523 |
fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt); |
|
524 |
||
525 |
val read_term_pattern = read_term_mode mode_pattern; |
|
526 |
val read_term_schematic = read_term_mode mode_schematic; |
|
527 |
val read_term_abbrev = read_term_mode mode_abbrev; |
|
528 |
||
529 |
||
19001 | 530 |
(* local abbreviations *) |
5819 | 531 |
|
24501 | 532 |
local |
533 |
||
42383
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
wenzelm
parents:
42381
diff
changeset
|
534 |
fun certify_consts ctxt = Consts.certify (Context.pretty ctxt) (tsig_of ctxt) |
27286 | 535 |
(not (abbrev_mode ctxt)) (consts_of ctxt); |
19001 | 536 |
|
38979
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
537 |
fun expand_binds ctxt = |
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
538 |
let |
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
539 |
val Mode {pattern, schematic, ...} = get_mode ctxt; |
5819 | 540 |
|
38979
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
541 |
fun reject_schematic (t as Var _) = |
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
542 |
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
|
543 |
| 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
|
544 |
| 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
|
545 |
| reject_schematic _ = (); |
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
546 |
in |
24495 | 547 |
if pattern then I |
548 |
else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic) |
|
549 |
end; |
|
5819 | 550 |
|
24501 | 551 |
in |
552 |
||
553 |
fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt; |
|
554 |
||
555 |
end; |
|
556 |
||
40879
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
39557
diff
changeset
|
557 |
val show_abbrevs_raw = Config.declare "show_abbrevs" (fn _ => Config.Bool true); |
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
39557
diff
changeset
|
558 |
val show_abbrevs = Config.bool show_abbrevs_raw; |
5819 | 559 |
|
24922 | 560 |
fun contract_abbrevs ctxt t = |
561 |
let |
|
562 |
val thy = theory_of ctxt; |
|
563 |
val consts = consts_of ctxt; |
|
564 |
val Mode {abbrev, ...} = get_mode ctxt; |
|
30566
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
wenzelm
parents:
30473
diff
changeset
|
565 |
val retrieve = Consts.retrieve_abbrevs consts (print_mode_value () @ [""]); |
9643f54c4184
reverted abbreviations: improved performance via Item_Net.T;
wenzelm
parents:
30473
diff
changeset
|
566 |
fun match_abbrev u = Option.map #1 (get_first (Pattern.match_rew thy u) (retrieve u)); |
24922 | 567 |
in |
40879
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
39557
diff
changeset
|
568 |
if abbrev orelse not (Config.get ctxt show_abbrevs) orelse not (can Term.type_of t) then t |
35211
5d2fe4e09354
Use top-down rewriting to contract abbreviations.
berghofe
parents:
35141
diff
changeset
|
569 |
else Pattern.rewrite_term_top thy [] [match_abbrev] t |
24922 | 570 |
end; |
571 |
||
572 |
||
24518 | 573 |
(* patterns *) |
574 |
||
32003 | 575 |
fun prepare_patternT ctxt T = |
576 |
let |
|
577 |
val Mode {pattern, schematic, ...} = get_mode ctxt; |
|
578 |
val _ = |
|
579 |
pattern orelse schematic orelse |
|
580 |
T |> Term.exists_subtype |
|
38979
60dbf0b3f6c7
prefer regular print functions over slightly low-level Term.string_of_vname;
wenzelm
parents:
38756
diff
changeset
|
581 |
(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
|
582 |
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
|
583 |
error ("Illegal schematic type variable: " ^ Syntax.string_of_typ ctxt T) |
32003 | 584 |
| _ => false) |
585 |
in T end; |
|
24518 | 586 |
|
22712 | 587 |
|
24505 | 588 |
local |
6550 | 589 |
|
42360 | 590 |
val dummies = Config.bool (Config.declare "Proof_Context.dummies" (K (Config.Bool false))); |
27259 | 591 |
|
592 |
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
|
593 |
if Config.get ctxt dummies then t |
27259 | 594 |
else Term.no_dummy_patterns t handle TERM _ => error "Illegal dummy pattern(s) in term"; |
595 |
||
24767 | 596 |
fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1); |
6762 | 597 |
|
27259 | 598 |
in |
6550 | 599 |
|
39508
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
wenzelm
parents:
39507
diff
changeset
|
600 |
val allow_dummies = Config.put dummies true; |
24505 | 601 |
|
24684 | 602 |
fun prepare_patterns ctxt = |
24518 | 603 |
let val Mode {pattern, ...} = get_mode ctxt in |
39296 | 604 |
Type_Infer.fixate ctxt #> |
24767 | 605 |
pattern ? Variable.polymorphic ctxt #> |
24684 | 606 |
(map o Term.map_types) (prepare_patternT ctxt) #> |
27259 | 607 |
(if pattern then prepare_dummies else map (check_dummies ctxt)) |
24505 | 608 |
end; |
609 |
||
610 |
end; |
|
611 |
||
6550 | 612 |
|
42250
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
wenzelm
parents:
42242
diff
changeset
|
613 |
(* sort constraints *) |
27286 | 614 |
|
36448
edb757388592
get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents:
36429
diff
changeset
|
615 |
fun get_sort ctxt raw_text = |
27286 | 616 |
let |
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
617 |
val tsig = tsig_of ctxt; |
27286 | 618 |
|
36448
edb757388592
get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents:
36429
diff
changeset
|
619 |
val text = distinct (op =) (map (apsnd (Type.minimize_sort tsig)) raw_text); |
36152 | 620 |
val _ = |
36448
edb757388592
get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents:
36429
diff
changeset
|
621 |
(case duplicates (eq_fst (op =)) text of |
36152 | 622 |
[] => () |
27286 | 623 |
| dups => error ("Inconsistent sort constraints for type variable(s) " |
624 |
^ commas_quote (map (Term.string_of_vname' o fst) dups))); |
|
625 |
||
36152 | 626 |
fun lookup xi = |
36448
edb757388592
get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents:
36429
diff
changeset
|
627 |
(case AList.lookup (op =) text xi of |
36152 | 628 |
NONE => NONE |
629 |
| SOME S => if S = dummyS then NONE else SOME S); |
|
630 |
||
27286 | 631 |
fun get xi = |
36152 | 632 |
(case (lookup xi, Variable.def_sort ctxt xi) of |
27286 | 633 |
(NONE, NONE) => Type.defaultS tsig |
634 |
| (NONE, SOME S) => S |
|
635 |
| (SOME S, NONE) => S |
|
636 |
| (SOME S, SOME S') => |
|
637 |
if Type.eq_sort tsig (S, S') then S' |
|
35669
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
638 |
else error ("Sort constraint " ^ Syntax.string_of_sort ctxt S ^ |
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
wenzelm
parents:
35616
diff
changeset
|
639 |
" inconsistent with default " ^ Syntax.string_of_sort ctxt S' ^ |
30815 | 640 |
" for type variable " ^ quote (Term.string_of_vname' xi))); |
36448
edb757388592
get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents:
36429
diff
changeset
|
641 |
in get end; |
27286 | 642 |
|
36152 | 643 |
fun check_tvar ctxt (xi, S) = (xi, get_sort ctxt [(xi, S)] xi); |
644 |
fun check_tfree ctxt (x, S) = apfst fst (check_tvar ctxt ((x, ~1), S)); |
|
645 |
||
5819 | 646 |
|
647 |
(* certify terms *) |
|
648 |
||
10554 | 649 |
local |
650 |
||
24684 | 651 |
fun gen_cert prop ctxt t = |
652 |
t |
|
653 |
|> expand_abbrevs ctxt |
|
42383
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
wenzelm
parents:
42381
diff
changeset
|
654 |
|> (fn t' => |
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
wenzelm
parents:
42381
diff
changeset
|
655 |
#1 (Sign.certify' prop (Context.pretty ctxt) false (consts_of ctxt) (theory_of ctxt) t') |
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
wenzelm
parents:
42381
diff
changeset
|
656 |
handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg); |
16501 | 657 |
|
10554 | 658 |
in |
8096 | 659 |
|
24684 | 660 |
val cert_term = gen_cert false; |
661 |
val cert_prop = gen_cert true; |
|
10554 | 662 |
|
663 |
end; |
|
5819 | 664 |
|
665 |
||
42405
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
wenzelm
parents:
42402
diff
changeset
|
666 |
(* check/uncheck *) |
22701 | 667 |
|
35616
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents:
35429
diff
changeset
|
668 |
fun def_type ctxt = |
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents:
35429
diff
changeset
|
669 |
let val Mode {pattern, ...} = get_mode ctxt |
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents:
35429
diff
changeset
|
670 |
in Variable.def_type ctxt pattern end; |
b342390d296f
provide ProofContext.def_type depending on "pattern" mode;
wenzelm
parents:
35429
diff
changeset
|
671 |
|
25406 | 672 |
local |
673 |
||
24518 | 674 |
fun standard_typ_check ctxt = |
675 |
map (cert_typ_mode (Type.get_mode ctxt) ctxt) #> |
|
676 |
map (prepare_patternT ctxt); |
|
677 |
||
24922 | 678 |
fun standard_term_uncheck ctxt = |
679 |
map (contract_abbrevs ctxt); |
|
680 |
||
24518 | 681 |
in |
682 |
||
42402
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42387
diff
changeset
|
683 |
val _ = Context.>> |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42387
diff
changeset
|
684 |
(Syntax.add_typ_check 0 "standard" standard_typ_check #> |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42387
diff
changeset
|
685 |
Syntax.add_term_check 100 "fixate" prepare_patterns #> |
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
wenzelm
parents:
42387
diff
changeset
|
686 |
Syntax.add_term_uncheck 0 "standard" standard_term_uncheck); |
22701 | 687 |
|
24518 | 688 |
end; |
22701 | 689 |
|
690 |
||
9553 | 691 |
|
21610 | 692 |
(** export results **) |
21531 | 693 |
|
20310 | 694 |
fun common_export is_goal inner outer = |
695 |
map (Assumption.export is_goal inner outer) #> |
|
696 |
Variable.export inner outer; |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
697 |
|
20310 | 698 |
val goal_export = common_export true; |
699 |
val export = common_export false; |
|
12704 | 700 |
|
21531 | 701 |
fun export_morphism inner outer = |
702 |
Assumption.export_morphism inner outer $> |
|
703 |
Variable.export_morphism inner outer; |
|
704 |
||
28396 | 705 |
fun norm_export_morphism inner outer = |
706 |
export_morphism inner outer $> |
|
707 |
Morphism.thm_morphism Goal.norm_result; |
|
708 |
||
21531 | 709 |
|
15758
07e382399a96
binds/thms: do not store options, but delete from table;
wenzelm
parents:
15750
diff
changeset
|
710 |
|
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
711 |
(** term bindings **) |
5819 | 712 |
|
8096 | 713 |
(* simult_matches *) |
714 |
||
19867 | 715 |
fun simult_matches ctxt (t, pats) = |
716 |
(case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of |
|
717 |
NONE => error "Pattern match failed!" |
|
32032 | 718 |
| SOME (env, _) => Vartab.fold (fn (v, (_, t)) => cons (v, t)) (Envir.term_env env) []); |
8096 | 719 |
|
720 |
||
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
721 |
(* bind_terms *) |
7925 | 722 |
|
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
723 |
val bind_terms = fold (fn (xi, t) => fn ctxt => |
24511
69d270cc7e4f
removed obsolete read/cert variations (cf. Syntax.read/check);
wenzelm
parents:
24505
diff
changeset
|
724 |
ctxt |
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
725 |
|> Variable.bind_term (xi, Option.map (cert_term (set_mode mode_default ctxt)) t)); |
5819 | 726 |
|
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
727 |
|
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
728 |
(* auto_bind *) |
10810 | 729 |
|
20330 | 730 |
fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b |
10554 | 731 |
| drop_schematic b = b; |
732 |
||
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
733 |
fun auto_bind f ts ctxt = ctxt |> bind_terms (map drop_schematic (f (theory_of ctxt) ts)); |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
734 |
|
33386 | 735 |
val auto_bind_goal = auto_bind Auto_Bind.goal; |
736 |
val auto_bind_facts = auto_bind Auto_Bind.facts; |
|
7925 | 737 |
|
5819 | 738 |
|
8096 | 739 |
(* match_bind(_i) *) |
5819 | 740 |
|
8096 | 741 |
local |
742 |
||
24684 | 743 |
fun gen_bind prep_terms gen raw_binds ctxt = |
5819 | 744 |
let |
24684 | 745 |
fun prep_bind (raw_pats, t) ctxt1 = |
746 |
let |
|
747 |
val T = Term.fastype_of t; |
|
748 |
val ctxt2 = Variable.declare_term t ctxt1; |
|
749 |
val pats = prep_terms (set_mode mode_pattern ctxt2) T raw_pats; |
|
750 |
val binds = simult_matches ctxt2 (t, pats); |
|
751 |
in (binds, ctxt2) end; |
|
7670 | 752 |
|
24686 | 753 |
val ts = prep_terms ctxt dummyT (map snd raw_binds); |
754 |
val (binds, ctxt') = apfst flat (fold_map prep_bind (map fst raw_binds ~~ ts) ctxt); |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
755 |
val binds' = |
19916
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
wenzelm
parents:
19897
diff
changeset
|
756 |
if gen then map #1 binds ~~ Variable.exportT_terms ctxt' ctxt (map #2 binds) |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
757 |
else binds; |
15531 | 758 |
val binds'' = map (apsnd SOME) binds'; |
18310 | 759 |
val ctxt'' = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
760 |
tap (Variable.warn_extra_tfrees ctxt) |
18310 | 761 |
(if gen then |
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
762 |
ctxt (*sic!*) |> fold Variable.declare_term (map #2 binds') |> bind_terms binds'' |
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
763 |
else ctxt' |> bind_terms binds''); |
18310 | 764 |
in (ts, ctxt'') end; |
8096 | 765 |
|
766 |
in |
|
5935 | 767 |
|
24684 | 768 |
fun read_terms ctxt T = |
39288 | 769 |
map (Syntax.parse_term ctxt #> Type.constraint T) #> Syntax.check_terms ctxt; |
24684 | 770 |
|
771 |
val match_bind = gen_bind read_terms; |
|
772 |
val match_bind_i = gen_bind (fn ctxt => fn _ => map (cert_term ctxt)); |
|
8096 | 773 |
|
774 |
end; |
|
5935 | 775 |
|
776 |
||
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
777 |
(* propositions with patterns *) |
5935 | 778 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
779 |
local |
8096 | 780 |
|
24684 | 781 |
fun prep_propp mode prep_props (context, args) = |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
782 |
let |
19585 | 783 |
fun prep (_, raw_pats) (ctxt, prop :: props) = |
24684 | 784 |
let val ctxt' = Variable.declare_term prop ctxt |
785 |
in ((prop, prep_props (set_mode mode_pattern ctxt') raw_pats), (ctxt', props)) end; |
|
786 |
||
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
787 |
val (propp, (context', _)) = (fold_map o fold_map) prep args |
24684 | 788 |
(context, prep_props (set_mode mode context) (maps (map fst) args)); |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
789 |
in (context', propp) end; |
5935 | 790 |
|
24684 | 791 |
fun gen_bind_propp mode parse_prop (ctxt, raw_args) = |
8096 | 792 |
let |
24684 | 793 |
val (ctxt', args) = prep_propp mode parse_prop (ctxt, raw_args); |
19585 | 794 |
val binds = flat (flat (map (map (simult_matches ctxt')) args)); |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
795 |
val propss = map (map #1) args; |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
796 |
|
10554 | 797 |
(*generalize result: context evaluated now, binds added later*) |
19916
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
wenzelm
parents:
19897
diff
changeset
|
798 |
val gen = Variable.exportT_terms ctxt' ctxt; |
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
799 |
fun gen_binds c = c |> bind_terms (map #1 binds ~~ map SOME (gen (map #2 binds))); |
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
800 |
in (ctxt' |> bind_terms (map (apsnd SOME) binds), (propss, gen_binds)) end; |
8096 | 801 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
802 |
in |
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
803 |
|
24684 | 804 |
val read_propp = prep_propp mode_default Syntax.read_props; |
805 |
val cert_propp = prep_propp mode_default (map o cert_prop); |
|
806 |
val read_propp_schematic = prep_propp mode_schematic Syntax.read_props; |
|
807 |
val cert_propp_schematic = prep_propp mode_schematic (map o cert_prop); |
|
10554 | 808 |
|
24684 | 809 |
val bind_propp = gen_bind_propp mode_default Syntax.read_props; |
810 |
val bind_propp_i = gen_bind_propp mode_default (map o cert_prop); |
|
811 |
val bind_propp_schematic = gen_bind_propp mode_schematic Syntax.read_props; |
|
812 |
val bind_propp_schematic_i = gen_bind_propp mode_schematic (map o cert_prop); |
|
6789 | 813 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
814 |
end; |
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
815 |
|
6789 | 816 |
|
5819 | 817 |
|
818 |
(** theorems **) |
|
819 |
||
18042 | 820 |
(* fact_tac *) |
821 |
||
27867 | 822 |
fun comp_incr_tac [] _ = no_tac |
823 |
| comp_incr_tac (th :: ths) i = |
|
824 |
(fn st => Goal.compose_hhf_tac (Drule.incr_indexes st th) i st) APPEND comp_incr_tac ths i; |
|
18042 | 825 |
|
21687 | 826 |
fun fact_tac facts = Goal.norm_hhf_tac THEN' comp_incr_tac facts; |
18122 | 827 |
|
27867 | 828 |
fun potential_facts ctxt prop = |
829 |
Facts.could_unify (facts_of ctxt) (Term.strip_all_body prop); |
|
830 |
||
831 |
fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => fact_tac (potential_facts ctxt goal) i); |
|
18042 | 832 |
|
833 |
||
6091 | 834 |
(* get_thm(s) *) |
5819 | 835 |
|
26361 | 836 |
local |
837 |
||
26687 | 838 |
fun retrieve_thms pick ctxt (Facts.Fact s) = |
16501 | 839 |
let |
27867 | 840 |
val (_, pos) = Syntax.read_token s; |
42502 | 841 |
val prop = |
842 |
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
|
843 |
|> singleton (Variable.polymorphic ctxt); |
42502 | 844 |
fun err msg = error (msg ^ Position.str_of pos ^ ":\n" ^ Syntax.string_of_term ctxt prop); |
27867 | 845 |
|
42502 | 846 |
val (prop', _) = Term.replace_dummy_patterns prop (Variable.maxidx_of ctxt + 1); |
27867 | 847 |
fun prove_fact th = |
42502 | 848 |
Goal.prove ctxt [] [] prop' (K (ALLGOALS (fact_tac [th]))); |
849 |
val results = map_filter (try prove_fact) (potential_facts ctxt prop'); |
|
27867 | 850 |
val res = |
42502 | 851 |
(case distinct Thm.eq_thm_prop results of |
852 |
[res] => res |
|
853 |
| [] => err "Failed to retrieve literal fact" |
|
854 |
| _ => err "Ambiguous specification of literal fact"); |
|
27867 | 855 |
in pick "" [res] end |
26687 | 856 |
| retrieve_thms pick ctxt xthmref = |
18042 | 857 |
let |
858 |
val thy = theory_of ctxt; |
|
26284 | 859 |
val local_facts = facts_of ctxt; |
26673 | 860 |
val thmref = Facts.map_name_of_ref (Facts.intern local_facts) xthmref; |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26321
diff
changeset
|
861 |
val name = Facts.name_of_ref thmref; |
27821 | 862 |
val pos = Facts.pos_of_ref xthmref; |
24012 | 863 |
val thms = |
864 |
if name = "" then [Thm.transfer thy Drule.dummy_thm] |
|
865 |
else |
|
26393 | 866 |
(case Facts.lookup (Context.Proof ctxt) local_facts name of |
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39441
diff
changeset
|
867 |
SOME (_, ths) => |
42378 | 868 |
(Context_Position.report ctxt pos |
42379 | 869 |
(Name_Space.markup (Facts.space_of local_facts) name); |
42378 | 870 |
map (Thm.transfer thy) (Facts.select thmref ths)) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39508
diff
changeset
|
871 |
| NONE => Global_Theory.get_fact (Context.Proof ctxt) thy xthmref); |
24012 | 872 |
in pick name thms end; |
5819 | 873 |
|
26361 | 874 |
in |
26346
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
875 |
|
26687 | 876 |
val get_fact = retrieve_thms (K I); |
877 |
val get_fact_single = retrieve_thms Facts.the_single; |
|
26346
17debd2fff8e
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26336
diff
changeset
|
878 |
|
26361 | 879 |
fun get_thms ctxt = get_fact ctxt o Facts.named; |
880 |
fun get_thm ctxt = get_fact_single ctxt o Facts.named; |
|
881 |
||
882 |
end; |
|
5819 | 883 |
|
884 |
||
26284 | 885 |
(* facts *) |
5819 | 886 |
|
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
887 |
local |
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
888 |
|
28965 | 889 |
fun update_thms _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b)) |
28861 | 890 |
| update_thms do_props (b, SOME ths) ctxt = ctxt |> map_facts |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42363
diff
changeset
|
891 |
(Facts.add_local ctxt do_props (naming_of ctxt) (b, ths) #> snd); |
5819 | 892 |
|
30761
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
893 |
in |
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
894 |
|
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
895 |
fun note_thmss kind = fold_map (fn ((b, more_attrs), raw_facts) => fn ctxt => |
5819 | 896 |
let |
28965 | 897 |
val name = full_name ctxt b; |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39508
diff
changeset
|
898 |
val facts = Global_Theory.name_thmss false name raw_facts; |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
899 |
fun app (th, attrs) x = |
30761
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
900 |
swap (Library.foldl_map |
ac7570d80c3d
renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
wenzelm
parents:
30758
diff
changeset
|
901 |
(Thm.proof_attributes (surround (Thm.kind kind) (attrs @ more_attrs))) (x, th)); |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
902 |
val (res, ctxt') = fold_map app facts ctxt; |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39508
diff
changeset
|
903 |
val thms = Global_Theory.name_thms false false name (flat res); |
24388
cf24894b81ff
added inner syntax mode, includes former type_mode and is_stmt;
wenzelm
parents:
24371
diff
changeset
|
904 |
val Mode {stmt, ...} = get_mode ctxt; |
28861 | 905 |
in ((name, thms), ctxt' |> update_thms stmt (b, SOME thms)) end); |
12711 | 906 |
|
28417 | 907 |
fun put_thms do_props thms ctxt = ctxt |
908 |
|> map_naming (K local_naming) |
|
33383 | 909 |
|> Context_Position.set_visible false |
28965 | 910 |
|> update_thms do_props (apfst Binding.name thms) |
33383 | 911 |
|> Context_Position.restore_visible ctxt |
28417 | 912 |
|> restore_naming ctxt; |
28082
37350f301128
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
913 |
|
12711 | 914 |
end; |
9196 | 915 |
|
5819 | 916 |
|
917 |
||
35680 | 918 |
(** basic logical entities **) |
17360
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents:
17221
diff
changeset
|
919 |
|
8096 | 920 |
(* variables *) |
921 |
||
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
922 |
fun declare_var (x, opt_T, mx) ctxt = |
42287
d98eb048a2e4
discontinued special treatment of structure Mixfix;
wenzelm
parents:
42267
diff
changeset
|
923 |
let val T = (case opt_T of SOME T => T | NONE => Mixfix.mixfixT mx) |
20163 | 924 |
in ((x, T, mx), ctxt |> Variable.declare_constraints (Free (x, T))) end; |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
925 |
|
10381 | 926 |
local |
927 |
||
25353
17f04d987f37
removed unused read_termTs_schematic, read/cert_vars_legacy, add_fixes_legacy;
wenzelm
parents:
25345
diff
changeset
|
928 |
fun prep_vars prep_typ internal = |
35129 | 929 |
fold_map (fn (b, raw_T, mx) => fn ctxt => |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
930 |
let |
42494 | 931 |
val x = Variable.check_name b; |
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
42287
diff
changeset
|
932 |
val _ = Lexicon.is_identifier (no_skolem internal x) orelse |
42381
309ec68442c6
added Binding.print convenience, which includes quote already;
wenzelm
parents:
42379
diff
changeset
|
933 |
error ("Illegal variable name: " ^ Binding.print b); |
12504 | 934 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
935 |
fun cond_tvars T = |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
936 |
if internal then T |
18678 | 937 |
else Type.no_tvars T handle TYPE (msg, _, _) => error msg; |
24277 | 938 |
val opt_T = Option.map (cond_tvars o cert_typ ctxt o prep_typ ctxt) raw_T; |
35129 | 939 |
val (_, ctxt') = ctxt |> declare_var (x, opt_T, mx); |
940 |
in ((b, opt_T, mx), ctxt') end); |
|
8096 | 941 |
|
10381 | 942 |
in |
943 |
||
25353
17f04d987f37
removed unused read_termTs_schematic, read/cert_vars_legacy, add_fixes_legacy;
wenzelm
parents:
25345
diff
changeset
|
944 |
val read_vars = prep_vars Syntax.parse_typ false; |
17f04d987f37
removed unused read_termTs_schematic, read/cert_vars_legacy, add_fixes_legacy;
wenzelm
parents:
25345
diff
changeset
|
945 |
val cert_vars = prep_vars (K I) true; |
8096 | 946 |
|
10381 | 947 |
end; |
948 |
||
8096 | 949 |
|
21744 | 950 |
(* notation *) |
951 |
||
24949 | 952 |
local |
953 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
954 |
fun type_syntax (Type (c, args), mx) = |
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
42287
diff
changeset
|
955 |
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
|
956 |
| type_syntax _ = NONE; |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
957 |
|
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
958 |
fun const_syntax _ (Free (x, T), mx) = SOME (Local_Syntax.Fixed, (x, T, mx)) |
24949 | 959 |
| const_syntax ctxt (Const (c, _), mx) = |
35255 | 960 |
(case try (Consts.type_scheme (consts_of ctxt)) c of |
42290
b1f544c84040
discontinued special treatment of structure Lexicon;
wenzelm
parents:
42287
diff
changeset
|
961 |
SOME T => SOME (Local_Syntax.Const, (Lexicon.mark_const c, T, mx)) |
35255 | 962 |
| NONE => NONE) |
24949 | 963 |
| const_syntax _ _ = NONE; |
964 |
||
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
965 |
fun gen_notation syntax add mode args ctxt = |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
966 |
ctxt |> map_syntax |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
967 |
(Local_Syntax.update_modesyntax (theory_of ctxt) add mode (map_filter (syntax ctxt) args)); |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
968 |
|
24949 | 969 |
in |
21744 | 970 |
|
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
971 |
val type_notation = gen_notation (K type_syntax); |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
972 |
val notation = gen_notation const_syntax; |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
973 |
|
38328 | 974 |
fun target_type_notation add mode args phi = |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
975 |
let |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
976 |
val args' = args |> map_filter (fn (T, mx) => |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
977 |
let |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
978 |
val T' = Morphism.typ phi T; |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
979 |
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
|
980 |
in if similar then SOME (T', mx) else NONE end); |
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
981 |
in Context.mapping (Sign.type_notation add mode args') (type_notation add mode args') end; |
24949 | 982 |
|
983 |
fun target_notation add mode args phi = |
|
33537
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents:
33519
diff
changeset
|
984 |
let |
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents:
33519
diff
changeset
|
985 |
val args' = args |> map_filter (fn (t, mx) => |
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents:
33519
diff
changeset
|
986 |
let val t' = Morphism.term phi t |
06c87d2c5b5a
locale_const/target_notation: uniform use of Term.aconv_untyped;
wenzelm
parents:
33519
diff
changeset
|
987 |
in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end); |
24949 | 988 |
in Context.mapping (Sign.notation add mode args') (notation add mode args') end; |
989 |
||
35680 | 990 |
end; |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
35399
diff
changeset
|
991 |
|
35680 | 992 |
|
993 |
(* aliases *) |
|
994 |
||
36450 | 995 |
fun class_alias b c ctxt = (map_tsig o apfst) (Type.class_alias (naming_of ctxt) b c) ctxt; |
996 |
fun type_alias b c ctxt = (map_tsig o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt; |
|
35680 | 997 |
fun const_alias b c ctxt = (map_consts o apfst) (Consts.alias (naming_of ctxt) b c) ctxt; |
21744 | 998 |
|
999 |
||
24767 | 1000 |
(* local constants *) |
1001 |
||
1002 |
fun add_const_constraint (c, opt_T) ctxt = |
|
1003 |
let |
|
1004 |
fun prepT raw_T = |
|
1005 |
let val T = cert_typ ctxt raw_T |
|
1006 |
in cert_term ctxt (Const (c, T)); T end; |
|
25039 | 1007 |
in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end; |
19001 | 1008 |
|
33173
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
wenzelm
parents:
33165
diff
changeset
|
1009 |
fun add_abbrev mode (b, raw_t) ctxt = |
19001 | 1010 |
let |
24675
2be1253a20d3
removed obsolete set_expand_abbrevs (superceded by mode_abbrev);
wenzelm
parents:
24612
diff
changeset
|
1011 |
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
|
1012 |
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
|
1013 |
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
|
1014 |
val ((lhs, rhs), consts') = consts_of ctxt |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42363
diff
changeset
|
1015 |
|> Consts.abbreviate ctxt (tsig_of ctxt) (naming_of ctxt) mode (b, t); |
19001 | 1016 |
in |
1017 |
ctxt |
|
25039 | 1018 |
|> (map_consts o apfst) (K consts') |
21803 | 1019 |
|> Variable.declare_term rhs |
1020 |
|> pair (lhs, rhs) |
|
21704 | 1021 |
end; |
19001 | 1022 |
|
25052 | 1023 |
fun revert_abbrev mode c = (map_consts o apfst) (Consts.revert_abbrev mode c); |
1024 |
||
19001 | 1025 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1026 |
(* fixes *) |
5819 | 1027 |
|
30763
6976521b4263
renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
wenzelm
parents:
30761
diff
changeset
|
1028 |
fun add_fixes raw_vars ctxt = |
8096 | 1029 |
let |
42491
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1030 |
val thy = theory_of ctxt; |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1031 |
val vars = #1 (cert_vars raw_vars ctxt); |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1032 |
in |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1033 |
ctxt |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1034 |
|> Variable.add_fixes_binding (map #1 vars) |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1035 |
|-> (fn xs => |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1036 |
fold_map declare_var (map2 (fn x => fn (_, T, mx) => (x, T, mx)) xs vars) |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1037 |
#-> (map_syntax o Local_Syntax.add_syntax thy o map (pair Local_Syntax.Fixed)) |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1038 |
#> pair xs) |
4bb5de0aae66
more precise position information via Variable.add_fixes_binding;
wenzelm
parents:
42488
diff
changeset
|
1039 |
end; |
5819 | 1040 |
|
32784 | 1041 |
fun auto_fixes (ctxt, (propss, x)) = |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21269
diff
changeset
|
1042 |
((fold o fold) Variable.auto_fixes propss ctxt, (propss, x)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1043 |
|
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
1044 |
|
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
1045 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1046 |
(** assumptions **) |
18187 | 1047 |
|
20209 | 1048 |
local |
1049 |
||
1050 |
fun gen_assms prepp exp args ctxt = |
|
1051 |
let |
|
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1052 |
val cert = Thm.cterm_of (theory_of ctxt); |
20209 | 1053 |
val (propss, ctxt1) = swap (prepp (ctxt, map snd args)); |
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1054 |
val _ = Variable.warn_extra_tfrees ctxt ctxt1; |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1055 |
val (premss, ctxt2) = fold_burrow (Assumption.add_assms exp o map cert) propss ctxt1; |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1056 |
in |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1057 |
ctxt2 |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1058 |
|> auto_bind_facts (flat propss) |
33644
5266a72e0889
eliminated slightly odd (unused) "axiom" and "assumption" -- collapsed to unspecific "";
wenzelm
parents:
33537
diff
changeset
|
1059 |
|> 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
|
1060 |
end; |
20209 | 1061 |
|
1062 |
in |
|
1063 |
||
1064 |
val add_assms = gen_assms (apsnd #1 o bind_propp); |
|
1065 |
val add_assms_i = gen_assms (apsnd #1 o bind_propp_i); |
|
1066 |
||
1067 |
end; |
|
1068 |
||
1069 |
||
5819 | 1070 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1071 |
(** cases **) |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1072 |
|
16147 | 1073 |
local |
1074 |
||
16668 | 1075 |
fun rem_case name = remove (fn (x: string, (y, _)) => x = y) name; |
16147 | 1076 |
|
18476 | 1077 |
fun add_case _ ("", _) cases = cases |
1078 |
| add_case _ (name, NONE) cases = rem_case name cases |
|
1079 |
| add_case is_proper (name, SOME c) cases = (name, (c, is_proper)) :: rem_case name cases; |
|
16147 | 1080 |
|
18678 | 1081 |
fun prep_case name fxs c = |
18609 | 1082 |
let |
1083 |
fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys |
|
1084 |
| replace [] ys = ys |
|
18678 | 1085 |
| replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name); |
33368 | 1086 |
val Rule_Cases.Case {fixes, assumes, binds, cases} = c; |
18609 | 1087 |
val fixes' = replace fxs fixes; |
1088 |
val binds' = map drop_schematic binds; |
|
1089 |
in |
|
1090 |
if null (fold (Term.add_tvarsT o snd) fixes []) andalso |
|
1091 |
null (fold (fold Term.add_vars o snd) assumes []) then |
|
33368 | 1092 |
Rule_Cases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases} |
18678 | 1093 |
else error ("Illegal schematic variable(s) in case " ^ quote name) |
18609 | 1094 |
end; |
1095 |
||
42501
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
1096 |
fun fix (b, T) ctxt = |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
1097 |
let val ([x], ctxt') = add_fixes [(b, SOME T, NoSyn)] ctxt |
2b8c34f53388
eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents:
42496
diff
changeset
|
1098 |
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
|
1099 |
|
16147 | 1100 |
in |
1101 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1102 |
fun add_cases is_proper = map_cases o fold (add_case is_proper); |
18609 | 1103 |
|
1104 |
fun case_result c ctxt = |
|
1105 |
let |
|
33368 | 1106 |
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
|
1107 |
val (ts, ctxt') = ctxt |> fold_map fix fixes; |
33368 | 1108 |
val Rule_Cases.Case {assumes, binds, cases, ...} = Rule_Cases.apply ts c; |
18609 | 1109 |
in |
1110 |
ctxt' |
|
30757
2d2076300185
replaced add_binds(_i) by bind_terms -- internal version only;
wenzelm
parents:
30723
diff
changeset
|
1111 |
|> bind_terms (map drop_schematic binds) |
18609 | 1112 |
|> add_cases true (map (apsnd SOME) cases) |
1113 |
|> pair (assumes, (binds, cases)) |
|
1114 |
end; |
|
1115 |
||
1116 |
val apply_case = apfst fst oo case_result; |
|
1117 |
||
16540 | 1118 |
fun get_case ctxt name xs = |
17184 | 1119 |
(case AList.lookup (op =) (cases_of ctxt) name of |
18678 | 1120 |
NONE => error ("Unknown case: " ^ quote name) |
1121 |
| SOME (c, _) => prep_case name xs c); |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1122 |
|
16147 | 1123 |
end; |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1124 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1125 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1126 |
|
10810 | 1127 |
(** print context information **) |
1128 |
||
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1129 |
(* local syntax *) |
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1130 |
|
12093 | 1131 |
val print_syntax = Syntax.print_syntax o syn_of; |
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1132 |
|
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1133 |
|
21728 | 1134 |
(* abbreviations *) |
18971 | 1135 |
|
21728 | 1136 |
fun pretty_abbrevs show_globals ctxt = |
18971 | 1137 |
let |
25052 | 1138 |
val ((space, consts), (_, globals)) = |
19033
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
1139 |
pairself (#constants o Consts.dest) (#consts (rep_context ctxt)); |
21803 | 1140 |
fun add_abbr (_, (_, NONE)) = I |
25406 | 1141 |
| add_abbr (c, (T, SOME t)) = |
21728 | 1142 |
if not show_globals andalso Symtab.defined globals c then I |
1143 |
else cons (c, Logic.mk_equals (Const (c, T), t)); |
|
42358
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents:
42357
diff
changeset
|
1144 |
val abbrevs = |
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents:
42357
diff
changeset
|
1145 |
Name_Space.extern_table ctxt (space, Symtab.make (Symtab.fold add_abbr consts [])); |
18971 | 1146 |
in |
35141 | 1147 |
if null abbrevs then [] |
21728 | 1148 |
else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)] |
18971 | 1149 |
end; |
1150 |
||
21728 | 1151 |
val print_abbrevs = Pretty.writeln o Pretty.chunks o pretty_abbrevs true; |
1152 |
||
18971 | 1153 |
|
10810 | 1154 |
(* term bindings *) |
1155 |
||
16540 | 1156 |
fun pretty_binds ctxt = |
10810 | 1157 |
let |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
1158 |
val binds = Variable.binds_of ctxt; |
21728 | 1159 |
fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t)); |
10810 | 1160 |
in |
35141 | 1161 |
if Vartab.is_empty binds then [] |
15758
07e382399a96
binds/thms: do not store options, but delete from table;
wenzelm
parents:
15750
diff
changeset
|
1162 |
else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))] |
10810 | 1163 |
end; |
1164 |
||
1165 |
val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds; |
|
1166 |
||
1167 |
||
1168 |
(* local theorems *) |
|
1169 |
||
16540 | 1170 |
fun pretty_lthms ctxt = |
20012 | 1171 |
let |
26284 | 1172 |
val local_facts = facts_of ctxt; |
1173 |
val props = Facts.props local_facts; |
|
26673 | 1174 |
val facts = |
35141 | 1175 |
(if null props then [] else [("<unnamed>", props)]) @ |
28212 | 1176 |
Facts.dest_static [] local_facts; |
20012 | 1177 |
in |
35141 | 1178 |
if null facts then [] |
28212 | 1179 |
else [Pretty.big_list "facts:" (map #1 (sort_wrt (#1 o #2) (map (`(pretty_fact ctxt)) facts)))] |
20012 | 1180 |
end; |
10810 | 1181 |
|
12057 | 1182 |
val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms; |
10810 | 1183 |
|
1184 |
||
1185 |
(* local contexts *) |
|
1186 |
||
26722 | 1187 |
local |
1188 |
||
1189 |
fun pretty_case (name, (fixes, ((asms, (lets, cs)), ctxt))) = |
|
10810 | 1190 |
let |
24922 | 1191 |
val prt_term = Syntax.pretty_term ctxt; |
12057 | 1192 |
|
10810 | 1193 |
fun prt_let (xi, t) = Pretty.block |
10818 | 1194 |
[Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1, |
10810 | 1195 |
Pretty.quote (prt_term t)]; |
1196 |
||
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1197 |
fun prt_asm (a, ts) = Pretty.block (Pretty.breaks |
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1198 |
((if a = "" then [] else [Pretty.str (a ^ ":")]) @ map (Pretty.quote o prt_term) ts)); |
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1199 |
|
10810 | 1200 |
fun prt_sect _ _ _ [] = [] |
35141 | 1201 |
| prt_sect s sep prt xs = |
1202 |
[Pretty.block (Pretty.breaks (Pretty.str s :: |
|
1203 |
flat (separate sep (map (single o prt) xs))))]; |
|
26722 | 1204 |
in |
1205 |
Pretty.block (Pretty.fbreaks |
|
10810 | 1206 |
(Pretty.str (name ^ ":") :: |
42496 | 1207 |
prt_sect "fix" [] (Pretty.str o Binding.name_of o fst) fixes @ |
10810 | 1208 |
prt_sect "let" [Pretty.str "and"] prt_let |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19422
diff
changeset
|
1209 |
(map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @ |
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1210 |
(if forall (null o #2) asms then [] |
18609 | 1211 |
else prt_sect "assume" [Pretty.str "and"] prt_asm asms) @ |
26722 | 1212 |
prt_sect "subcases:" [] (Pretty.str o fst) cs)) |
1213 |
end; |
|
16540 | 1214 |
|
26722 | 1215 |
in |
1216 |
||
1217 |
fun pretty_cases ctxt = |
|
1218 |
let |
|
18476 | 1219 |
fun add_case (_, (_, false)) = I |
33368 | 1220 |
| add_case (name, (c as Rule_Cases.Case {fixes, ...}, true)) = |
26722 | 1221 |
cons (name, (fixes, case_result c ctxt)); |
18476 | 1222 |
val cases = fold add_case (cases_of ctxt) []; |
10810 | 1223 |
in |
35141 | 1224 |
if null cases then [] |
26722 | 1225 |
else [Pretty.big_list "cases:" (map pretty_case cases)] |
10810 | 1226 |
end; |
1227 |
||
1228 |
val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases; |
|
1229 |
||
26722 | 1230 |
end; |
1231 |
||
10810 | 1232 |
|
12057 | 1233 |
(* core context *) |
10810 | 1234 |
|
42717
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
1235 |
val debug = Config.bool (Config.declare "Proof_Context.debug" (K (Config.Bool false))); |
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
1236 |
val verbose = Config.bool (Config.declare "Proof_Context.verbose" (K (Config.Bool false))); |
10810 | 1237 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1238 |
fun pretty_ctxt ctxt = |
42717
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
1239 |
if not (Config.get ctxt debug) then [] |
20310 | 1240 |
else |
1241 |
let |
|
24922 | 1242 |
val prt_term = Syntax.pretty_term ctxt; |
12057 | 1243 |
|
20310 | 1244 |
(*structures*) |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
1245 |
val {structs, ...} = Local_Syntax.idents_of (syntax_of ctxt); |
35139 | 1246 |
val prt_structs = |
1247 |
if null structs then [] |
|
20310 | 1248 |
else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 :: |
1249 |
Pretty.commas (map Pretty.str structs))]; |
|
12093 | 1250 |
|
20310 | 1251 |
(*fixes*) |
1252 |
fun prt_fix (x, x') = |
|
1253 |
if x = x' then Pretty.str x |
|
1254 |
else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')]; |
|
1255 |
val fixes = |
|
42488
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
1256 |
filter_out ((can Name.dest_internal orf member (op =) structs) o #1) |
4638622bcaa1
reorganized fixes as specialized (global) name space;
wenzelm
parents:
42469
diff
changeset
|
1257 |
(Variable.dest_fixes ctxt); |
35139 | 1258 |
val prt_fixes = |
1259 |
if null fixes then [] |
|
20310 | 1260 |
else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 :: |
1261 |
Pretty.commas (map prt_fix fixes))]; |
|
12057 | 1262 |
|
20310 | 1263 |
(*prems*) |
30473
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
wenzelm
parents:
30469
diff
changeset
|
1264 |
val prems = Assumption.all_prems_of ctxt; |
35139 | 1265 |
val prt_prems = |
1266 |
if null prems then [] |
|
39165 | 1267 |
else [Pretty.big_list "prems:" (map (Display.pretty_thm ctxt) prems)]; |
20310 | 1268 |
in prt_structs @ prt_fixes @ prt_prems end; |
10810 | 1269 |
|
1270 |
||
1271 |
(* main context *) |
|
1272 |
||
16540 | 1273 |
fun pretty_context ctxt = |
10810 | 1274 |
let |
42717
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
1275 |
val verbose = Config.get ctxt verbose; |
0bbb56867091
proper configuration options Proof_Context.debug and Proof_Context.verbose;
wenzelm
parents:
42502
diff
changeset
|
1276 |
fun verb f x = if verbose then f (x ()) else []; |
35141 | 1277 |
|
24922 | 1278 |
val prt_term = Syntax.pretty_term ctxt; |
1279 |
val prt_typ = Syntax.pretty_typ ctxt; |
|
1280 |
val prt_sort = Syntax.pretty_sort ctxt; |
|
10810 | 1281 |
|
1282 |
(*theory*) |
|
12057 | 1283 |
val pretty_thy = Pretty.block |
17384 | 1284 |
[Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)]; |
10810 | 1285 |
|
1286 |
(*defaults*) |
|
1287 |
fun prt_atom prt prtT (x, X) = Pretty.block |
|
1288 |
[prt x, Pretty.str " ::", Pretty.brk 1, prtT X]; |
|
1289 |
||
1290 |
fun prt_var (x, ~1) = prt_term (Syntax.free x) |
|
1291 |
| prt_var xi = prt_term (Syntax.var xi); |
|
1292 |
||
1293 |
fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) |
|
1294 |
| prt_varT xi = prt_typ (TVar (xi, [])); |
|
1295 |
||
1296 |
val prt_defT = prt_atom prt_var prt_typ; |
|
1297 |
val prt_defS = prt_atom prt_varT prt_sort; |
|
16540 | 1298 |
|
20163 | 1299 |
val (types, sorts) = Variable.constraints_of ctxt; |
10810 | 1300 |
in |
18609 | 1301 |
verb single (K pretty_thy) @ |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1302 |
pretty_ctxt ctxt @ |
21728 | 1303 |
verb (pretty_abbrevs false) (K ctxt) @ |
10810 | 1304 |
verb pretty_binds (K ctxt) @ |
12057 | 1305 |
verb pretty_lthms (K ctxt) @ |
10810 | 1306 |
verb pretty_cases (K ctxt) @ |
18609 | 1307 |
verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @ |
20163 | 1308 |
verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts))) |
10810 | 1309 |
end; |
1310 |
||
5819 | 1311 |
end; |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35412
diff
changeset
|
1312 |
|
42360 | 1313 |
val show_abbrevs = Proof_Context.show_abbrevs; |
40879
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
39557
diff
changeset
|
1314 |