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
|
|
10 |
val init: xstring option -> theory -> local_theory
|
|
11 |
val init_i: string option -> theory -> local_theory
|
20915
|
12 |
val exit: local_theory -> Proof.context * theory
|
|
13 |
val mapping: xstring option -> (local_theory -> local_theory) -> theory -> Proof.context * theory
|
20894
|
14 |
end;
|
|
15 |
|
|
16 |
structure TheoryTarget: THEORY_TARGET =
|
|
17 |
struct
|
|
18 |
|
|
19 |
(** locale targets **)
|
|
20 |
|
|
21 |
(* pretty *)
|
|
22 |
|
|
23 |
fun pretty loc ctxt =
|
|
24 |
let
|
|
25 |
val thy = ProofContext.theory_of ctxt;
|
|
26 |
val fixes = map (fn (x, T) => (x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt));
|
|
27 |
val assumes = map (fn A => (("", []), [(A, [])])) (Assumption.assms_of ctxt);
|
|
28 |
val elems =
|
|
29 |
(if null fixes then [] else [Element.Fixes fixes]) @
|
|
30 |
(if null assumes then [] else [Element.Assumes assumes]);
|
|
31 |
in
|
|
32 |
([Pretty.block
|
|
33 |
[Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy),
|
|
34 |
Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]] @
|
|
35 |
(if loc = "" then []
|
|
36 |
else if null elems then [Pretty.str ("locale " ^ Locale.extern thy loc)]
|
|
37 |
else
|
|
38 |
[Pretty.big_list ("locale " ^ Locale.extern thy loc ^ " =")
|
|
39 |
(map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)]))
|
|
40 |
|> Pretty.chunks
|
|
41 |
end;
|
|
42 |
|
|
43 |
|
|
44 |
(* consts *)
|
|
45 |
|
|
46 |
fun consts loc depends decls lthy =
|
|
47 |
let
|
|
48 |
val xs = filter depends (#1 (ProofContext.inferred_fixes lthy));
|
|
49 |
val ys = filter (Variable.newly_fixed lthy (LocalTheory.target_of lthy) o #1) xs;
|
|
50 |
|
|
51 |
fun const ((c, T), mx) thy =
|
|
52 |
let
|
|
53 |
val U = map #2 xs ---> T;
|
|
54 |
val mx' = if null ys then mx else NoSyn;
|
|
55 |
val mx'' = Syntax.unlocalize_mixfix (loc <> "") mx';
|
|
56 |
|
|
57 |
val t = Term.list_comb (Const (Sign.full_name thy c, U), map Free xs);
|
|
58 |
val defn = ((c, if loc <> "" then mx else NoSyn (* FIXME !? *)), (("", []), t));
|
|
59 |
val abbr = ((c, mx'), fold_rev (lambda o Free) ys t);
|
|
60 |
val thy' = Sign.add_consts_authentic [(c, U, mx'')] thy;
|
|
61 |
in ((defn, abbr), thy') end;
|
|
62 |
|
|
63 |
val ((defns, abbrs), lthy') = lthy
|
|
64 |
|> LocalTheory.theory_result (fold_map const decls) |>> split_list;
|
|
65 |
in
|
|
66 |
lthy'
|
|
67 |
|> K (loc <> "") ? LocalTheory.abbrevs Syntax.default_mode abbrs
|
|
68 |
|> LocalDefs.add_defs defns |>> map (apsnd snd)
|
|
69 |
end;
|
|
70 |
|
|
71 |
|
|
72 |
(* axioms *)
|
|
73 |
|
|
74 |
local
|
|
75 |
|
|
76 |
fun add_axiom hyps (name, prop) thy =
|
|
77 |
let
|
|
78 |
val name' = if name = "" then "axiom_" ^ serial_string () else name;
|
|
79 |
val prop' = Logic.list_implies (hyps, prop);
|
|
80 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')];
|
|
81 |
val axm = Drule.unvarify (Thm.get_axiom_i thy' (Sign.full_name thy' name'));
|
|
82 |
val prems = map (Thm.assume o Thm.cterm_of thy') hyps;
|
|
83 |
in (Drule.implies_elim_list axm prems, thy') end;
|
|
84 |
|
|
85 |
in
|
|
86 |
|
20915
|
87 |
fun axioms specs lthy =
|
|
88 |
let
|
|
89 |
val hyps = Assumption.assms_of lthy;
|
|
90 |
fun axiom ((name, atts), props) thy = thy
|
|
91 |
|> fold_map (add_axiom hyps) (PureThy.name_multi name props)
|
|
92 |
|-> (fn ths => pair ((name, atts), [(ths, [])]));
|
|
93 |
in
|
|
94 |
lthy
|
|
95 |
|> LocalTheory.theory_result (fold_map axiom specs)
|
|
96 |
|-> LocalTheory.notes
|
|
97 |
end;
|
20894
|
98 |
|
|
99 |
end;
|
|
100 |
|
|
101 |
|
|
102 |
(* defs *)
|
|
103 |
|
|
104 |
local
|
|
105 |
|
|
106 |
fun add_def (name, prop) =
|
|
107 |
Theory.add_defs_i false false [(name, prop)] #>
|
|
108 |
(fn thy => (Drule.unvarify (Thm.get_axiom_i thy (Sign.full_name thy name)), thy));
|
|
109 |
|
|
110 |
fun expand_defs lthy =
|
|
111 |
Drule.term_rule (ProofContext.theory_of lthy)
|
|
112 |
(Assumption.export false lthy (LocalTheory.target_of lthy));
|
|
113 |
|
|
114 |
in
|
|
115 |
|
|
116 |
fun defs args lthy =
|
|
117 |
let
|
|
118 |
fun def ((x, mx), ((name, atts), rhs)) lthy1 =
|
|
119 |
let
|
|
120 |
val name' = Thm.def_name_optional x name;
|
|
121 |
val T = Term.fastype_of rhs;
|
|
122 |
val rhs' = expand_defs lthy1 rhs;
|
|
123 |
val depends = member (op =) (Term.add_frees rhs' []);
|
|
124 |
val defined = filter_out depends (Term.add_frees rhs []);
|
|
125 |
|
|
126 |
val rhs_conv = rhs
|
|
127 |
|> Thm.cterm_of (ProofContext.theory_of lthy1)
|
|
128 |
|> Tactic.rewrite true (map_filter (LocalDefs.find_def lthy1 o #1) defined);
|
|
129 |
|
|
130 |
val ([(lhs, local_def)], lthy2) = lthy1
|
|
131 |
|> LocalTheory.consts depends [((x, T), mx)];
|
|
132 |
val lhs' = #2 (Logic.dest_equals (Thm.prop_of local_def));
|
|
133 |
|
|
134 |
val (global_def, lthy3) = lthy2
|
|
135 |
|> LocalTheory.theory_result (add_def (name', Logic.mk_equals (lhs', rhs')));
|
|
136 |
|
|
137 |
val eq = Thm.transitive (Thm.transitive local_def global_def) (Thm.symmetric rhs_conv);
|
|
138 |
in ((lhs, ((name', atts), [([eq], [])])), lthy3) end;
|
|
139 |
|
|
140 |
val ((lhss, facts), lthy') = lthy |> fold_map def args |>> split_list;
|
|
141 |
val (res, lthy'') = lthy' |> LocalTheory.notes facts;
|
|
142 |
in (lhss ~~ map (apsnd the_single) res, lthy'') end;
|
|
143 |
|
|
144 |
end;
|
|
145 |
|
|
146 |
|
|
147 |
(* notes *)
|
|
148 |
|
|
149 |
fun context_notes facts lthy =
|
|
150 |
let
|
|
151 |
val facts' = facts
|
|
152 |
|> Element.export_standard_facts lthy lthy
|
|
153 |
|> Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of lthy));
|
|
154 |
in
|
|
155 |
lthy
|
|
156 |
|> ProofContext.qualified_names
|
|
157 |
|> ProofContext.note_thmss_i facts'
|
|
158 |
||> ProofContext.restore_naming lthy
|
|
159 |
end;
|
|
160 |
|
20915
|
161 |
fun theory_notes keep_atts facts lthy = lthy |> LocalTheory.theory (fn thy =>
|
20894
|
162 |
let
|
|
163 |
val facts' = facts
|
|
164 |
|> Element.export_standard_facts lthy (ProofContext.init thy)
|
20915
|
165 |
|> Attrib.map_facts (if keep_atts then Attrib.attribute_i thy else K I);
|
20894
|
166 |
in
|
|
167 |
thy
|
|
168 |
|> Sign.qualified_names
|
|
169 |
|> PureThy.note_thmss_i "" facts' |> #2
|
|
170 |
|> Sign.restore_naming thy
|
|
171 |
end);
|
|
172 |
|
|
173 |
fun locale_notes loc facts lthy = lthy |> LocalTheory.target (fn ctxt =>
|
|
174 |
#2 (Locale.add_thmss "" loc (Element.export_standard_facts lthy ctxt facts) ctxt));
|
|
175 |
|
20915
|
176 |
fun notes "" facts = theory_notes true facts #> context_notes facts
|
|
177 |
| notes loc facts = theory_notes false facts #> locale_notes loc facts #> context_notes facts;
|
20894
|
178 |
|
|
179 |
|
|
180 |
(* declarations *)
|
|
181 |
|
|
182 |
fun term_syntax "" f = LocalTheory.theory (Context.theory_map f)
|
|
183 |
| term_syntax loc f = LocalTheory.target (Locale.add_term_syntax loc (Context.proof_map f));
|
|
184 |
|
|
185 |
fun declaration "" f = LocalTheory.theory (Context.theory_map f)
|
|
186 |
| declaration loc f = I; (* FIXME *)
|
|
187 |
|
|
188 |
|
|
189 |
|
|
190 |
(* init and exit *)
|
|
191 |
|
20915
|
192 |
fun target_operations loc exit : LocalTheory.operations =
|
20894
|
193 |
{pretty = pretty loc,
|
|
194 |
consts = consts loc,
|
|
195 |
axioms = axioms,
|
|
196 |
defs = defs,
|
|
197 |
notes = notes loc,
|
|
198 |
term_syntax = term_syntax loc,
|
20915
|
199 |
declaration = declaration loc,
|
|
200 |
exit = exit};
|
20894
|
201 |
|
20915
|
202 |
fun init_i NONE thy = LocalTheory.init (SOME "") (target_operations "" I) (ProofContext.init thy)
|
20894
|
203 |
| init_i (SOME loc) thy =
|
20915
|
204 |
LocalTheory.init (SOME (NameSpace.base loc)) (target_operations loc I) (Locale.init loc thy);
|
20894
|
205 |
|
|
206 |
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy;
|
|
207 |
|
20915
|
208 |
val exit = LocalTheory.exit #> (fn lthy => (lthy, ProofContext.theory_of lthy));
|
20894
|
209 |
|
|
210 |
fun mapping loc f = init loc #> f #> exit;
|
|
211 |
|
|
212 |
end;
|