author | wenzelm |
Wed, 29 Nov 2006 15:44:59 +0100 | |
changeset 21592 | 8831206d7f41 |
parent 21585 | 2444f1d1127b |
child 21594 | 2859c94d67d4 |
permissions | -rw-r--r-- |
20894 | 1 |
(* Title: Pure/Isar/theory_target.ML |
2 |
ID: $Id$ |
|
3 |
Author: Makarius |
|
4 |
||
5 |
Common theory targets. |
|
6 |
*) |
|
7 |
||
8 |
signature THEORY_TARGET = |
|
9 |
sig |
|
21006 | 10 |
val peek: local_theory -> string option |
20962 | 11 |
val begin: bstring -> Proof.context -> local_theory |
20894 | 12 |
val init: xstring option -> theory -> local_theory |
13 |
val init_i: string option -> theory -> local_theory |
|
14 |
end; |
|
15 |
||
16 |
structure TheoryTarget: THEORY_TARGET = |
|
17 |
struct |
|
18 |
||
19 |
(** locale targets **) |
|
20 |
||
21006 | 21 |
(* context data *) |
22 |
||
23 |
structure Data = ProofDataFun |
|
24 |
( |
|
25 |
val name = "Isar/theory_target"; |
|
26 |
type T = string option; |
|
27 |
fun init _ = NONE; |
|
28 |
fun print _ _ = (); |
|
29 |
); |
|
30 |
||
31 |
val _ = Context.add_setup Data.init; |
|
32 |
val peek = Data.get; |
|
33 |
||
34 |
||
20894 | 35 |
(* pretty *) |
36 |
||
37 |
fun pretty loc ctxt = |
|
38 |
let |
|
39 |
val thy = ProofContext.theory_of ctxt; |
|
40 |
val fixes = map (fn (x, T) => (x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt)); |
|
21585 | 41 |
val assumes = map (fn A => (("", []), [(A, [])])) (map Thm.term_of (Assumption.assms_of ctxt)); |
20894 | 42 |
val elems = |
43 |
(if null fixes then [] else [Element.Fixes fixes]) @ |
|
44 |
(if null assumes then [] else [Element.Assumes assumes]); |
|
45 |
in |
|
20984 | 46 |
if loc = "" then |
47 |
[Pretty.block |
|
48 |
[Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy), |
|
49 |
Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]] |
|
50 |
else if null elems then [Pretty.str ("locale " ^ Locale.extern thy loc)] |
|
51 |
else |
|
52 |
[Pretty.big_list ("locale " ^ Locale.extern thy loc ^ " =") |
|
53 |
(map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)] |
|
20894 | 54 |
end; |
55 |
||
56 |
||
57 |
(* consts *) |
|
58 |
||
59 |
fun consts loc depends decls lthy = |
|
60 |
let |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
61 |
val is_loc = loc <> ""; |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
62 |
val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy))); |
20894 | 63 |
|
64 |
fun const ((c, T), mx) thy = |
|
65 |
let |
|
66 |
val U = map #2 xs ---> T; |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
67 |
val t = Term.list_comb (Const (Sign.full_name thy c, U), map Free xs); |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
68 |
val thy' = Sign.add_consts_authentic [(c, U, Syntax.unlocalize_mixfix is_loc mx)] thy; |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
69 |
in (((c, mx), t), thy') end; |
20894 | 70 |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
71 |
val (abbrs, lthy') = lthy |> LocalTheory.theory_result (fold_map const decls); |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
72 |
val defs = abbrs |> map (fn (x, t) => (x, (("", []), t))); |
20894 | 73 |
in |
74 |
lthy' |
|
21570 | 75 |
|> is_loc ? LocalTheory.abbrevs Syntax.default_mode abbrs |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
76 |
|> LocalDefs.add_defs defs |>> map (apsnd snd) |
20894 | 77 |
end; |
78 |
||
79 |
||
80 |
(* axioms *) |
|
81 |
||
82 |
local |
|
83 |
||
84 |
fun add_axiom hyps (name, prop) thy = |
|
85 |
let |
|
86 |
val name' = if name = "" then "axiom_" ^ serial_string () else name; |
|
87 |
val prop' = Logic.list_implies (hyps, prop); |
|
88 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')]; |
|
89 |
val axm = Drule.unvarify (Thm.get_axiom_i thy' (Sign.full_name thy' name')); |
|
90 |
val prems = map (Thm.assume o Thm.cterm_of thy') hyps; |
|
91 |
in (Drule.implies_elim_list axm prems, thy') end; |
|
92 |
||
93 |
in |
|
94 |
||
21433 | 95 |
fun axioms kind specs lthy = |
20915 | 96 |
let |
21585 | 97 |
val hyps = map Thm.term_of (Assumption.assms_of lthy); |
20915 | 98 |
fun axiom ((name, atts), props) thy = thy |
99 |
|> fold_map (add_axiom hyps) (PureThy.name_multi name props) |
|
100 |
|-> (fn ths => pair ((name, atts), [(ths, [])])); |
|
101 |
in |
|
102 |
lthy |
|
20984 | 103 |
|> fold (fold Variable.declare_term o snd) specs |
20915 | 104 |
|> LocalTheory.theory_result (fold_map axiom specs) |
21433 | 105 |
|-> LocalTheory.notes kind |
20915 | 106 |
end; |
20894 | 107 |
|
108 |
end; |
|
109 |
||
110 |
||
111 |
(* defs *) |
|
112 |
||
21570 | 113 |
infix also; |
114 |
fun eq1 also eq2 = |
|
115 |
eq2 COMP (eq1 COMP (Drule.incr_indexes2 eq1 eq2 transitive_thm)); |
|
116 |
||
20894 | 117 |
local |
118 |
||
21570 | 119 |
fun expand_defs ctxt t = |
120 |
let |
|
121 |
val thy = ProofContext.theory_of ctxt; |
|
122 |
val thy_ctxt = ProofContext.init thy; |
|
123 |
val ct = Thm.cterm_of thy t; |
|
124 |
val (defs, ct') = LocalDefs.expand_defs ctxt thy_ctxt (Drule.mk_term ct) ||> Drule.dest_term; |
|
125 |
in (Thm.term_of ct', Tactic.rewrite true defs ct) end; |
|
126 |
||
20894 | 127 |
fun add_def (name, prop) = |
128 |
Theory.add_defs_i false false [(name, prop)] #> |
|
129 |
(fn thy => (Drule.unvarify (Thm.get_axiom_i thy (Sign.full_name thy name)), thy)); |
|
130 |
||
131 |
in |
|
132 |
||
21570 | 133 |
fun defs kind args lthy0 = |
20894 | 134 |
let |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
135 |
fun def ((c, mx), ((name, atts), rhs)) lthy1 = |
20894 | 136 |
let |
21570 | 137 |
val (rhs', rhs_conv) = expand_defs lthy0 rhs; |
138 |
val xs = Variable.add_fixed (LocalTheory.target_of lthy0) rhs' []; |
|
20894 | 139 |
|
21570 | 140 |
val ([(lhs, lhs_def)], lthy2) = lthy1 |
141 |
|> LocalTheory.consts (member (op =) xs) [((c, Term.fastype_of rhs), mx)]; |
|
142 |
val lhs' = #2 (Logic.dest_equals (Thm.prop_of lhs_def)); |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
143 |
|
21570 | 144 |
val name' = Thm.def_name_optional c name; |
145 |
val (def, lthy3) = lthy2 |
|
20894 | 146 |
|> LocalTheory.theory_result (add_def (name', Logic.mk_equals (lhs', rhs'))); |
147 |
||
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
148 |
val eq = |
21570 | 149 |
(*c == loc.c xs*) lhs_def |
150 |
(*lhs' == rhs'*) also def |
|
151 |
(*rhs' == rhs*) also Thm.symmetric rhs_conv; |
|
20894 | 152 |
in ((lhs, ((name', atts), [([eq], [])])), lthy3) end; |
153 |
||
21570 | 154 |
val ((lhss, facts), lthy') = lthy0 |> fold_map def args |>> split_list; |
21433 | 155 |
val (res, lthy'') = lthy' |> LocalTheory.notes kind facts; |
20894 | 156 |
in (lhss ~~ map (apsnd the_single) res, lthy'') end; |
157 |
||
158 |
end; |
|
159 |
||
160 |
||
161 |
(* notes *) |
|
162 |
||
21585 | 163 |
(* FIXME tmp *) |
164 |
||
165 |
val maxidx_atts = fold Args.maxidx_values; |
|
166 |
||
167 |
fun export_standard_facts inner outer facts = |
|
168 |
let |
|
169 |
val thy = ProofContext.theory_of inner; |
|
170 |
val maxidx = |
|
171 |
fold (fn ((_, atts), bs) => maxidx_atts atts #> fold (maxidx_atts o #2) bs) facts ~1; |
|
172 |
val exp_fact = map (Thm.adjust_maxidx_thm maxidx) #> ProofContext.export_standard inner outer; |
|
173 |
val exp_term = Drule.term_rule thy (singleton exp_fact); |
|
174 |
val exp_typ = Logic.type_map exp_term; |
|
175 |
val morphism = |
|
176 |
Morphism.morphism {name = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact}; |
|
177 |
in Element.facts_map (Element.morph_ctxt morphism) facts end; |
|
178 |
||
179 |
||
21433 | 180 |
fun context_notes kind facts lthy = |
20894 | 181 |
let |
182 |
val facts' = facts |
|
21585 | 183 |
|> export_standard_facts lthy lthy |
20894 | 184 |
|> Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of lthy)); |
185 |
in |
|
186 |
lthy |
|
21433 | 187 |
|> ProofContext.set_stmt true |
20894 | 188 |
|> ProofContext.qualified_names |
21433 | 189 |
|> ProofContext.note_thmss_i kind facts' |
20894 | 190 |
||> ProofContext.restore_naming lthy |
21433 | 191 |
||> ProofContext.restore_stmt lthy |
20894 | 192 |
end; |
193 |
||
21433 | 194 |
fun theory_notes keep_atts kind facts lthy = lthy |> LocalTheory.theory (fn thy => |
20894 | 195 |
let |
196 |
val facts' = facts |
|
21585 | 197 |
|> export_standard_facts lthy (ProofContext.init thy) |
20915 | 198 |
|> Attrib.map_facts (if keep_atts then Attrib.attribute_i thy else K I); |
20894 | 199 |
in |
200 |
thy |
|
201 |
|> Sign.qualified_names |
|
21433 | 202 |
|> PureThy.note_thmss_i kind facts' |> #2 |
20894 | 203 |
|> Sign.restore_naming thy |
204 |
end); |
|
205 |
||
21433 | 206 |
fun locale_notes loc kind facts lthy = lthy |> LocalTheory.target (fn ctxt => |
21585 | 207 |
Locale.add_thmss loc kind (export_standard_facts lthy ctxt facts) ctxt); |
20894 | 208 |
|
21433 | 209 |
fun notes "" kind facts = theory_notes true kind facts #> context_notes kind facts |
210 |
| notes loc kind facts = theory_notes false kind facts #> |
|
211 |
locale_notes loc kind facts #> context_notes kind facts; |
|
20894 | 212 |
|
213 |
||
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
214 |
(* target declarations *) |
20894 | 215 |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
216 |
fun decl _ "" f = |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
217 |
LocalTheory.theory (Context.theory_map (f Morphism.identity)) #> |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
218 |
LocalTheory.target (Context.proof_map (f Morphism.identity)) |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
219 |
| decl add loc f = |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
220 |
LocalTheory.target (add loc (Context.proof_map o f)); |
20894 | 221 |
|
21498
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents:
21467
diff
changeset
|
222 |
val type_syntax = decl Locale.add_type_syntax; |
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents:
21467
diff
changeset
|
223 |
val term_syntax = decl Locale.add_term_syntax; |
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents:
21467
diff
changeset
|
224 |
val declaration = decl Locale.add_declaration; |
20894 | 225 |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
226 |
fun target_morphism loc lthy = |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
227 |
ProofContext.export_standard_morphism lthy (LocalTheory.target_of lthy) $> |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
228 |
Morphism.name_morphism (NameSpace.qualified loc); |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
229 |
|
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
230 |
fun target_name "" lthy = Sign.full_name (ProofContext.theory_of lthy) |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
231 |
| target_name _ lthy = ProofContext.full_name (LocalTheory.target_of lthy); |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
232 |
|
20894 | 233 |
|
234 |
(* init and exit *) |
|
235 |
||
21293 | 236 |
fun begin loc = |
237 |
Data.put (if loc = "" then NONE else SOME loc) #> |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
238 |
LocalTheory.init (NameSpace.base loc) |
21293 | 239 |
{pretty = pretty loc, |
240 |
consts = consts loc, |
|
241 |
axioms = axioms, |
|
242 |
defs = defs, |
|
243 |
notes = notes loc, |
|
21498
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents:
21467
diff
changeset
|
244 |
type_syntax = type_syntax loc, |
21293 | 245 |
term_syntax = term_syntax loc, |
246 |
declaration = declaration loc, |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
247 |
target_morphism = target_morphism loc, |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
248 |
target_name = target_name loc, |
21293 | 249 |
reinit = fn _ => begin loc o (if loc = "" then ProofContext.init else Locale.init loc), |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
250 |
exit = LocalTheory.target_of}; |
20894 | 251 |
|
20962 | 252 |
fun init_i NONE thy = begin "" (ProofContext.init thy) |
253 |
| init_i (SOME loc) thy = begin loc (Locale.init loc thy); |
|
20894 | 254 |
|
21276 | 255 |
fun init (SOME "-") thy = init_i NONE thy |
256 |
| init loc thy = init_i (Option.map (Locale.intern thy) loc) thy; |
|
20894 | 257 |
|
258 |
end; |