author | haftmann |
Thu, 04 Jan 2007 17:11:09 +0100 | |
changeset 21993 | 4b802a9e0738 |
parent 21845 | da545169fe06 |
child 22203 | efc0cdc01307 |
permissions | -rw-r--r-- |
20894 | 1 |
(* Title: Pure/Isar/theory_target.ML |
2 |
ID: $Id$ |
|
3 |
Author: Makarius |
|
4 |
||
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
5 |
Common theory/locale targets. |
20894 | 6 |
*) |
7 |
||
8 |
signature THEORY_TARGET = |
|
9 |
sig |
|
21006 | 10 |
val peek: local_theory -> string option |
21804 | 11 |
val fork_mixfix: bool -> bool -> mixfix -> mixfix * mixfix |
12 |
val begin: bool -> bstring -> Proof.context -> local_theory |
|
20894 | 13 |
val init: xstring option -> theory -> local_theory |
14 |
val init_i: string option -> theory -> local_theory |
|
15 |
end; |
|
16 |
||
17 |
structure TheoryTarget: THEORY_TARGET = |
|
18 |
struct |
|
19 |
||
20 |
(** locale targets **) |
|
21 |
||
21006 | 22 |
(* context data *) |
23 |
||
24 |
structure Data = ProofDataFun |
|
25 |
( |
|
26 |
val name = "Isar/theory_target"; |
|
27 |
type T = string option; |
|
28 |
fun init _ = NONE; |
|
29 |
fun print _ _ = (); |
|
30 |
); |
|
31 |
||
32 |
val _ = Context.add_setup Data.init; |
|
33 |
val peek = Data.get; |
|
34 |
||
35 |
||
20894 | 36 |
(* pretty *) |
37 |
||
38 |
fun pretty loc ctxt = |
|
39 |
let |
|
40 |
val thy = ProofContext.theory_of ctxt; |
|
41 |
val fixes = map (fn (x, T) => (x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt)); |
|
21585 | 42 |
val assumes = map (fn A => (("", []), [(A, [])])) (map Thm.term_of (Assumption.assms_of ctxt)); |
20894 | 43 |
val elems = |
44 |
(if null fixes then [] else [Element.Fixes fixes]) @ |
|
45 |
(if null assumes then [] else [Element.Assumes assumes]); |
|
46 |
in |
|
20984 | 47 |
if loc = "" then |
48 |
[Pretty.block |
|
49 |
[Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy), |
|
50 |
Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]] |
|
51 |
else if null elems then [Pretty.str ("locale " ^ Locale.extern thy loc)] |
|
52 |
else |
|
53 |
[Pretty.big_list ("locale " ^ Locale.extern thy loc ^ " =") |
|
54 |
(map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)] |
|
20894 | 55 |
end; |
56 |
||
57 |
||
21804 | 58 |
(* consts *) |
59 |
||
60 |
fun fork_mixfix is_class is_loc mx = |
|
61 |
let |
|
62 |
val mx' = Syntax.unlocalize_mixfix mx; |
|
63 |
val mx1 = if is_class orelse (is_loc andalso mx = mx') then NoSyn else mx'; |
|
64 |
val mx2 = if is_loc then mx else NoSyn; |
|
65 |
in (mx1, mx2) end; |
|
66 |
||
21825 | 67 |
fun internal_abbrev prmode ((c, mx), t) = |
68 |
LocalTheory.term_syntax (ProofContext.target_abbrev prmode ((c, mx), t)) #> |
|
69 |
ProofContext.add_abbrev Syntax.internalM (c, t) #> snd; |
|
21804 | 70 |
|
71 |
fun consts is_class is_loc depends decls lthy = |
|
72 |
let |
|
73 |
val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy))); |
|
21747 | 74 |
|
21804 | 75 |
fun const ((c, T), mx) thy = |
76 |
let |
|
77 |
val U = map #2 xs ---> T; |
|
78 |
val t = Term.list_comb (Const (Sign.full_name thy c, U), map Free xs); |
|
79 |
val (mx1, mx2) = fork_mixfix is_class is_loc mx; |
|
80 |
val thy' = Sign.add_consts_authentic [(c, U, mx1)] thy; |
|
81 |
in (((c, mx2), t), thy') end; |
|
82 |
||
83 |
val (abbrs, lthy') = lthy |> LocalTheory.theory_result (fold_map const decls); |
|
84 |
val defs = map (apsnd (pair ("", []))) abbrs; |
|
85 |
in |
|
86 |
lthy' |
|
21825 | 87 |
|> is_loc ? fold (internal_abbrev Syntax.default_mode) abbrs |
21804 | 88 |
|> LocalDefs.add_defs defs |>> map (apsnd snd) |
89 |
end; |
|
90 |
||
91 |
||
92 |
(* abbrev *) |
|
21747 | 93 |
|
94 |
fun occ_params ctxt ts = |
|
95 |
#1 (ProofContext.inferred_fixes ctxt) |
|
96 |
|> filter (member (op =) (fold (Variable.add_fixed ctxt) ts [])); |
|
97 |
||
21804 | 98 |
fun abbrev is_class is_loc prmode ((raw_c, mx), raw_t) lthy = |
21747 | 99 |
let |
100 |
val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy); |
|
101 |
val target = LocalTheory.target_of lthy; |
|
102 |
val target_morphism = LocalTheory.target_morphism lthy; |
|
103 |
||
104 |
val c = Morphism.name target_morphism raw_c; |
|
105 |
val t = Morphism.term target_morphism raw_t; |
|
106 |
val xs = map Free (occ_params target [t]); |
|
107 |
val u = fold_rev Term.lambda xs t; |
|
21804 | 108 |
val U = Term.fastype_of u; |
21747 | 109 |
val u' = singleton (Variable.export_terms (Variable.declare_term u target) thy_ctxt) u; |
110 |
||
21804 | 111 |
val ((lhs as Const (full_c, _), rhs), lthy1) = lthy |
21747 | 112 |
|> LocalTheory.theory_result (Sign.add_abbrev (#1 prmode) (c, u')); |
21804 | 113 |
val (mx1, mx2) = fork_mixfix is_class is_loc mx; |
21747 | 114 |
in |
115 |
lthy1 |
|
21804 | 116 |
|> LocalTheory.theory (Sign.add_notation prmode [(lhs, mx1)]) |
21825 | 117 |
|> is_loc ? internal_abbrev prmode ((c, mx2), Term.list_comb (Const (full_c, U), xs)) |
21804 | 118 |
|> ProofContext.local_abbrev (c, rhs) |
21747 | 119 |
end; |
120 |
||
121 |
||
20894 | 122 |
(* defs *) |
123 |
||
21845 | 124 |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
125 |
|
21845 | 126 |
local |
21570 | 127 |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
128 |
fun expand_term ctxt t = |
21570 | 129 |
let |
130 |
val thy = ProofContext.theory_of ctxt; |
|
131 |
val thy_ctxt = ProofContext.init thy; |
|
132 |
val ct = Thm.cterm_of thy t; |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
133 |
val (defs, ct') = LocalDefs.export ctxt thy_ctxt (Drule.mk_term ct) ||> Drule.dest_term; |
21708 | 134 |
in (Thm.term_of ct', MetaSimplifier.rewrite true defs ct) end; |
21570 | 135 |
|
20894 | 136 |
fun add_def (name, prop) = |
137 |
Theory.add_defs_i false false [(name, prop)] #> |
|
138 |
(fn thy => (Drule.unvarify (Thm.get_axiom_i thy (Sign.full_name thy name)), thy)); |
|
139 |
||
140 |
in |
|
141 |
||
21570 | 142 |
fun defs kind args lthy0 = |
20894 | 143 |
let |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
144 |
fun def ((c, mx), ((name, atts), rhs)) lthy1 = |
20894 | 145 |
let |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
146 |
val (rhs', rhs_conv) = expand_term lthy0 rhs; |
21570 | 147 |
val xs = Variable.add_fixed (LocalTheory.target_of lthy0) rhs' []; |
20894 | 148 |
|
21570 | 149 |
val ([(lhs, lhs_def)], lthy2) = lthy1 |
150 |
|> LocalTheory.consts (member (op =) xs) [((c, Term.fastype_of rhs), mx)]; |
|
151 |
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
|
152 |
|
21570 | 153 |
val name' = Thm.def_name_optional c name; |
154 |
val (def, lthy3) = lthy2 |
|
20894 | 155 |
|> LocalTheory.theory_result (add_def (name', Logic.mk_equals (lhs', rhs'))); |
156 |
||
21845 | 157 |
val eq = LocalDefs.trans_terms lthy3 |
158 |
[(*c == loc.c xs*) lhs_def, |
|
159 |
(*lhs' == rhs'*) def, |
|
160 |
(*rhs' == rhs*) Thm.symmetric rhs_conv]; |
|
20894 | 161 |
in ((lhs, ((name', atts), [([eq], [])])), lthy3) end; |
162 |
||
21570 | 163 |
val ((lhss, facts), lthy') = lthy0 |> fold_map def args |>> split_list; |
21433 | 164 |
val (res, lthy'') = lthy' |> LocalTheory.notes kind facts; |
20894 | 165 |
in (lhss ~~ map (apsnd the_single) res, lthy'') end; |
166 |
||
167 |
end; |
|
168 |
||
169 |
||
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
170 |
(* axioms *) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
171 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
172 |
local |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
173 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
174 |
fun add_axiom hyps (name, prop) thy = |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
175 |
let |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
176 |
val name' = if name = "" then "axiom_" ^ serial_string () else name; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
177 |
val prop' = Logic.list_implies (hyps, prop); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
178 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')]; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
179 |
val axm = Drule.unvarify (Thm.get_axiom_i thy' (Sign.full_name thy' name')); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
180 |
val prems = map (Thm.assume o Thm.cterm_of thy') hyps; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
181 |
in (Drule.implies_elim_list axm prems, thy') end; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
182 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
183 |
in |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
184 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
185 |
fun axioms kind specs lthy = |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
186 |
let |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
187 |
val hyps = map Thm.term_of (Assumption.assms_of lthy); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
188 |
fun axiom ((name, atts), props) thy = thy |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
189 |
|> fold_map (add_axiom hyps) (PureThy.name_multi name props) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
190 |
|-> (fn ths => pair ((name, atts), [(ths, [])])); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
191 |
in |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
192 |
lthy |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
193 |
|> fold (fold Variable.declare_term o snd) specs |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
194 |
|> LocalTheory.theory_result (fold_map axiom specs) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
195 |
|-> LocalTheory.notes kind |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
196 |
end; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
197 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
198 |
end; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
199 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
200 |
|
20894 | 201 |
(* notes *) |
202 |
||
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
203 |
fun import_export_proof ctxt (name, raw_th) = |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
204 |
let |
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
205 |
val thy = ProofContext.theory_of ctxt; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
206 |
val thy_ctxt = ProofContext.init thy; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
207 |
val certT = Thm.ctyp_of thy; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
208 |
val cert = Thm.cterm_of thy; |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
209 |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
210 |
(*export assumes/defines*) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
211 |
val th = Goal.norm_result raw_th; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
212 |
val (defs, th') = LocalDefs.export ctxt thy_ctxt th; |
21708 | 213 |
val concl_conv = MetaSimplifier.rewrite true defs (Thm.cprop_of th); |
214 |
val assms = map (MetaSimplifier.rewrite_rule defs o Thm.assume) (Assumption.assms_of ctxt); |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
215 |
val nprems = Thm.nprems_of th' - Thm.nprems_of th; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
216 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
217 |
(*export fixes*) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
218 |
val tfrees = map TFree (Drule.fold_terms Term.add_tfrees th' []); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
219 |
val frees = map Free (Drule.fold_terms Term.add_frees th' []); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
220 |
val (th'' :: vs) = (th' :: map (Drule.mk_term o cert) (map Logic.mk_type tfrees @ frees)) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
221 |
|> Variable.export ctxt thy_ctxt |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
222 |
|> Drule.zero_var_indexes_list; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
223 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
224 |
(*thm definition*) |
21644 | 225 |
val result = th'' |
226 |
|> PureThy.name_thm true true "" |
|
227 |
|> Goal.close_result |
|
228 |
|> PureThy.name_thm true true name; |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
229 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
230 |
(*import fixes*) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
231 |
val (tvars, vars) = |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
232 |
chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
233 |
|>> map Logic.dest_type; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
234 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
235 |
val instT = map_filter (fn (TVar v, T) => SOME (v, T) | _ => NONE) (tvars ~~ tfrees); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
236 |
val inst = filter (is_Var o fst) (vars ~~ frees); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
237 |
val cinstT = map (pairself certT o apfst TVar) instT; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
238 |
val cinst = map (pairself (cert o Term.map_types (TermSubst.instantiateT instT))) inst; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
239 |
val result' = Thm.instantiate (cinstT, cinst) result; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
240 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
241 |
(*import assumes/defines*) |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
242 |
val assm_tac = FIRST' (map (fn assm => Tactic.compose_tac (false, assm, 0)) assms); |
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
243 |
val result'' = |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
244 |
(case SINGLE (Seq.INTERVAL assm_tac 1 nprems) result' of |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
245 |
NONE => raise THM ("Failed to re-import result", 0, [result']) |
21845 | 246 |
| SOME res => LocalDefs.trans_props ctxt [res, Thm.symmetric concl_conv]) |
21644 | 247 |
|> Goal.norm_result |
248 |
|> PureThy.name_thm false false name; |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
249 |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
250 |
in (result'', result) end; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
251 |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
252 |
fun notes loc kind facts lthy = |
21585 | 253 |
let |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
254 |
val is_loc = loc <> ""; |
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
255 |
val thy = ProofContext.theory_of lthy; |
21615
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
wenzelm
parents:
21613
diff
changeset
|
256 |
val full = LocalTheory.full_name lthy; |
21585 | 257 |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
258 |
val facts' = facts |
21615
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
wenzelm
parents:
21613
diff
changeset
|
259 |
|> map (fn (a, bs) => (a, PureThy.burrow_fact (PureThy.name_multi (full (fst a))) bs)) |
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
wenzelm
parents:
21613
diff
changeset
|
260 |
|> PureThy.map_facts (import_export_proof lthy); |
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
261 |
val local_facts = PureThy.map_facts #1 facts' |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
262 |
|> Attrib.map_facts (Attrib.attribute_i thy); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
263 |
val target_facts = PureThy.map_facts #1 facts' |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
264 |
|> is_loc ? Element.facts_map (Element.morph_ctxt (LocalTheory.target_morphism lthy)); |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
265 |
val global_facts = PureThy.map_facts #2 facts' |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
266 |
|> Attrib.map_facts (if is_loc then K I else Attrib.attribute_i thy); |
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
267 |
in |
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
268 |
lthy |> LocalTheory.theory |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
269 |
(Sign.qualified_names |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
270 |
#> PureThy.note_thmss_i kind global_facts #> snd |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
271 |
#> Sign.restore_naming thy) |
21585 | 272 |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
273 |
|> is_loc ? LocalTheory.target (Locale.add_thmss loc kind target_facts) |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
274 |
|
21433 | 275 |
|> ProofContext.set_stmt true |
20894 | 276 |
|> ProofContext.qualified_names |
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
277 |
|> ProofContext.note_thmss_i kind local_facts |
20894 | 278 |
||> ProofContext.restore_naming lthy |
21433 | 279 |
||> ProofContext.restore_stmt lthy |
20894 | 280 |
end; |
281 |
||
282 |
||
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
283 |
(* target declarations *) |
20894 | 284 |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
285 |
fun target_decl _ "" f = |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
286 |
LocalTheory.theory (Context.theory_map (f Morphism.identity)) #> |
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
287 |
LocalTheory.target (Context.proof_map (f Morphism.identity)) |
21668 | 288 |
| target_decl add loc f = LocalTheory.target (add loc f); |
20894 | 289 |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
290 |
val type_syntax = target_decl Locale.add_type_syntax; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
291 |
val term_syntax = target_decl Locale.add_term_syntax; |
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
292 |
val declaration = target_decl Locale.add_declaration; |
20894 | 293 |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
294 |
fun target_morphism loc lthy = |
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
295 |
ProofContext.export_morphism lthy (LocalTheory.target_of lthy) $> |
21613
ae3bb930b50f
notes: more careful treatment of Goal.close_result;
wenzelm
parents:
21611
diff
changeset
|
296 |
Morphism.thm_morphism Goal.norm_result; |
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
297 |
|
21668 | 298 |
fun target_name loc lthy = |
299 |
(if loc = "" then Sign.naming_of (ProofContext.theory_of lthy) |
|
300 |
else ProofContext.naming_of (LocalTheory.target_of lthy)) |
|
301 |
|> NameSpace.qualified_names |
|
302 |
|> NameSpace.full; |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
303 |
|
20894 | 304 |
|
305 |
(* init and exit *) |
|
306 |
||
21804 | 307 |
fun begin is_class loc = |
21747 | 308 |
let val is_loc = loc <> "" in |
309 |
Data.put (if is_loc then SOME loc else NONE) #> |
|
310 |
LocalTheory.init (NameSpace.base loc) |
|
311 |
{pretty = pretty loc, |
|
21804 | 312 |
consts = consts is_class is_loc, |
21747 | 313 |
axioms = axioms, |
21804 | 314 |
abbrev = abbrev is_class is_loc, |
21747 | 315 |
defs = defs, |
316 |
notes = notes loc, |
|
317 |
type_syntax = type_syntax loc, |
|
318 |
term_syntax = term_syntax loc, |
|
319 |
declaration = declaration loc, |
|
320 |
target_morphism = target_morphism loc, |
|
321 |
target_name = target_name loc, |
|
21804 | 322 |
reinit = fn _ => |
323 |
begin is_class loc o (if is_loc then Locale.init loc else ProofContext.init), |
|
21747 | 324 |
exit = LocalTheory.target_of} |
325 |
end; |
|
20894 | 326 |
|
21804 | 327 |
fun init_i NONE thy = begin false "" (ProofContext.init thy) |
328 |
| init_i (SOME loc) thy = |
|
329 |
begin (can (Sign.certify_class thy) loc) (* FIXME approximation *) |
|
330 |
loc (Locale.init loc thy); |
|
20894 | 331 |
|
21276 | 332 |
fun init (SOME "-") thy = init_i NONE thy |
333 |
| init loc thy = init_i (Option.map (Locale.intern thy) loc) thy; |
|
20894 | 334 |
|
335 |
end; |