author | wenzelm |
Fri, 27 Jul 2007 21:55:22 +0200 | |
changeset 24012 | e48e1b4557c8 |
parent 23922 | 707639e9497d |
child 24277 | 6442fde2daaa |
permissions | -rw-r--r-- |
5819 | 1 |
(* Title: Pure/Isar/proof_context.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
19001 | 5 |
The key concept of Isar proof contexts: elevates primitive local |
6 |
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
|
7 |
elements. See also structure Variable and Assumption. |
5819 | 8 |
*) |
9 |
||
10 |
signature PROOF_CONTEXT = |
|
11 |
sig |
|
20310 | 12 |
val theory_of: Proof.context -> theory |
13 |
val init: theory -> Proof.context |
|
21667
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
14 |
val is_stmt: Proof.context -> bool |
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
15 |
val set_stmt: bool -> Proof.context -> Proof.context |
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
16 |
val restore_stmt: Proof.context -> Proof.context -> Proof.context |
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
17 |
val naming_of: Proof.context -> NameSpace.naming |
20310 | 18 |
val full_name: Proof.context -> bstring -> string |
19 |
val consts_of: Proof.context -> Consts.T |
|
21183 | 20 |
val const_syntax_name: Proof.context -> string -> string |
20784 | 21 |
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context |
20310 | 22 |
val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context |
22358 | 23 |
val theorems_of: Proof.context -> thm list NameSpace.table |
20310 | 24 |
val fact_index_of: Proof.context -> FactIndex.T |
25 |
val transfer: theory -> Proof.context -> Proof.context |
|
20367 | 26 |
val theory: (theory -> theory) -> Proof.context -> Proof.context |
27 |
val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context |
|
20310 | 28 |
val pretty_term: Proof.context -> term -> Pretty.T |
21728 | 29 |
val pretty_term_abbrev: Proof.context -> term -> Pretty.T |
20310 | 30 |
val pretty_typ: Proof.context -> typ -> Pretty.T |
31 |
val pretty_sort: Proof.context -> sort -> Pretty.T |
|
32 |
val pretty_classrel: Proof.context -> class list -> Pretty.T |
|
33 |
val pretty_arity: Proof.context -> arity -> Pretty.T |
|
34 |
val pp: Proof.context -> Pretty.pp |
|
22874 | 35 |
val pretty_thm_legacy: thm -> Pretty.T |
20310 | 36 |
val pretty_thm: Proof.context -> thm -> Pretty.T |
37 |
val pretty_thms: Proof.context -> thm list -> Pretty.T |
|
38 |
val pretty_fact: Proof.context -> string * thm list -> Pretty.T |
|
39 |
val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T |
|
40 |
val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T |
|
41 |
val string_of_typ: Proof.context -> typ -> string |
|
42 |
val string_of_term: Proof.context -> term -> string |
|
43 |
val string_of_thm: Proof.context -> thm -> string |
|
44 |
val read_typ: Proof.context -> string -> typ |
|
45 |
val read_typ_syntax: Proof.context -> string -> typ |
|
46 |
val read_typ_abbrev: Proof.context -> string -> typ |
|
47 |
val cert_typ: Proof.context -> typ -> typ |
|
48 |
val cert_typ_syntax: Proof.context -> typ -> typ |
|
49 |
val cert_typ_abbrev: Proof.context -> typ -> typ |
|
50 |
val get_skolem: Proof.context -> string -> string |
|
51 |
val revert_skolem: Proof.context -> string -> string |
|
52 |
val revert_skolems: Proof.context -> term -> term |
|
22763 | 53 |
val decode_term: Proof.context -> term -> term |
20310 | 54 |
val read_termTs_schematic: Proof.context -> (string -> bool) -> (indexname -> typ option) |
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
55 |
-> (indexname -> sort option) -> string list -> (string * typ) list |
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
56 |
-> term list * (indexname * typ) list |
22712 | 57 |
val read_prop_legacy: Proof.context -> string -> term |
20310 | 58 |
val read_term: Proof.context -> string -> term |
59 |
val read_prop: Proof.context -> string -> term |
|
60 |
val read_prop_schematic: Proof.context -> string -> term |
|
22773 | 61 |
val read_termTs: Proof.context -> (string * typ) list -> term list |
62 |
val read_terms: Proof.context -> string list -> term list |
|
20310 | 63 |
val read_term_pats: typ -> Proof.context -> string list -> term list |
64 |
val read_prop_pats: Proof.context -> string list -> term list |
|
21728 | 65 |
val read_term_abbrev: Proof.context -> string -> term |
20310 | 66 |
val cert_term: Proof.context -> term -> term |
22773 | 67 |
val cert_terms: Proof.context -> term list -> term list |
20310 | 68 |
val cert_prop: Proof.context -> term -> term |
69 |
val cert_term_pats: typ -> Proof.context -> term list -> term list |
|
70 |
val cert_prop_pats: Proof.context -> term list -> term list |
|
22728 | 71 |
val infer_types: Proof.context -> term list -> term list |
72 |
val infer_types_pats: Proof.context -> term list -> term list |
|
20310 | 73 |
val infer_type: Proof.context -> string -> typ |
74 |
val inferred_param: string -> Proof.context -> (string * typ) * Proof.context |
|
75 |
val inferred_fixes: Proof.context -> (string * typ) list * Proof.context |
|
76 |
val read_tyname: Proof.context -> string -> typ |
|
77 |
val read_const: Proof.context -> string -> term |
|
78 |
val goal_export: Proof.context -> Proof.context -> thm list -> thm list |
|
79 |
val export: Proof.context -> Proof.context -> thm list -> thm list |
|
21531 | 80 |
val export_morphism: Proof.context -> Proof.context -> morphism |
20310 | 81 |
val add_binds: (indexname * string option) list -> Proof.context -> Proof.context |
82 |
val add_binds_i: (indexname * term option) list -> Proof.context -> Proof.context |
|
83 |
val auto_bind_goal: term list -> Proof.context -> Proof.context |
|
84 |
val auto_bind_facts: term list -> Proof.context -> Proof.context |
|
85 |
val match_bind: bool -> (string list * string) list -> Proof.context -> term list * Proof.context |
|
86 |
val match_bind_i: bool -> (term list * term) list -> Proof.context -> term list * Proof.context |
|
87 |
val read_propp: Proof.context * (string * string list) list list |
|
88 |
-> Proof.context * (term * term list) list list |
|
89 |
val cert_propp: Proof.context * (term * term list) list list |
|
90 |
-> Proof.context * (term * term list) list list |
|
91 |
val read_propp_schematic: Proof.context * (string * string list) list list |
|
92 |
-> Proof.context * (term * term list) list list |
|
93 |
val cert_propp_schematic: Proof.context * (term * term list) list list |
|
94 |
-> Proof.context * (term * term list) list list |
|
95 |
val bind_propp: Proof.context * (string * string list) list list |
|
96 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
97 |
val bind_propp_i: Proof.context * (term * term list) list list |
|
98 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
99 |
val bind_propp_schematic: Proof.context * (string * string list) list list |
|
100 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
101 |
val bind_propp_schematic_i: Proof.context * (term * term list) list list |
|
102 |
-> Proof.context * (term list list * (Proof.context -> Proof.context)) |
|
18042 | 103 |
val fact_tac: thm list -> int -> tactic |
20310 | 104 |
val some_fact_tac: Proof.context -> int -> tactic |
105 |
val get_thm: Proof.context -> thmref -> thm |
|
106 |
val get_thm_closure: Proof.context -> thmref -> thm |
|
107 |
val get_thms: Proof.context -> thmref -> thm list |
|
108 |
val get_thms_closure: Proof.context -> thmref -> thm list |
|
109 |
val valid_thms: Proof.context -> string * thm list -> bool |
|
110 |
val lthms_containing: Proof.context -> FactIndex.spec -> (string * thm list) list |
|
22352 | 111 |
val add_path: string -> Proof.context -> Proof.context |
20310 | 112 |
val no_base_names: Proof.context -> Proof.context |
113 |
val qualified_names: Proof.context -> Proof.context |
|
114 |
val sticky_prefix: string -> Proof.context -> Proof.context |
|
115 |
val restore_naming: Proof.context -> Proof.context -> Proof.context |
|
21728 | 116 |
val reset_naming: Proof.context -> Proof.context |
20310 | 117 |
val hide_thms: bool -> string list -> Proof.context -> Proof.context |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
118 |
val put_thms: string * thm list option -> Proof.context -> Proof.context |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
119 |
val note_thmss: string -> |
18728 | 120 |
((bstring * attribute list) * (thmref * attribute list) list) list -> |
20310 | 121 |
Proof.context -> (bstring * thm list) list * Proof.context |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
122 |
val note_thmss_i: string -> |
18728 | 123 |
((bstring * attribute list) * (thm list * attribute list) list) list -> |
20310 | 124 |
Proof.context -> (bstring * thm list) list * Proof.context |
125 |
val read_vars: (string * string option * mixfix) list -> Proof.context -> |
|
126 |
(string * typ option * mixfix) list * Proof.context |
|
127 |
val cert_vars: (string * typ option * mixfix) list -> Proof.context -> |
|
128 |
(string * typ option * mixfix) list * Proof.context |
|
129 |
val read_vars_legacy: (string * string option * mixfix) list -> Proof.context -> |
|
130 |
(string * typ option * mixfix) list * Proof.context |
|
131 |
val cert_vars_legacy: (string * typ option * mixfix) list -> Proof.context -> |
|
132 |
(string * typ option * mixfix) list * Proof.context |
|
133 |
val add_fixes: (string * string option * mixfix) list -> |
|
134 |
Proof.context -> string list * Proof.context |
|
135 |
val add_fixes_i: (string * typ option * mixfix) list -> |
|
136 |
Proof.context -> string list * Proof.context |
|
137 |
val add_fixes_legacy: (string * typ option * mixfix) list -> |
|
138 |
Proof.context -> string list * Proof.context |
|
139 |
val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a) |
|
140 |
val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context |
|
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
141 |
val add_assms: Assumption.export -> |
19585 | 142 |
((string * attribute list) * (string * string list) list) list -> |
20310 | 143 |
Proof.context -> (bstring * thm list) list * Proof.context |
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
144 |
val add_assms_i: Assumption.export -> |
19585 | 145 |
((string * attribute list) * (term * term list) list) list -> |
20310 | 146 |
Proof.context -> (bstring * thm list) list * Proof.context |
147 |
val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context |
|
148 |
val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context |
|
149 |
val get_case: Proof.context -> string -> string option list -> RuleCases.T |
|
21208
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
150 |
val add_notation: Syntax.mode -> (term * mixfix) list -> |
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
151 |
Proof.context -> Proof.context |
21744 | 152 |
val target_notation: Syntax.mode -> (term * mixfix) list -> morphism -> |
153 |
Context.generic -> Context.generic |
|
21728 | 154 |
val set_expand_abbrevs: bool -> Proof.context -> Proof.context |
21803 | 155 |
val add_abbrev: string -> bstring * term -> Proof.context -> (term * term) * Proof.context |
10810 | 156 |
val verbose: bool ref |
157 |
val setmp_verbose: ('a -> 'b) -> 'a -> 'b |
|
20310 | 158 |
val print_syntax: Proof.context -> unit |
21728 | 159 |
val print_abbrevs: Proof.context -> unit |
20310 | 160 |
val print_binds: Proof.context -> unit |
161 |
val print_lthms: Proof.context -> unit |
|
162 |
val print_cases: Proof.context -> unit |
|
163 |
val debug: bool ref |
|
10810 | 164 |
val prems_limit: int ref |
20310 | 165 |
val pretty_ctxt: Proof.context -> Pretty.T list |
166 |
val pretty_context: Proof.context -> Pretty.T list |
|
5819 | 167 |
end; |
168 |
||
16540 | 169 |
structure ProofContext: PROOF_CONTEXT = |
5819 | 170 |
struct |
171 |
||
16540 | 172 |
val theory_of = Context.theory_of_proof; |
19001 | 173 |
val tsig_of = Sign.tsig_of o theory_of; |
174 |
||
16540 | 175 |
val init = Context.init_proof; |
12057 | 176 |
|
7270 | 177 |
|
5819 | 178 |
|
16540 | 179 |
(** Isar proof context information **) |
5819 | 180 |
|
16540 | 181 |
datatype ctxt = |
182 |
Ctxt of |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
183 |
{is_stmt: bool, (*inner statement mode*) |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
184 |
naming: NameSpace.naming, (*local naming conventions*) |
19001 | 185 |
syntax: LocalSyntax.T, (*local syntax*) |
19033
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
186 |
consts: Consts.T * Consts.T, (*global/local consts*) |
19001 | 187 |
thms: thm list NameSpace.table * FactIndex.T, (*local thms*) |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
188 |
cases: (string * (RuleCases.T * bool)) list}; (*local contexts*) |
5819 | 189 |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
190 |
fun make_ctxt (is_stmt, naming, syntax, consts, thms, cases) = |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
191 |
Ctxt {is_stmt = is_stmt, naming = naming, syntax = syntax, |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
192 |
consts = consts, thms = thms, cases = cases}; |
5819 | 193 |
|
19079
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
wenzelm
parents:
19062
diff
changeset
|
194 |
val local_naming = NameSpace.default_naming |> NameSpace.add_path "local"; |
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
wenzelm
parents:
19062
diff
changeset
|
195 |
|
16540 | 196 |
structure ContextData = ProofDataFun |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
197 |
( |
16540 | 198 |
type T = ctxt; |
199 |
fun init thy = |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
200 |
make_ctxt (false, local_naming, LocalSyntax.init thy, |
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
201 |
(Sign.consts_of thy, Sign.consts_of thy), |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
202 |
(NameSpace.empty_table, FactIndex.empty), []); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
203 |
); |
5819 | 204 |
|
16540 | 205 |
fun rep_context ctxt = ContextData.get ctxt |> (fn Ctxt args => args); |
5819 | 206 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
207 |
fun map_context f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
208 |
ContextData.map (fn Ctxt {is_stmt, naming, syntax, consts, thms, cases} => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
209 |
make_ctxt (f (is_stmt, naming, syntax, consts, thms, cases))); |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
210 |
|
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
211 |
fun set_stmt b = |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
212 |
map_context (fn (_, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
213 |
(b, naming, syntax, consts, thms, cases)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
214 |
|
19001 | 215 |
fun map_naming f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
216 |
map_context (fn (is_stmt, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
217 |
(is_stmt, f naming, syntax, consts, thms, cases)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
218 |
|
19001 | 219 |
fun map_syntax f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
220 |
map_context (fn (is_stmt, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
221 |
(is_stmt, naming, f syntax, consts, thms, cases)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
222 |
|
19001 | 223 |
fun map_consts f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
224 |
map_context (fn (is_stmt, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
225 |
(is_stmt, naming, syntax, f consts, thms, cases)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
226 |
|
19001 | 227 |
fun map_thms f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
228 |
map_context (fn (is_stmt, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
229 |
(is_stmt, naming, syntax, consts, f thms, cases)); |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
230 |
|
19001 | 231 |
fun map_cases f = |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
232 |
map_context (fn (is_stmt, naming, syntax, consts, thms, cases) => |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
233 |
(is_stmt, naming, syntax, consts, thms, f cases)); |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
234 |
|
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
235 |
val is_stmt = #is_stmt o rep_context; |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
236 |
fun restore_stmt ctxt = set_stmt (is_stmt ctxt); |
19001 | 237 |
|
238 |
val naming_of = #naming o rep_context; |
|
19387 | 239 |
val full_name = NameSpace.full o naming_of; |
5819 | 240 |
|
16540 | 241 |
val syntax_of = #syntax o rep_context; |
19001 | 242 |
val syn_of = LocalSyntax.syn_of o syntax_of; |
19543 | 243 |
val set_syntax_mode = map_syntax o LocalSyntax.set_mode; |
244 |
val restore_syntax_mode = map_syntax o LocalSyntax.restore_mode o syntax_of; |
|
19001 | 245 |
|
19033
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
246 |
val consts_of = #2 o #consts o rep_context; |
21183 | 247 |
val const_syntax_name = Consts.syntax_name o consts_of; |
5819 | 248 |
|
16540 | 249 |
val thms_of = #thms o rep_context; |
22358 | 250 |
val theorems_of = #1 o thms_of; |
19001 | 251 |
val fact_index_of = #2 o thms_of; |
16540 | 252 |
|
253 |
val cases_of = #cases o rep_context; |
|
5819 | 254 |
|
255 |
||
20367 | 256 |
(* theory transfer *) |
12093 | 257 |
|
19001 | 258 |
fun transfer_syntax thy = |
259 |
map_syntax (LocalSyntax.rebuild thy) #> |
|
19033
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
260 |
map_consts (fn consts as (global_consts, local_consts) => |
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
261 |
let val thy_consts = Sign.consts_of thy in |
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
262 |
if Consts.eq_consts (thy_consts, global_consts) then consts |
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
263 |
else (thy_consts, Consts.merge (thy_consts, local_consts)) |
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
264 |
end); |
17072 | 265 |
|
19001 | 266 |
fun transfer thy = Context.transfer_proof thy #> transfer_syntax thy; |
17072 | 267 |
|
20367 | 268 |
fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt; |
269 |
||
270 |
fun theory_result f ctxt = |
|
271 |
let val (res, thy') = f (theory_of ctxt) |
|
272 |
in (res, ctxt |> transfer thy') end; |
|
19019 | 273 |
|
12093 | 274 |
|
275 |
||
14828 | 276 |
(** pretty printing **) |
277 |
||
19310 | 278 |
local |
279 |
||
21744 | 280 |
fun rewrite_term warn thy rews t = |
19371 | 281 |
if can Term.type_of t then Pattern.rewrite_term thy rews [] t |
21744 | 282 |
else (if warn then warning "Printing malformed term -- cannot expand abbreviations" else (); t); |
19371 | 283 |
|
19001 | 284 |
fun pretty_term' abbrevs ctxt t = |
18971 | 285 |
let |
19001 | 286 |
val thy = theory_of ctxt; |
287 |
val syntax = syntax_of ctxt; |
|
18971 | 288 |
val consts = consts_of ctxt; |
22587 | 289 |
val do_abbrevs = abbrevs andalso not (print_mode_active "no_abbrevs"); |
19001 | 290 |
val t' = t |
21744 | 291 |
|> do_abbrevs ? rewrite_term true thy (Consts.abbrevs_of consts (! print_mode @ [""])) |
21824 | 292 |
|> do_abbrevs ? rewrite_term false thy (Consts.abbrevs_of consts [Syntax.internalM]) |
19371 | 293 |
|> Sign.extern_term (Consts.extern_early consts) thy |
294 |
|> LocalSyntax.extern_term syntax; |
|
21772 | 295 |
in Sign.pretty_term' ctxt (LocalSyntax.syn_of syntax) (Consts.extern consts) t' end; |
18971 | 296 |
|
19310 | 297 |
in |
298 |
||
19001 | 299 |
val pretty_term = pretty_term' true; |
21728 | 300 |
val pretty_term_abbrev = pretty_term' false; |
19310 | 301 |
|
302 |
end; |
|
303 |
||
16458 | 304 |
fun pretty_typ ctxt T = Sign.pretty_typ (theory_of ctxt) T; |
305 |
fun pretty_sort ctxt S = Sign.pretty_sort (theory_of ctxt) S; |
|
306 |
fun pretty_classrel ctxt cs = Sign.pretty_classrel (theory_of ctxt) cs; |
|
307 |
fun pretty_arity ctxt ar = Sign.pretty_arity (theory_of ctxt) ar; |
|
14828 | 308 |
|
14974
b1ecb7859c99
avoid premature evaluation of syn_of (wastes time in conjunction with pp);
wenzelm
parents:
14901
diff
changeset
|
309 |
fun pp ctxt = Pretty.pp (pretty_term ctxt, pretty_typ ctxt, pretty_sort ctxt, |
b1ecb7859c99
avoid premature evaluation of syn_of (wastes time in conjunction with pp);
wenzelm
parents:
14901
diff
changeset
|
310 |
pretty_classrel ctxt, pretty_arity ctxt); |
14828 | 311 |
|
22874 | 312 |
fun pretty_thm_legacy th = |
313 |
let val thy = Thm.theory_of_thm th |
|
314 |
in Display.pretty_thm_aux (pp (init thy)) true false [] th end; |
|
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
315 |
|
17451 | 316 |
fun pretty_thm ctxt th = |
22874 | 317 |
let val asms = map Thm.term_of (Assumption.assms_of ctxt) |
318 |
in Display.pretty_thm_aux (pp ctxt) false true asms th end; |
|
14828 | 319 |
|
320 |
fun pretty_thms ctxt [th] = pretty_thm ctxt th |
|
321 |
| pretty_thms ctxt ths = Pretty.blk (0, Pretty.fbreaks (map (pretty_thm ctxt) ths)); |
|
322 |
||
323 |
fun pretty_fact ctxt ("", ths) = pretty_thms ctxt ths |
|
324 |
| pretty_fact ctxt (a, [th]) = |
|
325 |
Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, pretty_thm ctxt th] |
|
326 |
| pretty_fact ctxt (a, ths) = |
|
327 |
Pretty.block (Pretty.fbreaks (Pretty.str (a ^ ":") :: map (pretty_thm ctxt) ths)); |
|
328 |
||
17072 | 329 |
fun pretty_proof ctxt prf = |
21728 | 330 |
pretty_term_abbrev (ctxt |> transfer_syntax (ProofSyntax.proof_syntax prf (theory_of ctxt))) |
17072 | 331 |
(ProofSyntax.term_of_proof prf); |
332 |
||
333 |
fun pretty_proof_of ctxt full th = |
|
334 |
pretty_proof ctxt (ProofSyntax.proof_of full th); |
|
335 |
||
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
336 |
val string_of_typ = Pretty.string_of oo pretty_typ; |
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
337 |
val string_of_term = Pretty.string_of oo pretty_term; |
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
338 |
val string_of_thm = Pretty.string_of oo pretty_thm; |
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
339 |
|
14828 | 340 |
|
341 |
||
5819 | 342 |
(** prepare types **) |
343 |
||
9504 | 344 |
local |
345 |
||
346 |
fun read_typ_aux read ctxt s = |
|
21772 | 347 |
read (syn_of ctxt) ctxt (Variable.def_sort ctxt) s; |
5819 | 348 |
|
10554 | 349 |
fun cert_typ_aux cert ctxt raw_T = |
18678 | 350 |
cert (theory_of ctxt) raw_T handle TYPE (msg, _, _) => error msg; |
9504 | 351 |
|
352 |
in |
|
353 |
||
16348 | 354 |
val read_typ = read_typ_aux Sign.read_typ'; |
355 |
val read_typ_syntax = read_typ_aux Sign.read_typ_syntax'; |
|
356 |
val read_typ_abbrev = read_typ_aux Sign.read_typ_abbrev'; |
|
357 |
val cert_typ = cert_typ_aux Sign.certify_typ; |
|
358 |
val cert_typ_syntax = cert_typ_aux Sign.certify_typ_syntax; |
|
359 |
val cert_typ_abbrev = cert_typ_aux Sign.certify_typ_abbrev; |
|
9504 | 360 |
|
361 |
end; |
|
362 |
||
5819 | 363 |
|
7679 | 364 |
(* internalize Skolem constants *) |
365 |
||
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
366 |
val lookup_skolem = AList.lookup (op =) o Variable.fixes_of; |
18187 | 367 |
fun get_skolem ctxt x = the_default x (lookup_skolem ctxt x); |
7679 | 368 |
|
18678 | 369 |
fun no_skolem internal x = |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
370 |
if can Name.dest_skolem x then |
18678 | 371 |
error ("Illegal reference to internal Skolem constant: " ^ quote x) |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
372 |
else if not internal andalso can Name.dest_internal x then |
18678 | 373 |
error ("Illegal reference to internal variable: " ^ quote x) |
7679 | 374 |
else x; |
375 |
||
376 |
||
20244
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
377 |
(* revert Skolem constants -- approximation only! *) |
18255 | 378 |
|
20244
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
379 |
fun revert_skolem ctxt = |
9133 | 380 |
let |
20244
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
381 |
val rev_fixes = map Library.swap (Variable.fixes_of ctxt); |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
382 |
val revert = AList.lookup (op =) rev_fixes; |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
383 |
in |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
384 |
fn x => |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
385 |
(case revert x of |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
386 |
SOME x' => x' |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
387 |
| NONE => perhaps (try Name.dest_skolem) x) |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
388 |
end; |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
389 |
|
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
390 |
fun revert_skolems ctxt = |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
391 |
let |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
392 |
val revert = revert_skolem ctxt; |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
393 |
fun reverts (Free (x, T)) = Free (revert x, T) |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
394 |
| reverts (t $ u) = reverts t $ reverts u |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
395 |
| reverts (Abs (x, T, t)) = Abs (x, T, reverts t) |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
396 |
| reverts a = a; |
89a407400874
replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents:
20234
diff
changeset
|
397 |
in reverts end |
9133 | 398 |
|
8096 | 399 |
|
18187 | 400 |
|
5819 | 401 |
(** prepare terms and propositions **) |
402 |
||
403 |
(* |
|
16501 | 404 |
(1) read / certify wrt. theory of context |
5819 | 405 |
(2) intern Skolem constants |
406 |
(3) expand term bindings |
|
407 |
*) |
|
408 |
||
409 |
||
16501 | 410 |
(* read wrt. theory *) (*exception ERROR*) |
5819 | 411 |
|
22712 | 412 |
fun read_def_termTs freeze pp syn ctxt (types, sorts, used) fixed sTs = |
413 |
Sign.read_def_terms' pp (Sign.is_logtype (theory_of ctxt)) syn (consts_of ctxt) fixed |
|
21772 | 414 |
ctxt (types, sorts) used freeze sTs; |
5874 | 415 |
|
22712 | 416 |
fun read_def_termT freeze pp syn ctxt defaults fixed sT = |
417 |
apfst hd (read_def_termTs freeze pp syn ctxt defaults fixed [sT]); |
|
14828 | 418 |
|
22712 | 419 |
fun read_term_thy freeze pp syn thy defaults fixed s = |
420 |
#1 (read_def_termT freeze pp syn thy defaults fixed (s, dummyT)); |
|
5874 | 421 |
|
22712 | 422 |
fun read_prop_thy freeze pp syn thy defaults fixed s = |
423 |
#1 (read_def_termT freeze pp syn thy defaults fixed (s, propT)); |
|
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
424 |
|
22712 | 425 |
fun read_terms_thy freeze pp syn thy defaults fixed = |
22773 | 426 |
#1 o read_def_termTs freeze pp syn thy defaults fixed; |
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
427 |
|
22712 | 428 |
fun read_props_thy freeze pp syn thy defaults fixed = |
429 |
#1 o read_def_termTs freeze pp syn thy defaults fixed o map (rpair propT); |
|
5819 | 430 |
|
431 |
||
19001 | 432 |
(* local abbreviations *) |
5819 | 433 |
|
19371 | 434 |
fun certify_consts ctxt = |
19001 | 435 |
Consts.certify (pp ctxt) (tsig_of ctxt) (consts_of ctxt); |
436 |
||
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
437 |
fun reject_schematic (Var (xi, _)) = |
22678 | 438 |
error ("Unbound schematic variable: " ^ Term.string_of_vname xi) |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
439 |
| reject_schematic (Abs (_, _, t)) = reject_schematic t |
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
440 |
| reject_schematic (t $ u) = (reject_schematic t; reject_schematic u) |
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
441 |
| reject_schematic _ = (); |
5819 | 442 |
|
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
443 |
fun expand_binds ctxt schematic = |
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
444 |
Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic); |
5819 | 445 |
|
446 |
||
22712 | 447 |
(* schematic type variables *) |
448 |
||
449 |
val reject_tvars = (Term.map_types o Term.map_atyps) |
|
450 |
(fn TVar (xi, _) => error ("Illegal schematic type variable: " ^ Term.string_of_vname xi) |
|
451 |
| T => T); |
|
452 |
||
453 |
||
6550 | 454 |
(* dummy patterns *) |
455 |
||
18310 | 456 |
val prepare_dummies = |
457 |
let val next = ref 1 in |
|
23922
707639e9497d
marked some CRITICAL sections (for multithreading);
wenzelm
parents:
23361
diff
changeset
|
458 |
fn t => CRITICAL (fn () => |
18310 | 459 |
let val (i, u) = Term.replace_dummy_patterns (! next, t) |
23922
707639e9497d
marked some CRITICAL sections (for multithreading);
wenzelm
parents:
23361
diff
changeset
|
460 |
in next := i; u end) |
18310 | 461 |
end; |
6762 | 462 |
|
18678 | 463 |
fun reject_dummies t = Term.no_dummy_patterns t |
464 |
handle TERM _ => error "Illegal dummy pattern(s) in term"; |
|
6550 | 465 |
|
466 |
||
22763 | 467 |
(* decoding raw terms (syntax trees) *) |
468 |
||
469 |
fun legacy_intern_skolem ctxt internal def_type x = (* FIXME cleanup this mess *) |
|
470 |
let |
|
471 |
val sko = lookup_skolem ctxt x; |
|
472 |
val is_const = can (Consts.read_const (consts_of ctxt)) x; |
|
473 |
val is_free = is_some sko orelse not is_const; |
|
474 |
val is_internal = internal x; |
|
475 |
val is_declared = is_some (def_type (x, ~1)); |
|
476 |
val res = |
|
477 |
if is_free andalso not is_internal then (no_skolem false x; sko) |
|
23361 | 478 |
else ((no_skolem false x; ()) handle ERROR msg => |
479 |
legacy_feature (msg ^ ContextPosition.str_of ctxt); |
|
22763 | 480 |
if is_internal andalso is_declared then SOME x else NONE); |
481 |
in if is_some res then res else if is_declared then SOME x else NONE end; |
|
482 |
||
483 |
fun decode_term ctxt = |
|
484 |
let |
|
485 |
val thy = theory_of ctxt; |
|
486 |
val get_sort = Sign.get_sort thy (Variable.def_sort ctxt); |
|
487 |
val map_const = try (#1 o Term.dest_Const o Consts.read_const (consts_of ctxt)); |
|
488 |
val map_free = legacy_intern_skolem ctxt (K false) (Variable.def_type ctxt false); |
|
489 |
val map_type = Sign.intern_tycons thy; |
|
490 |
val map_sort = Sign.intern_sort thy; |
|
491 |
in Syntax.decode_term get_sort map_const map_free map_type map_sort end; |
|
492 |
||
493 |
||
5819 | 494 |
(* read terms *) |
495 |
||
10554 | 496 |
local |
497 |
||
15531 | 498 |
fun append_env e1 e2 x = (case e2 x of NONE => e1 x | some => some); |
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
499 |
|
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
500 |
fun gen_read' read app pattern schematic |
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
501 |
ctxt internal more_types more_sorts more_used s = |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13629
diff
changeset
|
502 |
let |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
503 |
val types = append_env (Variable.def_type ctxt pattern) more_types; |
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
504 |
val sorts = append_env (Variable.def_sort ctxt) more_sorts; |
20163 | 505 |
val used = fold Name.declare more_used (Variable.names_of ctxt); |
14174
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
ballarin
parents:
13629
diff
changeset
|
506 |
in |
22763 | 507 |
(read (pp ctxt) (syn_of ctxt) ctxt (types, sorts, used) |
508 |
(legacy_intern_skolem ctxt internal types) s |
|
18678 | 509 |
handle TERM (msg, _) => error msg) |
19371 | 510 |
|> app (certify_consts ctxt) |
19001 | 511 |
|> app (if pattern then I else expand_binds ctxt schematic) |
22712 | 512 |
|> app (if pattern orelse schematic then I else reject_tvars) |
18678 | 513 |
|> app (if pattern then prepare_dummies else reject_dummies) |
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
514 |
end; |
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
515 |
|
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
516 |
fun gen_read read app pattern schematic ctxt = |
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
517 |
gen_read' read app pattern schematic ctxt (K false) (K NONE) (K NONE) []; |
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
518 |
|
10554 | 519 |
in |
520 |
||
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
521 |
val read_termTs_schematic = gen_read' (read_def_termTs false) (apfst o map) false true; |
8096 | 522 |
|
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
523 |
fun read_term_pats T ctxt = |
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
524 |
#1 o gen_read (read_def_termTs false) (apfst o map) true false ctxt o map (rpair T); |
8096 | 525 |
val read_prop_pats = read_term_pats propT; |
526 |
||
22712 | 527 |
fun read_prop_legacy ctxt = |
528 |
gen_read' (read_prop_thy true) I false false ctxt (K true) (K NONE) (K NONE) []; |
|
14720
ceff6d4fb836
cleanup up read functions, include liberal versions;
wenzelm
parents:
14707
diff
changeset
|
529 |
|
19143 | 530 |
val read_term = gen_read (read_term_thy true) I false false; |
531 |
val read_prop = gen_read (read_prop_thy true) I false false; |
|
532 |
val read_prop_schematic = gen_read (read_prop_thy true) I false true; |
|
22773 | 533 |
val read_termTs = gen_read (read_terms_thy true) map false false; |
534 |
fun read_terms ctxt = read_termTs ctxt o map (rpair dummyT); |
|
19143 | 535 |
fun read_props schematic = gen_read (read_props_thy true) map false schematic; |
5819 | 536 |
|
10554 | 537 |
end; |
538 |
||
5819 | 539 |
|
540 |
(* certify terms *) |
|
541 |
||
10554 | 542 |
local |
543 |
||
18971 | 544 |
fun gen_cert prop pattern schematic ctxt t = t |
19371 | 545 |
|> certify_consts ctxt |
19001 | 546 |
|> (if pattern then I else expand_binds ctxt schematic) |
19371 | 547 |
|> (fn t' => #1 (Sign.certify' false prop (pp ctxt) (consts_of ctxt) (theory_of ctxt) t') |
18678 | 548 |
handle TYPE (msg, _, _) => error msg |
549 |
| TERM (msg, _) => error msg); |
|
16501 | 550 |
|
10554 | 551 |
in |
8096 | 552 |
|
18971 | 553 |
val cert_term = gen_cert false false false; |
22773 | 554 |
val cert_terms = map o cert_term; |
18971 | 555 |
val cert_prop = gen_cert true false false; |
556 |
val cert_props = map oo gen_cert true false; |
|
10554 | 557 |
|
18971 | 558 |
fun cert_term_pats _ = map o gen_cert false true false; |
559 |
val cert_prop_pats = map o gen_cert true true false; |
|
10554 | 560 |
|
561 |
end; |
|
5819 | 562 |
|
563 |
||
22701 | 564 |
(* type inference *) |
565 |
||
566 |
local |
|
567 |
||
22728 | 568 |
fun gen_infer_types pattern ctxt ts = |
22701 | 569 |
TypeInfer.infer_types (pp ctxt) (tsig_of ctxt) (try (Consts.the_constraint (consts_of ctxt))) |
22728 | 570 |
(Variable.def_type ctxt pattern) (Variable.names_of ctxt) (not pattern) (map (rpair dummyT) ts) |
571 |
|> #1 |> map #1; |
|
22701 | 572 |
|
573 |
in |
|
574 |
||
575 |
val infer_types = gen_infer_types false; |
|
22728 | 576 |
val infer_types_pats = gen_infer_types true; |
22701 | 577 |
|
578 |
end; |
|
579 |
||
580 |
||
18770 | 581 |
(* inferred types of parameters *) |
582 |
||
583 |
fun infer_type ctxt x = |
|
22728 | 584 |
Term.fastype_of (singleton (infer_types ctxt) (Free (x, dummyT))); |
18770 | 585 |
|
586 |
fun inferred_param x ctxt = |
|
18619 | 587 |
let val T = infer_type ctxt x |
20163 | 588 |
in ((x, T), ctxt |> Variable.declare_term (Free (x, T))) end; |
18619 | 589 |
|
18770 | 590 |
fun inferred_fixes ctxt = |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
591 |
fold_map inferred_param (rev (map #2 (Variable.fixes_of ctxt))) ctxt; |
5819 | 592 |
|
593 |
||
15703 | 594 |
(* type and constant names *) |
595 |
||
596 |
fun read_tyname ctxt c = |
|
20163 | 597 |
if Syntax.is_tid c then |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
598 |
TFree (c, the_default (Sign.defaultS (theory_of ctxt)) (Variable.def_sort ctxt (c, ~1))) |
16458 | 599 |
else Sign.read_tyname (theory_of ctxt) c; |
15703 | 600 |
|
601 |
fun read_const ctxt c = |
|
602 |
(case lookup_skolem ctxt c of |
|
21208
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
603 |
SOME x => Free (x, infer_type ctxt x) |
18971 | 604 |
| NONE => Consts.read_const (consts_of ctxt) c); |
15703 | 605 |
|
606 |
||
9553 | 607 |
|
21610 | 608 |
(** export results **) |
21531 | 609 |
|
20310 | 610 |
fun common_export is_goal inner outer = |
611 |
map (Assumption.export is_goal inner outer) #> |
|
612 |
Variable.export inner outer; |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
613 |
|
20310 | 614 |
val goal_export = common_export true; |
615 |
val export = common_export false; |
|
12704 | 616 |
|
21531 | 617 |
fun export_morphism inner outer = |
618 |
Assumption.export_morphism inner outer $> |
|
619 |
Variable.export_morphism inner outer; |
|
620 |
||
621 |
||
15758
07e382399a96
binds/thms: do not store options, but delete from table;
wenzelm
parents:
15750
diff
changeset
|
622 |
|
5819 | 623 |
(** bindings **) |
624 |
||
8096 | 625 |
(* simult_matches *) |
626 |
||
19867 | 627 |
fun simult_matches ctxt (t, pats) = |
628 |
(case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of |
|
629 |
NONE => error "Pattern match failed!" |
|
630 |
| SOME (env, _) => map (apsnd snd) (Envir.alist_of env)); |
|
8096 | 631 |
|
632 |
||
633 |
(* add_binds(_i) *) |
|
5819 | 634 |
|
7925 | 635 |
local |
636 |
||
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
637 |
fun gen_bind prep (xi as (x, _), raw_t) ctxt = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
638 |
ctxt |> Variable.add_binds [(xi, Option.map (prep ctxt) raw_t)]; |
5819 | 639 |
|
10810 | 640 |
in |
641 |
||
20330 | 642 |
fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b |
10554 | 643 |
| drop_schematic b = b; |
644 |
||
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
645 |
val add_binds = fold (gen_bind read_term); |
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
646 |
val add_binds_i = fold (gen_bind cert_term); |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
647 |
|
16458 | 648 |
fun auto_bind f ts ctxt = ctxt |> add_binds_i (map drop_schematic (f (theory_of ctxt) ts)); |
12147 | 649 |
val auto_bind_goal = auto_bind AutoBind.goal; |
650 |
val auto_bind_facts = auto_bind AutoBind.facts; |
|
7925 | 651 |
|
652 |
end; |
|
5819 | 653 |
|
654 |
||
8096 | 655 |
(* match_bind(_i) *) |
5819 | 656 |
|
8096 | 657 |
local |
658 |
||
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
659 |
fun prep_bind prep_pats (raw_pats, t) ctxt = |
5819 | 660 |
let |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
661 |
val ctxt' = Variable.declare_term t ctxt; |
19585 | 662 |
val pats = prep_pats (Term.fastype_of t) ctxt' raw_pats; |
663 |
val binds = simult_matches ctxt' (t, pats); |
|
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
664 |
in (binds, ctxt') end; |
7670 | 665 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
666 |
fun gen_binds prep_terms prep_pats gen raw_binds ctxt = |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
667 |
let |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
668 |
val ts = prep_terms ctxt (map snd raw_binds); |
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
669 |
val (binds, ctxt') = |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19422
diff
changeset
|
670 |
apfst flat (fold_map (prep_bind prep_pats) (map fst raw_binds ~~ ts) ctxt); |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
671 |
val binds' = |
19916
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
wenzelm
parents:
19897
diff
changeset
|
672 |
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
|
673 |
else binds; |
15531 | 674 |
val binds'' = map (apsnd SOME) binds'; |
18310 | 675 |
val ctxt'' = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
676 |
tap (Variable.warn_extra_tfrees ctxt) |
18310 | 677 |
(if gen then |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
678 |
ctxt (*sic!*) |> fold Variable.declare_term (map #2 binds') |> add_binds_i binds'' |
18310 | 679 |
else ctxt' |> add_binds_i binds''); |
680 |
in (ts, ctxt'') end; |
|
8096 | 681 |
|
682 |
in |
|
5935 | 683 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
684 |
val match_bind = gen_binds read_terms read_term_pats; |
22773 | 685 |
val match_bind_i = gen_binds cert_terms cert_term_pats; |
8096 | 686 |
|
687 |
end; |
|
5935 | 688 |
|
689 |
||
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
690 |
(* propositions with patterns *) |
5935 | 691 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
692 |
local |
8096 | 693 |
|
10554 | 694 |
fun prep_propp schematic prep_props prep_pats (context, args) = |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
695 |
let |
19585 | 696 |
fun prep (_, raw_pats) (ctxt, prop :: props) = |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
697 |
let val ctxt' = Variable.declare_term prop ctxt |
19585 | 698 |
in ((prop, prep_pats ctxt' raw_pats), (ctxt', props)) end |
17860
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
699 |
| prep _ _ = sys_error "prep_propp"; |
b4cf247ea0d2
note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents:
17756
diff
changeset
|
700 |
val (propp, (context', _)) = (fold_map o fold_map) prep args |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19422
diff
changeset
|
701 |
(context, prep_props schematic context (maps (map fst) args)); |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
702 |
in (context', propp) end; |
5935 | 703 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
704 |
fun gen_bind_propp prepp (ctxt, raw_args) = |
8096 | 705 |
let |
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
706 |
val (ctxt', args) = prepp (ctxt, raw_args); |
19585 | 707 |
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
|
708 |
val propss = map (map #1) args; |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
709 |
|
10554 | 710 |
(*generalize result: context evaluated now, binds added later*) |
19916
3bbb9cc5d4f1
export: simultaneous facts, refer to Variable.export;
wenzelm
parents:
19897
diff
changeset
|
711 |
val gen = Variable.exportT_terms ctxt' ctxt; |
15531 | 712 |
fun gen_binds c = c |> add_binds_i (map #1 binds ~~ map SOME (gen (map #2 binds))); |
713 |
in (ctxt' |> add_binds_i (map (apsnd SOME) binds), (propss, gen_binds)) end; |
|
8096 | 714 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
715 |
in |
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
716 |
|
11925 | 717 |
val read_propp = prep_propp false read_props read_prop_pats; |
718 |
val cert_propp = prep_propp false cert_props cert_prop_pats; |
|
10554 | 719 |
val read_propp_schematic = prep_propp true read_props read_prop_pats; |
720 |
val cert_propp_schematic = prep_propp true cert_props cert_prop_pats; |
|
721 |
||
11925 | 722 |
val bind_propp = gen_bind_propp read_propp; |
723 |
val bind_propp_i = gen_bind_propp cert_propp; |
|
724 |
val bind_propp_schematic = gen_bind_propp read_propp_schematic; |
|
10554 | 725 |
val bind_propp_schematic_i = gen_bind_propp cert_propp_schematic; |
6789 | 726 |
|
10465
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
727 |
end; |
4aa6f8b5cdc4
added read_terms, read_props (simulataneous type-inference);
wenzelm
parents:
10381
diff
changeset
|
728 |
|
6789 | 729 |
|
5819 | 730 |
|
731 |
(** theorems **) |
|
732 |
||
18042 | 733 |
(* fact_tac *) |
734 |
||
18122 | 735 |
fun comp_incr_tac [] _ st = no_tac st |
736 |
| comp_incr_tac (th :: ths) i st = |
|
737 |
(Goal.compose_hhf_tac (Drule.incr_indexes st th) i APPEND comp_incr_tac ths i) st; |
|
18042 | 738 |
|
21687 | 739 |
fun fact_tac facts = Goal.norm_hhf_tac THEN' comp_incr_tac facts; |
18122 | 740 |
|
741 |
fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => |
|
18042 | 742 |
let |
19001 | 743 |
val index = fact_index_of ctxt; |
18042 | 744 |
val facts = FactIndex.could_unify index (Term.strip_all_body goal); |
745 |
in fact_tac facts i end); |
|
746 |
||
747 |
||
6091 | 748 |
(* get_thm(s) *) |
5819 | 749 |
|
18042 | 750 |
fun retrieve_thms _ pick ctxt (Fact s) = |
16501 | 751 |
let |
20049 | 752 |
val th = Goal.prove ctxt [] [] (read_prop ctxt s) (K (ALLGOALS (some_fact_tac ctxt))) |
18678 | 753 |
handle ERROR msg => cat_error msg "Failed to retrieve literal fact."; |
18042 | 754 |
in pick "" [th] end |
755 |
| retrieve_thms from_thy pick ctxt xthmref = |
|
756 |
let |
|
757 |
val thy = theory_of ctxt; |
|
19001 | 758 |
val (space, tab) = #1 (thms_of ctxt); |
16501 | 759 |
val thmref = PureThy.map_name_of_thmref (NameSpace.intern space) xthmref; |
760 |
val name = PureThy.name_of_thmref thmref; |
|
24012 | 761 |
val thms = |
762 |
if name = "" then [Thm.transfer thy Drule.dummy_thm] |
|
763 |
else |
|
764 |
(case Symtab.lookup tab name of |
|
765 |
SOME ths => map (Thm.transfer thy) (PureThy.select_thm thmref ths) |
|
766 |
| NONE => from_thy thy xthmref); |
|
767 |
in pick name thms end; |
|
5819 | 768 |
|
9566 | 769 |
val get_thm = retrieve_thms PureThy.get_thms PureThy.single_thm; |
770 |
val get_thm_closure = retrieve_thms PureThy.get_thms_closure PureThy.single_thm; |
|
771 |
val get_thms = retrieve_thms PureThy.get_thms (K I); |
|
772 |
val get_thms_closure = retrieve_thms PureThy.get_thms_closure (K I); |
|
5819 | 773 |
|
774 |
||
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
775 |
(* valid_thms *) |
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
776 |
|
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
777 |
fun valid_thms ctxt (name, ths) = |
18678 | 778 |
(case try (fn () => get_thms ctxt (Name name)) () of |
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
779 |
NONE => false |
16147 | 780 |
| SOME ths' => Thm.eq_thms (ths, ths')); |
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
781 |
|
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
782 |
|
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
783 |
(* lthms_containing *) |
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
784 |
|
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
785 |
fun lthms_containing ctxt spec = |
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
786 |
FactIndex.find (fact_index_of ctxt) spec |
21569 | 787 |
|> map (fn (name, ths) => |
788 |
if valid_thms ctxt (name, ths) then (name, ths) |
|
789 |
else (NameSpace.hidden (if name = "" then "unnamed" else name), ths)); |
|
16031
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
790 |
|
fbf3471214d6
moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents:
15979
diff
changeset
|
791 |
|
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
792 |
(* name space operations *) |
12309
03e9287be350
name space for local thms (export cond_extern, qualified);
wenzelm
parents:
12291
diff
changeset
|
793 |
|
22352 | 794 |
val add_path = map_naming o NameSpace.add_path; |
19062
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents:
19033
diff
changeset
|
795 |
val no_base_names = map_naming NameSpace.no_base_names; |
16147 | 796 |
val qualified_names = map_naming NameSpace.qualified_names; |
19062
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents:
19033
diff
changeset
|
797 |
val sticky_prefix = map_naming o NameSpace.sticky_prefix; |
0fd52e819c24
replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents:
19033
diff
changeset
|
798 |
val restore_naming = map_naming o K o naming_of; |
21728 | 799 |
val reset_naming = map_naming (K local_naming); |
12309
03e9287be350
name space for local thms (export cond_extern, qualified);
wenzelm
parents:
12291
diff
changeset
|
800 |
|
19001 | 801 |
fun hide_thms fully names = map_thms (fn ((space, tab), index) => |
802 |
((fold (NameSpace.hide fully) names space, tab), index)); |
|
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
803 |
|
12309
03e9287be350
name space for local thms (export cond_extern, qualified);
wenzelm
parents:
12291
diff
changeset
|
804 |
|
17360
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents:
17221
diff
changeset
|
805 |
(* put_thms *) |
5819 | 806 |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
807 |
fun update_thms _ ("", NONE) ctxt = ctxt |
21648 | 808 |
| update_thms opts ("", SOME ths) ctxt = ctxt |> map_thms (fn (facts, index) => |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
809 |
let |
21648 | 810 |
val index' = uncurry FactIndex.add_local opts ("", ths) index; |
19001 | 811 |
in (facts, index') end) |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
812 |
| update_thms _ (bname, NONE) ctxt = ctxt |> map_thms (fn ((space, tab), index) => |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
813 |
let |
19387 | 814 |
val name = full_name ctxt bname; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
815 |
val tab' = Symtab.delete_safe name tab; |
19001 | 816 |
in ((space, tab'), index) end) |
21648 | 817 |
| update_thms opts (bname, SOME ths) ctxt = ctxt |> map_thms (fn ((space, tab), index) => |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
818 |
let |
19387 | 819 |
val name = full_name ctxt bname; |
19001 | 820 |
val space' = NameSpace.declare (naming_of ctxt) name space; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
821 |
val tab' = Symtab.update (name, ths) tab; |
21648 | 822 |
val index' = uncurry FactIndex.add_local opts (name, ths) index; |
19001 | 823 |
in ((space', tab'), index') end); |
5819 | 824 |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
825 |
fun put_thms thms ctxt = |
21648 | 826 |
ctxt |> map_naming (K local_naming) |> update_thms (true, false) thms |> restore_naming ctxt; |
19079
9a7678a0736d
added put_thms_internal: local_naming, no fact index;
wenzelm
parents:
19062
diff
changeset
|
827 |
|
7606 | 828 |
|
14564 | 829 |
(* note_thmss *) |
5819 | 830 |
|
12711 | 831 |
local |
16147 | 832 |
|
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
833 |
fun gen_note_thmss get k = fold_map (fn ((bname, more_attrs), raw_facts) => fn ctxt => |
5819 | 834 |
let |
21643
bdf3e74727df
note_thmss: added kind tag and non-official name;
wenzelm
parents:
21622
diff
changeset
|
835 |
val name = full_name ctxt bname; |
bdf3e74727df
note_thmss: added kind tag and non-official name;
wenzelm
parents:
21622
diff
changeset
|
836 |
val facts = PureThy.name_thmss false name (map (apfst (get ctxt)) raw_facts); |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
837 |
fun app (th, attrs) x = |
21643
bdf3e74727df
note_thmss: added kind tag and non-official name;
wenzelm
parents:
21622
diff
changeset
|
838 |
swap (foldl_map (Thm.proof_attributes (attrs @ more_attrs @ [PureThy.kind k])) (x, th)); |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
839 |
val (res, ctxt') = fold_map app facts ctxt; |
21643
bdf3e74727df
note_thmss: added kind tag and non-official name;
wenzelm
parents:
21622
diff
changeset
|
840 |
val thms = PureThy.name_thms false false name (flat res); |
21648 | 841 |
in ((bname, thms), ctxt' |> update_thms (is_stmt ctxt, true) (bname, SOME thms)) end); |
12711 | 842 |
|
843 |
in |
|
844 |
||
21622 | 845 |
fun note_thmss k = gen_note_thmss get_thms k; |
846 |
fun note_thmss_i k = gen_note_thmss (K I) k; |
|
15696 | 847 |
|
12711 | 848 |
end; |
9196 | 849 |
|
5819 | 850 |
|
851 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
852 |
(** parameters **) |
17360
fa1f262dbc4e
added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents:
17221
diff
changeset
|
853 |
|
8096 | 854 |
(* variables *) |
855 |
||
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
856 |
fun declare_var (x, opt_T, mx) ctxt = |
22701 | 857 |
let val T = (case opt_T of SOME T => T | NONE => Syntax.mixfixT mx) |
20163 | 858 |
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
|
859 |
|
10381 | 860 |
local |
861 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
862 |
fun prep_vars prep_typ internal legacy = |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
863 |
fold_map (fn (raw_x, raw_T, raw_mx) => fn ctxt => |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
864 |
let |
19371 | 865 |
val (x, mx) = Syntax.const_mixfix raw_x raw_mx; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
866 |
val _ = |
18678 | 867 |
if not legacy andalso not (Syntax.is_identifier (no_skolem internal x)) then |
868 |
error ("Illegal variable name: " ^ quote x) |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
869 |
else (); |
12504 | 870 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
871 |
fun cond_tvars T = |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
872 |
if internal then T |
18678 | 873 |
else Type.no_tvars T handle TYPE (msg, _, _) => error msg; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
874 |
val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
875 |
val var = (x, opt_T, mx); |
19001 | 876 |
in (var, ctxt |> declare_var var |> #2) end); |
8096 | 877 |
|
10381 | 878 |
in |
879 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
880 |
val read_vars = prep_vars read_typ false false; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
881 |
val cert_vars = prep_vars cert_typ true false; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
882 |
val read_vars_legacy = prep_vars read_typ true true; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
883 |
val cert_vars_legacy = prep_vars cert_typ true true; |
8096 | 884 |
|
10381 | 885 |
end; |
886 |
||
8096 | 887 |
|
19681 | 888 |
(* authentic constants *) |
19663 | 889 |
|
21208
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
890 |
fun const_syntax ctxt (Free (x, T), mx) = SOME (true, (x, T, mx)) |
21667
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
891 |
| const_syntax ctxt (Const (c, _), mx) = |
ce813b82c88b
add_notation: permissive about undeclared consts;
wenzelm
parents:
21648
diff
changeset
|
892 |
Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx)) |
21208
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
893 |
| const_syntax _ _ = NONE; |
62ccdaf9369a
replaced const_syntax by notation, which operates on terms;
wenzelm
parents:
21183
diff
changeset
|
894 |
|
21772 | 895 |
fun context_const_ast_tr ctxt [Syntax.Variable c] = |
19681 | 896 |
let |
21772 | 897 |
val consts = consts_of ctxt; |
19681 | 898 |
val c' = Consts.intern consts c; |
21728 | 899 |
val _ = Consts.the_constraint consts c' handle TYPE (msg, _, _) => error msg; |
19681 | 900 |
in Syntax.Constant (Syntax.constN ^ c') end |
901 |
| context_const_ast_tr _ asts = raise Syntax.AST ("const_ast_tr", asts); |
|
902 |
||
903 |
val _ = Context.add_setup |
|
904 |
(Sign.add_syntax |
|
905 |
[("_context_const", "id => 'a", Delimfix "CONST _"), |
|
906 |
("_context_const", "longid => 'a", Delimfix "CONST _")] #> |
|
907 |
Sign.add_advanced_trfuns ([("_context_const", context_const_ast_tr)], [], [], [])); |
|
908 |
||
19663 | 909 |
|
21744 | 910 |
(* notation *) |
911 |
||
912 |
fun add_notation mode args ctxt = |
|
913 |
ctxt |> map_syntax |
|
914 |
(LocalSyntax.add_modesyntax (theory_of ctxt) mode (map_filter (const_syntax ctxt) args)); |
|
915 |
||
21803 | 916 |
fun target_notation mode args phi = |
917 |
let val args' = filter (fn (t, _) => Term.equiv_types (t, Morphism.term phi t)) args; |
|
21744 | 918 |
in Context.mapping (Sign.add_notation mode args') (add_notation mode args') end; |
919 |
||
920 |
||
19001 | 921 |
(* abbreviations *) |
922 |
||
21728 | 923 |
val set_expand_abbrevs = map_consts o apsnd o Consts.set_expand; |
924 |
fun read_term_abbrev ctxt = read_term (set_expand_abbrevs false ctxt); |
|
19371 | 925 |
|
21704 | 926 |
fun add_abbrev mode (c, raw_t) ctxt = |
19001 | 927 |
let |
21728 | 928 |
val t0 = cert_term (ctxt |> set_expand_abbrevs false) raw_t |
21681 | 929 |
handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote c); |
20008
8d9d770e1f06
add_abbrevs/polymorphic: Variable.exportT_terms avoids over-generalization;
wenzelm
parents:
19916
diff
changeset
|
930 |
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
|
931 |
val ((lhs, rhs), consts') = consts_of ctxt |
21728 | 932 |
|> Consts.abbreviate (pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode (c, t); |
19001 | 933 |
in |
934 |
ctxt |
|
21681 | 935 |
|> map_consts (apsnd (K consts')) |
21803 | 936 |
|> Variable.declare_term rhs |
937 |
|> pair (lhs, rhs) |
|
21704 | 938 |
end; |
19001 | 939 |
|
940 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
941 |
(* fixes *) |
5819 | 942 |
|
8096 | 943 |
local |
944 |
||
19001 | 945 |
fun prep_mixfix (x, T, mx) = |
19019 | 946 |
if mx <> NoSyn andalso mx <> Structure andalso |
20086
94ca946fb689
adapted to more efficient Name/Variable implementation;
wenzelm
parents:
20049
diff
changeset
|
947 |
(can Name.dest_internal x orelse can Name.dest_skolem x) then |
19001 | 948 |
error ("Illegal mixfix syntax for internal/skolem constant " ^ quote x) |
949 |
else (true, (x, T, mx)); |
|
950 |
||
18844 | 951 |
fun gen_fixes prep raw_vars ctxt = |
8096 | 952 |
let |
22712 | 953 |
val (vars, _) = prep raw_vars ctxt; |
954 |
val (xs', ctxt') = Variable.add_fixes (map #1 vars) ctxt; |
|
8096 | 955 |
in |
22712 | 956 |
ctxt' |
19001 | 957 |
|> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars) |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
958 |
|-> (map_syntax o LocalSyntax.add_syntax (theory_of ctxt) o map prep_mixfix) |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
959 |
|> pair xs' |
8096 | 960 |
end; |
5819 | 961 |
|
8096 | 962 |
in |
7679 | 963 |
|
18844 | 964 |
val add_fixes = gen_fixes read_vars; |
965 |
val add_fixes_i = gen_fixes cert_vars; |
|
966 |
val add_fixes_legacy = gen_fixes cert_vars_legacy; |
|
8096 | 967 |
|
968 |
end; |
|
5819 | 969 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
970 |
|
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
971 |
(* fixes vs. frees *) |
12016 | 972 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
973 |
fun auto_fixes (arg as (ctxt, (propss, x))) = |
21370
d9dd7b4e5e69
replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents:
21269
diff
changeset
|
974 |
((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
|
975 |
|
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
976 |
fun bind_fixes xs ctxt = |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
977 |
let |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
978 |
val (_, ctxt') = ctxt |> add_fixes_i (map (fn x => (x, NONE, NoSyn)) xs); |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
979 |
fun bind (t as Free (x, T)) = |
18340 | 980 |
if member (op =) xs x then |
15531 | 981 |
(case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t) |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
982 |
else t |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
983 |
| bind (t $ u) = bind t $ bind u |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
984 |
| bind (Abs (x, T, t)) = Abs (x, T, bind t) |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
985 |
| bind a = a; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
986 |
in (bind, ctxt') end; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
987 |
|
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
988 |
|
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
989 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
990 |
(** assumptions **) |
18187 | 991 |
|
20209 | 992 |
local |
993 |
||
994 |
fun gen_assms prepp exp args ctxt = |
|
995 |
let |
|
20234
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
996 |
val cert = Thm.cterm_of (theory_of ctxt); |
20209 | 997 |
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
|
998 |
val _ = Variable.warn_extra_tfrees ctxt ctxt1; |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
999 |
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
|
1000 |
in |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1001 |
ctxt2 |
7e0693474bcd
added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents:
20209
diff
changeset
|
1002 |
|> auto_bind_facts (flat propss) |
21443
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
1003 |
|> put_thms (AutoBind.premsN, SOME (Assumption.prems_of ctxt2)) |
cc5095d57da4
added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents:
21370
diff
changeset
|
1004 |
|> note_thmss_i Thm.assumptionK (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
|
1005 |
end; |
20209 | 1006 |
|
1007 |
in |
|
1008 |
||
1009 |
val add_assms = gen_assms (apsnd #1 o bind_propp); |
|
1010 |
val add_assms_i = gen_assms (apsnd #1 o bind_propp_i); |
|
1011 |
||
1012 |
end; |
|
1013 |
||
1014 |
||
5819 | 1015 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1016 |
(** cases **) |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1017 |
|
16147 | 1018 |
local |
1019 |
||
16668 | 1020 |
fun rem_case name = remove (fn (x: string, (y, _)) => x = y) name; |
16147 | 1021 |
|
18476 | 1022 |
fun add_case _ ("", _) cases = cases |
1023 |
| add_case _ (name, NONE) cases = rem_case name cases |
|
1024 |
| add_case is_proper (name, SOME c) cases = (name, (c, is_proper)) :: rem_case name cases; |
|
16147 | 1025 |
|
18678 | 1026 |
fun prep_case name fxs c = |
18609 | 1027 |
let |
1028 |
fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys |
|
1029 |
| replace [] ys = ys |
|
18678 | 1030 |
| replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name); |
18609 | 1031 |
val RuleCases.Case {fixes, assumes, binds, cases} = c; |
1032 |
val fixes' = replace fxs fixes; |
|
1033 |
val binds' = map drop_schematic binds; |
|
1034 |
in |
|
1035 |
if null (fold (Term.add_tvarsT o snd) fixes []) andalso |
|
1036 |
null (fold (fold Term.add_vars o snd) assumes []) then |
|
1037 |
RuleCases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases} |
|
18678 | 1038 |
else error ("Illegal schematic variable(s) in case " ^ quote name) |
18609 | 1039 |
end; |
1040 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1041 |
fun fix (x, T) ctxt = |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1042 |
let |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1043 |
val (bind, ctxt') = bind_fixes [x] ctxt; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1044 |
val t = bind (Free (x, T)); |
20163 | 1045 |
in (t, ctxt' |> Variable.declare_constraints t) end; |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1046 |
|
16147 | 1047 |
in |
1048 |
||
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1049 |
fun add_cases is_proper = map_cases o fold (add_case is_proper); |
18609 | 1050 |
|
1051 |
fun case_result c ctxt = |
|
1052 |
let |
|
1053 |
val RuleCases.Case {fixes, ...} = c; |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1054 |
val (ts, ctxt') = ctxt |> fold_map fix fixes; |
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1055 |
val RuleCases.Case {assumes, binds, cases, ...} = RuleCases.apply ts c; |
18609 | 1056 |
in |
1057 |
ctxt' |
|
18699
f3bfe81b6e58
case_result: drop_schematic, i.e. be permissive about illegal binds;
wenzelm
parents:
18678
diff
changeset
|
1058 |
|> add_binds_i (map drop_schematic binds) |
18609 | 1059 |
|> add_cases true (map (apsnd SOME) cases) |
1060 |
|> pair (assumes, (binds, cases)) |
|
1061 |
end; |
|
1062 |
||
1063 |
val apply_case = apfst fst oo case_result; |
|
1064 |
||
16540 | 1065 |
fun get_case ctxt name xs = |
17184 | 1066 |
(case AList.lookup (op =) (cases_of ctxt) name of |
18678 | 1067 |
NONE => error ("Unknown case: " ^ quote name) |
1068 |
| SOME (c, _) => prep_case name xs c); |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1069 |
|
16147 | 1070 |
end; |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1071 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1072 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1073 |
|
10810 | 1074 |
(** print context information **) |
1075 |
||
20310 | 1076 |
val debug = ref false; |
1077 |
||
10810 | 1078 |
val verbose = ref false; |
1079 |
fun verb f x = if ! verbose then f (x ()) else []; |
|
1080 |
||
1081 |
fun setmp_verbose f x = Library.setmp verbose true f x; |
|
1082 |
||
1083 |
||
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1084 |
(* local syntax *) |
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1085 |
|
12093 | 1086 |
val print_syntax = Syntax.print_syntax o syn_of; |
12072
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1087 |
|
4281198fb8cd
local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents:
12066
diff
changeset
|
1088 |
|
21728 | 1089 |
(* abbreviations *) |
18971 | 1090 |
|
21728 | 1091 |
fun pretty_abbrevs show_globals ctxt = |
18971 | 1092 |
let |
19033
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
1093 |
val ((_, globals), (space, consts)) = |
24e251657e56
consts: maintain thy version for efficient transfer;
wenzelm
parents:
19019
diff
changeset
|
1094 |
pairself (#constants o Consts.dest) (#consts (rep_context ctxt)); |
21803 | 1095 |
fun add_abbr (_, (_, NONE)) = I |
1096 |
| add_abbr (c, (T, SOME (t, _))) = |
|
21728 | 1097 |
if not show_globals andalso Symtab.defined globals c then I |
1098 |
else cons (c, Logic.mk_equals (Const (c, T), t)); |
|
21803 | 1099 |
val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts [])); |
18971 | 1100 |
in |
1101 |
if null abbrevs andalso not (! verbose) then [] |
|
21728 | 1102 |
else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)] |
18971 | 1103 |
end; |
1104 |
||
21728 | 1105 |
val print_abbrevs = Pretty.writeln o Pretty.chunks o pretty_abbrevs true; |
1106 |
||
18971 | 1107 |
|
10810 | 1108 |
(* term bindings *) |
1109 |
||
16540 | 1110 |
fun pretty_binds ctxt = |
10810 | 1111 |
let |
19897
fe661eb3b0e7
ProofContext: moved variable operations to struct Variable;
wenzelm
parents:
19882
diff
changeset
|
1112 |
val binds = Variable.binds_of ctxt; |
21728 | 1113 |
fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t)); |
10810 | 1114 |
in |
15758
07e382399a96
binds/thms: do not store options, but delete from table;
wenzelm
parents:
15750
diff
changeset
|
1115 |
if Vartab.is_empty binds andalso not (! verbose) then [] |
07e382399a96
binds/thms: do not store options, but delete from table;
wenzelm
parents:
15750
diff
changeset
|
1116 |
else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))] |
10810 | 1117 |
end; |
1118 |
||
1119 |
val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds; |
|
1120 |
||
1121 |
||
1122 |
(* local theorems *) |
|
1123 |
||
16540 | 1124 |
fun pretty_lthms ctxt = |
20012 | 1125 |
let |
1126 |
val props = FactIndex.props (fact_index_of ctxt); |
|
1127 |
val facts = |
|
1128 |
(if null props then I else cons ("unnamed", props)) |
|
1129 |
(NameSpace.extern_table (#1 (thms_of ctxt))); |
|
1130 |
in |
|
1131 |
if null facts andalso not (! verbose) then [] |
|
1132 |
else [Pretty.big_list "facts:" (map (pretty_fact ctxt) facts)] |
|
1133 |
end; |
|
10810 | 1134 |
|
12057 | 1135 |
val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms; |
10810 | 1136 |
|
1137 |
||
1138 |
(* local contexts *) |
|
1139 |
||
16540 | 1140 |
fun pretty_cases ctxt = |
10810 | 1141 |
let |
12057 | 1142 |
val prt_term = pretty_term ctxt; |
1143 |
||
10810 | 1144 |
fun prt_let (xi, t) = Pretty.block |
10818 | 1145 |
[Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1, |
10810 | 1146 |
Pretty.quote (prt_term t)]; |
1147 |
||
13425
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1148 |
fun prt_asm (a, ts) = Pretty.block (Pretty.breaks |
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
wenzelm
parents:
13415
diff
changeset
|
1149 |
((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
|
1150 |
|
10810 | 1151 |
fun prt_sect _ _ _ [] = [] |
1152 |
| prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s :: |
|
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19422
diff
changeset
|
1153 |
flat (Library.separate sep (map (Library.single o prt) xs))))]; |
10810 | 1154 |
|
18609 | 1155 |
fun prt_case (name, (fixes, (asms, (lets, cs)))) = Pretty.block (Pretty.fbreaks |
10810 | 1156 |
(Pretty.str (name ^ ":") :: |
11915 | 1157 |
prt_sect "fix" [] (Pretty.str o fst) fixes @ |
10810 | 1158 |
prt_sect "let" [Pretty.str "and"] prt_let |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19422
diff
changeset
|
1159 |
(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
|
1160 |
(if forall (null o #2) asms then [] |
18609 | 1161 |
else prt_sect "assume" [Pretty.str "and"] prt_asm asms) @ |
1162 |
prt_sect "subcases:" [] (Pretty.str o fst) cs)); |
|
16540 | 1163 |
|
18476 | 1164 |
fun add_case (_, (_, false)) = I |
18609 | 1165 |
| add_case (name, (c as RuleCases.Case {fixes, ...}, true)) = |
1166 |
cons (name, (fixes, #1 (case_result c ctxt))); |
|
18476 | 1167 |
val cases = fold add_case (cases_of ctxt) []; |
10810 | 1168 |
in |
1169 |
if null cases andalso not (! verbose) then [] |
|
18476 | 1170 |
else [Pretty.big_list "cases:" (map prt_case cases)] |
10810 | 1171 |
end; |
1172 |
||
1173 |
val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases; |
|
1174 |
||
1175 |
||
12057 | 1176 |
(* core context *) |
10810 | 1177 |
|
20367 | 1178 |
val prems_limit = ref ~1; |
10810 | 1179 |
|
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1180 |
fun pretty_ctxt ctxt = |
20310 | 1181 |
if ! prems_limit < 0 andalso not (! debug) then [] |
1182 |
else |
|
1183 |
let |
|
1184 |
val prt_term = pretty_term ctxt; |
|
12057 | 1185 |
|
20310 | 1186 |
(*structures*) |
1187 |
val structs = LocalSyntax.structs_of (syntax_of ctxt); |
|
1188 |
val prt_structs = if null structs then [] |
|
1189 |
else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 :: |
|
1190 |
Pretty.commas (map Pretty.str structs))]; |
|
12093 | 1191 |
|
20310 | 1192 |
(*fixes*) |
1193 |
fun prt_fix (x, x') = |
|
1194 |
if x = x' then Pretty.str x |
|
1195 |
else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')]; |
|
1196 |
val fixes = |
|
1197 |
rev (filter_out ((can Name.dest_internal orf member (op =) structs) o #1) |
|
1198 |
(Variable.fixes_of ctxt)); |
|
1199 |
val prt_fixes = if null fixes then [] |
|
1200 |
else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 :: |
|
1201 |
Pretty.commas (map prt_fix fixes))]; |
|
12057 | 1202 |
|
20310 | 1203 |
(*prems*) |
1204 |
val prems = Assumption.prems_of ctxt; |
|
1205 |
val len = length prems; |
|
20367 | 1206 |
val suppressed = len - ! prems_limit; |
20310 | 1207 |
val prt_prems = if null prems then [] |
20367 | 1208 |
else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @ |
1209 |
map (pretty_thm ctxt) (Library.drop (suppressed, prems)))]; |
|
20310 | 1210 |
in prt_structs @ prt_fixes @ prt_prems end; |
10810 | 1211 |
|
1212 |
||
1213 |
(* main context *) |
|
1214 |
||
16540 | 1215 |
fun pretty_context ctxt = |
10810 | 1216 |
let |
12057 | 1217 |
val prt_term = pretty_term ctxt; |
1218 |
val prt_typ = pretty_typ ctxt; |
|
1219 |
val prt_sort = pretty_sort ctxt; |
|
10810 | 1220 |
|
1221 |
(*theory*) |
|
12057 | 1222 |
val pretty_thy = Pretty.block |
17384 | 1223 |
[Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)]; |
10810 | 1224 |
|
1225 |
(*defaults*) |
|
1226 |
fun prt_atom prt prtT (x, X) = Pretty.block |
|
1227 |
[prt x, Pretty.str " ::", Pretty.brk 1, prtT X]; |
|
1228 |
||
1229 |
fun prt_var (x, ~1) = prt_term (Syntax.free x) |
|
1230 |
| prt_var xi = prt_term (Syntax.var xi); |
|
1231 |
||
1232 |
fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) |
|
1233 |
| prt_varT xi = prt_typ (TVar (xi, [])); |
|
1234 |
||
1235 |
val prt_defT = prt_atom prt_var prt_typ; |
|
1236 |
val prt_defS = prt_atom prt_varT prt_sort; |
|
16540 | 1237 |
|
20163 | 1238 |
val (types, sorts) = Variable.constraints_of ctxt; |
10810 | 1239 |
in |
18609 | 1240 |
verb single (K pretty_thy) @ |
18672
ac1a048ca7dd
uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents:
18619
diff
changeset
|
1241 |
pretty_ctxt ctxt @ |
21728 | 1242 |
verb (pretty_abbrevs false) (K ctxt) @ |
10810 | 1243 |
verb pretty_binds (K ctxt) @ |
12057 | 1244 |
verb pretty_lthms (K ctxt) @ |
10810 | 1245 |
verb pretty_cases (K ctxt) @ |
18609 | 1246 |
verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @ |
20163 | 1247 |
verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts))) |
10810 | 1248 |
end; |
1249 |
||
5819 | 1250 |
end; |