| author | wenzelm |
| Sat, 06 Oct 2007 16:50:08 +0200 | |
| changeset 24868 | 2990c327d8c6 |
| parent 24848 | 5dbbd33c3236 |
| child 24901 | d3cbf79769b9 |
| 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 |
| 22353 | 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 |
type T = string option; |
|
26 |
fun init _ = NONE; |
|
27 |
); |
|
28 |
||
29 |
val peek = Data.get; |
|
30 |
||
31 |
||
| 20894 | 32 |
(* pretty *) |
33 |
||
34 |
fun pretty loc ctxt = |
|
35 |
let |
|
36 |
val thy = ProofContext.theory_of ctxt; |
|
37 |
val fixes = map (fn (x, T) => (x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt)); |
|
| 21585 | 38 |
val assumes = map (fn A => (("", []), [(A, [])])) (map Thm.term_of (Assumption.assms_of ctxt));
|
| 20894 | 39 |
val elems = |
40 |
(if null fixes then [] else [Element.Fixes fixes]) @ |
|
41 |
(if null assumes then [] else [Element.Assumes assumes]); |
|
42 |
in |
|
| 20984 | 43 |
if loc = "" then |
44 |
[Pretty.block |
|
45 |
[Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy), |
|
46 |
Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]] |
|
47 |
else if null elems then [Pretty.str ("locale " ^ Locale.extern thy loc)]
|
|
48 |
else |
|
49 |
[Pretty.big_list ("locale " ^ Locale.extern thy loc ^ " =")
|
|
50 |
(map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)] |
|
| 20894 | 51 |
end; |
52 |
||
53 |
||
| 24838 | 54 |
(* target declarations *) |
55 |
||
56 |
fun target_morphism loc lthy = |
|
57 |
ProofContext.export_morphism lthy (LocalTheory.target_of lthy) $> |
|
58 |
Morphism.thm_morphism Goal.norm_result; |
|
59 |
||
60 |
fun target_decl add loc d lthy = |
|
61 |
let |
|
62 |
val d' = Morphism.transform (target_morphism loc lthy) d; |
|
63 |
val d0 = Morphism.form d'; |
|
64 |
in |
|
65 |
if loc = "" then |
|
66 |
lthy |
|
67 |
|> LocalTheory.theory (Context.theory_map d0) |
|
68 |
|> LocalTheory.target (Context.proof_map d0) |
|
69 |
else |
|
70 |
lthy |
|
71 |
|> LocalTheory.target (add loc d') |
|
72 |
end; |
|
73 |
||
74 |
val type_syntax = target_decl Locale.add_type_syntax; |
|
75 |
val term_syntax = target_decl Locale.add_term_syntax; |
|
76 |
val declaration = target_decl Locale.add_declaration; |
|
77 |
||
78 |
fun target_naming loc lthy = |
|
79 |
(if loc = "" then Sign.naming_of (ProofContext.theory_of lthy) |
|
80 |
else ProofContext.naming_of (LocalTheory.target_of lthy)) |
|
81 |
|> NameSpace.qualified_names; |
|
82 |
||
83 |
||
| 21804 | 84 |
(* consts *) |
85 |
||
|
23205
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
86 |
fun target_abbrev prmode ((c, mx), rhs) phi = |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
87 |
let |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
88 |
val c' = Morphism.name phi c; |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
89 |
val rhs' = Morphism.term phi rhs; |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
90 |
val arg' = (c', rhs'); |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
91 |
val eq_heads = |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
92 |
(case pairself Term.head_of (rhs, rhs') of |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
93 |
(Const (a, _), Const (a', _)) => a = a' |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
94 |
| _ => false); |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
95 |
in |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
96 |
eq_heads ? (Context.mapping_result |
| 24779 | 97 |
(Sign.add_abbrev Syntax.internalM [] arg') (ProofContext.add_abbrev Syntax.internalM [] arg') |
|
23205
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
98 |
#-> (fn (lhs, _) => |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
99 |
Term.equiv_types (rhs, rhs') ? |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
100 |
Morphism.form (ProofContext.target_notation prmode [(lhs, mx)]))) |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
101 |
end; |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
102 |
|
| 24779 | 103 |
fun internal_abbrev prmode ((c, mx), t) lthy = lthy |
| 22380 | 104 |
(* FIXME really permissive *) |
| 24779 | 105 |
|> LocalTheory.term_syntax (perhaps o try o target_abbrev prmode ((c, mx), t)) |
106 |
|> ProofContext.add_abbrev Syntax.internalM (ContextPosition.properties_of lthy) (c, t) |
|
107 |
|> snd; |
|
| 21804 | 108 |
|
| 22353 | 109 |
fun consts is_loc some_class depends decls lthy = |
| 21804 | 110 |
let |
111 |
val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy))); |
|
| 21747 | 112 |
|
| 21804 | 113 |
fun const ((c, T), mx) thy = |
114 |
let |
|
115 |
val U = map #2 xs ---> T; |
|
116 |
val t = Term.list_comb (Const (Sign.full_name thy c, U), map Free xs); |
|
| 24218 | 117 |
val (mx1, mx2) = Class.fork_mixfix is_loc some_class mx; |
| 24779 | 118 |
val thy' = Sign.add_consts_authentic (ContextPosition.properties_of lthy) [(c, U, mx1)] thy; |
| 21804 | 119 |
in (((c, mx2), t), thy') end; |
120 |
||
| 23087 | 121 |
fun const_class (SOME class) ((c, _), mx) (_, t) = |
| 24218 | 122 |
Class.add_const_in_class class ((c, t), mx) |
| 23087 | 123 |
| const_class NONE _ _ = I; |
| 24657 | 124 |
fun hide_abbrev (SOME class) abbrs thy = |
125 |
let |
|
126 |
val raw_cs = map (fst o fst) abbrs; |
|
127 |
val prfx = (Logic.const_of_class o NameSpace.base) class; |
|
128 |
fun mk_name c = |
|
129 |
let |
|
130 |
val n1 = Sign.full_name thy c; |
|
131 |
val n2 = NameSpace.qualifier n1; |
|
132 |
val n3 = NameSpace.base n1; |
|
133 |
in NameSpace.implode [n2, prfx, prfx, n3] end; |
|
134 |
val cs = map mk_name raw_cs; |
|
135 |
in |
|
136 |
Sign.hide_consts_i true cs thy |
|
137 |
end |
|
138 |
| hide_abbrev NONE _ thy = thy; |
|
| 23087 | 139 |
val (abbrs, lthy') = lthy |
140 |
|> LocalTheory.theory_result (fold_map const decls) |
|
| 21804 | 141 |
val defs = map (apsnd (pair ("", []))) abbrs;
|
|
24818
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
142 |
|
| 21804 | 143 |
in |
144 |
lthy' |
|
| 24657 | 145 |
|> LocalTheory.raw_theory (fold2 (const_class some_class) decls abbrs) |
| 21825 | 146 |
|> is_loc ? fold (internal_abbrev Syntax.default_mode) abbrs |
| 24657 | 147 |
|> LocalTheory.raw_theory (hide_abbrev some_class abbrs) |
148 |
(*FIXME abbreviations should never occur*) |
|
| 23087 | 149 |
|> LocalDefs.add_defs defs |
150 |
|>> map (apsnd snd) |
|
| 21804 | 151 |
end; |
152 |
||
153 |
||
154 |
(* abbrev *) |
|
| 21747 | 155 |
|
156 |
fun occ_params ctxt ts = |
|
157 |
#1 (ProofContext.inferred_fixes ctxt) |
|
158 |
|> filter (member (op =) (fold (Variable.add_fixed ctxt) ts [])); |
|
159 |
||
|
23205
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
160 |
fun local_abbrev (x, rhs) = |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
161 |
Variable.add_fixes [x] #-> (fn [x'] => |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
162 |
let |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
163 |
val T = Term.fastype_of rhs; |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
164 |
val lhs = Free (x', T); |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
165 |
in |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
166 |
Variable.declare_term lhs #> |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
167 |
Variable.declare_term rhs #> |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
168 |
Assumption.add_assms (K (K (I, Envir.expand_term_frees [((x', T), rhs)]))) [] #> snd #> |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
169 |
pair (lhs, rhs) |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
170 |
end); |
|
b12f1c03cc9a
added target/local_abbrev (from proof_context.ML);
wenzelm
parents:
23087
diff
changeset
|
171 |
|
| 23087 | 172 |
fun abbrev is_loc some_class prmode ((raw_c, mx), raw_t) lthy = |
| 21747 | 173 |
let |
174 |
val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy); |
|
175 |
val target = LocalTheory.target_of lthy; |
|
176 |
val target_morphism = LocalTheory.target_morphism lthy; |
|
177 |
||
178 |
val c = Morphism.name target_morphism raw_c; |
|
179 |
val t = Morphism.term target_morphism raw_t; |
|
180 |
val xs = map Free (occ_params target [t]); |
|
181 |
val u = fold_rev Term.lambda xs t; |
|
| 21804 | 182 |
val U = Term.fastype_of u; |
| 21747 | 183 |
val u' = singleton (Variable.export_terms (Variable.declare_term u target) thy_ctxt) u; |
| 24218 | 184 |
val (mx1, mx2) = Class.fork_mixfix is_loc some_class mx; |
| 21747 | 185 |
in |
| 24838 | 186 |
lthy |
187 |
|> LocalTheory.theory_result |
|
188 |
(Sign.add_abbrev (#1 prmode) (ContextPosition.properties_of lthy) (c, u')) |
|
189 |
|-> (fn (lhs as Const (full_c, _), rhs) => LocalTheory.theory (Sign.add_notation prmode [(lhs, mx1)]) |
|
190 |
#> is_loc ? internal_abbrev prmode ((c, mx2), Term.list_comb (Const (full_c, U), xs)) |
|
191 |
#> local_abbrev (c, rhs)) |
|
| 21747 | 192 |
end; |
193 |
||
194 |
||
| 20894 | 195 |
(* defs *) |
196 |
||
| 21845 | 197 |
local |
| 21570 | 198 |
|
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
199 |
fun expand_term ctxt t = |
| 21570 | 200 |
let |
201 |
val thy = ProofContext.theory_of ctxt; |
|
202 |
val thy_ctxt = ProofContext.init thy; |
|
203 |
val ct = Thm.cterm_of thy t; |
|
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
204 |
val (defs, ct') = LocalDefs.export ctxt thy_ctxt (Drule.mk_term ct) ||> Drule.dest_term; |
| 21708 | 205 |
in (Thm.term_of ct', MetaSimplifier.rewrite true defs ct) end; |
| 21570 | 206 |
|
| 20894 | 207 |
fun add_def (name, prop) = |
208 |
Theory.add_defs_i false false [(name, prop)] #> |
|
209 |
(fn thy => (Drule.unvarify (Thm.get_axiom_i thy (Sign.full_name thy name)), thy)); |
|
210 |
||
|
24818
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
211 |
fun add_def_new (name, prop) thy = (* FIXME inactive --- breaks codegen *) |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
212 |
let |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
213 |
val tfrees = rev (map TFree (Term.add_tfrees prop [])); |
| 24848 | 214 |
val tfrees' = map (fn a => TFree (a, [])) (Name.invents Name.context Name.aT (length tfrees)); |
|
24818
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
215 |
val strip_sorts = tfrees ~~ tfrees'; |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
216 |
val recover_sorts = map (pairself (Thm.ctyp_of thy o Logic.varifyT)) (tfrees' ~~ tfrees); |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
217 |
|
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
218 |
val prop' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip_sorts))) prop; |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
219 |
val thy' = Theory.add_defs_i false false [(name, prop')] thy; |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
220 |
val axm = Thm.get_axiom_i thy' (Sign.full_name thy' name); |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
221 |
val def = Thm.instantiate (recover_sorts, []) axm; |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
222 |
in (Drule.unvarify def, thy') end; |
|
d07e56a9a0c2
added add_defs_new, which strips sorts for axioms (presently inactive);
wenzelm
parents:
24779
diff
changeset
|
223 |
|
| 20894 | 224 |
in |
225 |
||
| 24303 | 226 |
fun defs kind args lthy0 = |
| 20894 | 227 |
let |
|
21533
bada5ac1216a
simplified consts: no auxiliary params, sane mixfix syntax;
wenzelm
parents:
21498
diff
changeset
|
228 |
fun def ((c, mx), ((name, atts), rhs)) lthy1 = |
| 20894 | 229 |
let |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
230 |
val (rhs', rhs_conv) = expand_term lthy0 rhs; |
| 21570 | 231 |
val xs = Variable.add_fixed (LocalTheory.target_of lthy0) rhs' []; |
| 23087 | 232 |
val T = Term.fastype_of rhs; |
| 20894 | 233 |
|
| 21570 | 234 |
val ([(lhs, lhs_def)], lthy2) = lthy1 |
| 23087 | 235 |
|> LocalTheory.consts (member (op =) xs) [((c, T), mx)]; |
| 21570 | 236 |
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
|
237 |
|
| 21570 | 238 |
val name' = Thm.def_name_optional c name; |
239 |
val (def, lthy3) = lthy2 |
|
| 20894 | 240 |
|> LocalTheory.theory_result (add_def (name', Logic.mk_equals (lhs', rhs'))); |
| 21845 | 241 |
val eq = LocalDefs.trans_terms lthy3 |
242 |
[(*c == loc.c xs*) lhs_def, |
|
243 |
(*lhs' == rhs'*) def, |
|
244 |
(*rhs' == rhs*) Thm.symmetric rhs_conv]; |
|
| 23087 | 245 |
in ((lhs, ((name', atts), [([eq], [])])), lthy3) end; |
| 20894 | 246 |
|
| 21570 | 247 |
val ((lhss, facts), lthy') = lthy0 |> fold_map def args |>> split_list; |
| 21433 | 248 |
val (res, lthy'') = lthy' |> LocalTheory.notes kind facts; |
| 22424 | 249 |
|
| 20894 | 250 |
in (lhss ~~ map (apsnd the_single) res, lthy'') end; |
251 |
||
252 |
end; |
|
253 |
||
254 |
||
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
255 |
(* axioms *) |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
256 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
257 |
local |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
258 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
259 |
fun add_axiom hyps (name, prop) thy = |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
260 |
let |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
261 |
val name' = if name = "" then "axiom_" ^ serial_string () else name; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
262 |
val prop' = Logic.list_implies (hyps, prop); |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
263 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')]; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
264 |
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
|
265 |
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
|
266 |
in (Drule.implies_elim_list axm prems, thy') end; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
267 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
268 |
in |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
269 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
270 |
fun axioms kind specs lthy = |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
271 |
let |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
272 |
val hyps = map Thm.term_of (Assumption.assms_of lthy); |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
273 |
fun axiom ((name, atts), props) thy = thy |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
274 |
|> fold_map (add_axiom hyps) (PureThy.name_multi name props) |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
275 |
|-> (fn ths => pair ((name, atts), [(ths, [])])); |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
276 |
in |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
277 |
lthy |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
278 |
|> fold (fold Variable.declare_term o snd) specs |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
279 |
|> LocalTheory.theory_result (fold_map axiom specs) |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
280 |
|-> LocalTheory.notes kind |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
281 |
end; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
282 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
283 |
end; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
284 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
285 |
|
| 20894 | 286 |
(* notes *) |
287 |
||
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
288 |
fun import_export_proof ctxt (name, raw_th) = |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
289 |
let |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
290 |
val thy = ProofContext.theory_of ctxt; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
291 |
val thy_ctxt = ProofContext.init thy; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
292 |
val certT = Thm.ctyp_of thy; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
293 |
val cert = Thm.cterm_of thy; |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
294 |
|
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
295 |
(*export assumes/defines*) |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
296 |
val th = Goal.norm_result raw_th; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
297 |
val (defs, th') = LocalDefs.export ctxt thy_ctxt th; |
| 21708 | 298 |
val concl_conv = MetaSimplifier.rewrite true defs (Thm.cprop_of th); |
299 |
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
|
300 |
val nprems = Thm.nprems_of th' - Thm.nprems_of th; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
301 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
302 |
(*export fixes*) |
| 22692 | 303 |
val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []); |
304 |
val frees = map Free (Thm.fold_terms Term.add_frees th' []); |
|
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
305 |
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
|
306 |
|> Variable.export ctxt thy_ctxt |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
307 |
|> Drule.zero_var_indexes_list; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
308 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
309 |
(*thm definition*) |
| 21644 | 310 |
val result = th'' |
311 |
|> PureThy.name_thm true true "" |
|
312 |
|> Goal.close_result |
|
| 24779 | 313 |
|> fold PureThy.tag_rule (ContextPosition.properties_of ctxt) |
| 21644 | 314 |
|> PureThy.name_thm true true name; |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
315 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
316 |
(*import fixes*) |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
317 |
val (tvars, vars) = |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
318 |
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
|
319 |
|>> map Logic.dest_type; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
320 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
321 |
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
|
322 |
val inst = filter (is_Var o fst) (vars ~~ frees); |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
323 |
val cinstT = map (pairself certT o apfst TVar) instT; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
324 |
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
|
325 |
val result' = Thm.instantiate (cinstT, cinst) result; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
326 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
327 |
(*import assumes/defines*) |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
328 |
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
|
329 |
val result'' = |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
330 |
(case SINGLE (Seq.INTERVAL assm_tac 1 nprems) result' of |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
331 |
NONE => raise THM ("Failed to re-import result", 0, [result'])
|
| 21845 | 332 |
| SOME res => LocalDefs.trans_props ctxt [res, Thm.symmetric concl_conv]) |
| 21644 | 333 |
|> Goal.norm_result |
334 |
|> PureThy.name_thm false false name; |
|
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
335 |
|
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
336 |
in (result'', result) end; |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
337 |
|
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
338 |
fun notes loc kind facts lthy = |
| 21585 | 339 |
let |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
340 |
val is_loc = loc <> ""; |
|
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
341 |
val thy = ProofContext.theory_of lthy; |
|
21615
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
wenzelm
parents:
21613
diff
changeset
|
342 |
val full = LocalTheory.full_name lthy; |
| 21585 | 343 |
|
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
344 |
val facts' = facts |
|
21615
1bd558879c44
notes: proper naming of thm proof, activated import_export_proof;
wenzelm
parents:
21613
diff
changeset
|
345 |
|> 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
|
346 |
|> PureThy.map_facts (import_export_proof lthy); |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
347 |
val local_facts = PureThy.map_facts #1 facts' |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
348 |
|> Attrib.map_facts (Attrib.attribute_i thy); |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
349 |
val target_facts = PureThy.map_facts #1 facts' |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
350 |
|> 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
|
351 |
val global_facts = PureThy.map_facts #2 facts' |
|
21594
2859c94d67d4
reworked notes: towards proper import/export of proof terms;
wenzelm
parents:
21585
diff
changeset
|
352 |
|> 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
|
353 |
in |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
354 |
lthy |> LocalTheory.theory |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
355 |
(Sign.qualified_names |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
356 |
#> PureThy.note_thmss_i kind global_facts #> snd |
|
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
357 |
#> Sign.restore_naming thy) |
| 21585 | 358 |
|
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
359 |
|> 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
|
360 |
|
| 21433 | 361 |
|> ProofContext.set_stmt true |
| 20894 | 362 |
|> ProofContext.qualified_names |
|
21611
fc95ff1fe738
notes: proper import/export of proofs (still inactive);
wenzelm
parents:
21594
diff
changeset
|
363 |
|> ProofContext.note_thmss_i kind local_facts |
| 20894 | 364 |
||> ProofContext.restore_naming lthy |
| 24389 | 365 |
||> ProofContext.restore_mode lthy |
| 20894 | 366 |
end; |
367 |
||
368 |
||
369 |
(* init and exit *) |
|
370 |
||
| 22353 | 371 |
fun begin loc ctxt = |
372 |
let |
|
373 |
val thy = ProofContext.theory_of ctxt; |
|
374 |
val is_loc = loc <> ""; |
|
| 24218 | 375 |
val some_class = Class.class_of_locale thy loc; |
| 24748 | 376 |
fun class_init (SOME class) = |
377 |
Class.init [class] |
|
378 |
| class_init NONE = |
|
379 |
I; |
|
| 22353 | 380 |
in |
381 |
ctxt |
|
382 |
|> Data.put (if is_loc then SOME loc else NONE) |
|
| 24748 | 383 |
|> class_init some_class |
384 |
|> LocalTheory.init (NameSpace.base loc) |
|
| 21747 | 385 |
{pretty = pretty loc,
|
| 22353 | 386 |
consts = consts is_loc some_class, |
| 21747 | 387 |
axioms = axioms, |
| 22353 | 388 |
abbrev = abbrev is_loc some_class, |
| 24303 | 389 |
defs = defs, |
| 21747 | 390 |
notes = notes loc, |
391 |
type_syntax = type_syntax loc, |
|
392 |
term_syntax = term_syntax loc, |
|
393 |
declaration = declaration loc, |
|
394 |
target_morphism = target_morphism loc, |
|
| 22203 | 395 |
target_naming = target_naming loc, |
| 21804 | 396 |
reinit = fn _ => |
| 24657 | 397 |
(if is_loc then Locale.init loc else ProofContext.init) |
398 |
#> begin loc, |
|
| 24748 | 399 |
exit = LocalTheory.target_of } |
| 21747 | 400 |
end; |
| 20894 | 401 |
|
| 22353 | 402 |
fun init_i NONE thy = begin "" (ProofContext.init thy) |
403 |
| init_i (SOME loc) thy = begin loc (Locale.init loc thy); |
|
| 20894 | 404 |
|
| 21276 | 405 |
fun init (SOME "-") thy = init_i NONE thy |
406 |
| init loc thy = init_i (Option.map (Locale.intern thy) loc) thy; |
|
| 20894 | 407 |
|
408 |
end; |