author | blanchet |
Tue, 05 Apr 2011 10:37:21 +0200 | |
changeset 42233 | aab49f3cf802 |
parent 40844 | 5895c525739d |
child 42360 | da8817d01e7c |
permissions | -rw-r--r-- |
13402 | 1 |
(* Title: Pure/Proof/extraction.ML |
2 |
Author: Stefan Berghofer, TU Muenchen |
|
3 |
||
4 |
Extraction of programs from proofs. |
|
5 |
*) |
|
6 |
||
7 |
signature EXTRACTION = |
|
8 |
sig |
|
16458 | 9 |
val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory |
13402 | 10 |
val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory |
11 |
val add_realizes_eqns : string list -> theory -> theory |
|
12 |
val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory |
|
13 |
val add_typeof_eqns : string list -> theory -> theory |
|
14 |
val add_realizers_i : (string * (string list * term * Proofterm.proof)) list |
|
15 |
-> theory -> theory |
|
16 |
val add_realizers : (thm * (string list * string * string)) list |
|
17 |
-> theory -> theory |
|
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
18 |
val add_expand_thm : bool -> thm -> theory -> theory |
13732 | 19 |
val add_types : (xstring * ((term -> term option) list * |
20 |
(term -> typ -> term -> typ -> term) option)) list -> theory -> theory |
|
21 |
val extract : (thm * string list) list -> theory -> theory |
|
13402 | 22 |
val nullT : typ |
23 |
val nullt : term |
|
13714 | 24 |
val mk_typ : typ -> term |
25 |
val etype_of : theory -> string list -> typ list -> term -> typ |
|
26 |
val realizes_of: theory -> string list -> term -> term -> term |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
27 |
val abs_corr_shyps: theory -> thm -> string list -> term list -> Proofterm.proof -> Proofterm.proof |
13402 | 28 |
end; |
29 |
||
30 |
structure Extraction : EXTRACTION = |
|
31 |
struct |
|
32 |
||
33 |
(**** tools ****) |
|
34 |
||
35 |
fun add_syntax thy = |
|
36 |
thy |
|
37 |
|> Theory.copy |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24707
diff
changeset
|
38 |
|> Sign.root_path |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
29635
diff
changeset
|
39 |
|> Sign.add_types [(Binding.name "Type", 0, NoSyn), (Binding.name "Null", 0, NoSyn)] |
22796 | 40 |
|> Sign.add_consts |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
29635
diff
changeset
|
41 |
[(Binding.name "typeof", "'b::{} => Type", NoSyn), |
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
29635
diff
changeset
|
42 |
(Binding.name "Type", "'a::{} itself => Type", NoSyn), |
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
29635
diff
changeset
|
43 |
(Binding.name "Null", "Null", NoSyn), |
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
29635
diff
changeset
|
44 |
(Binding.name "realizes", "'a::{} => 'b::{} => 'b", NoSyn)]; |
13402 | 45 |
|
46 |
val nullT = Type ("Null", []); |
|
47 |
val nullt = Const ("Null", nullT); |
|
48 |
||
49 |
fun mk_typ T = |
|
19391 | 50 |
Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T; |
13402 | 51 |
|
52 |
fun typeof_proc defaultS vs (Const ("typeof", _) $ u) = |
|
15531 | 53 |
SOME (mk_typ (case strip_comb u of |
13402 | 54 |
(Var ((a, i), _), _) => |
20664 | 55 |
if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS) |
13402 | 56 |
else nullT |
57 |
| (Free (a, _), _) => |
|
20664 | 58 |
if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT |
13402 | 59 |
| _ => nullT)) |
15531 | 60 |
| typeof_proc _ _ _ = NONE; |
13402 | 61 |
|
15531 | 62 |
fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = SOME t |
13732 | 63 |
| rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) = |
64 |
(case strip_comb t of |
|
15531 | 65 |
(Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts)) |
66 |
| (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts)) |
|
67 |
| _ => NONE) |
|
68 |
| rlz_proc _ = NONE; |
|
13402 | 69 |
|
70 |
val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o |
|
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40132
diff
changeset
|
71 |
take_prefix (fn s => s <> ":") o raw_explode; |
13402 | 72 |
|
73 |
type rules = |
|
74 |
{next: int, rs: ((term * term) list * (term * term)) list, |
|
75 |
net: (int * ((term * term) list * (term * term))) Net.net}; |
|
76 |
||
77 |
val empty_rules : rules = {next = 0, rs = [], net = Net.empty}; |
|
78 |
||
33337 | 79 |
fun add_rule (r as (_, (lhs, _))) ({next, rs, net} : rules) = |
16800 | 80 |
{next = next - 1, rs = r :: rs, net = Net.insert_term (K false) |
18956 | 81 |
(Envir.eta_contract lhs, (next, r)) net}; |
13402 | 82 |
|
32784 | 83 |
fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) = |
33337 | 84 |
fold_rev add_rule (subtract (op =) rs1 rs2) {next = next, rs = rs1, net = net}; |
13402 | 85 |
|
16458 | 86 |
fun condrew thy rules procs = |
13402 | 87 |
let |
88 |
fun rew tm = |
|
17203 | 89 |
Pattern.rewrite_term thy [] (condrew' :: procs) tm |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
90 |
and condrew' tm = |
13402 | 91 |
let |
32738 | 92 |
val cache = Unsynchronized.ref ([] : (term * term) list); |
17232 | 93 |
fun lookup f x = (case AList.lookup (op =) (!cache) x of |
15531 | 94 |
NONE => |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
95 |
let val y = f x |
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
96 |
in (cache := (x, y) :: !cache; y) end |
15531 | 97 |
| SOME y => y); |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
98 |
in |
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
99 |
get_first (fn (_, (prems, (tm1, tm2))) => |
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
100 |
let |
19466 | 101 |
fun ren t = the_default t (Term.rename_abs tm1 tm t); |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
102 |
val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1); |
18184 | 103 |
val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty); |
32035 | 104 |
val prems' = map (pairself (Envir.subst_term env o inc o ren)) prems; |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
105 |
val env' = Envir.Envir |
32032 | 106 |
{maxidx = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u) prems' ~1, |
107 |
tenv = tenv, tyenv = Tenv}; |
|
18184 | 108 |
val env'' = fold (Pattern.unify thy o pairself (lookup rew)) prems' env'; |
15531 | 109 |
in SOME (Envir.norm_term env'' (inc (ren tm2))) |
110 |
end handle Pattern.MATCH => NONE | Pattern.Unif => NONE) |
|
16486 | 111 |
(sort (int_ord o pairself fst) |
18956 | 112 |
(Net.match_term rules (Envir.eta_contract tm))) |
15399
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
berghofe
parents:
14981
diff
changeset
|
113 |
end; |
13402 | 114 |
|
115 |
in rew end; |
|
116 |
||
37310 | 117 |
val chtype = Proofterm.change_type o SOME; |
13402 | 118 |
|
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30344
diff
changeset
|
119 |
fun extr_name s vs = Long_Name.append "extr" (space_implode "_" (s :: vs)); |
16195 | 120 |
fun corr_name s vs = extr_name s vs ^ "_correctness"; |
13402 | 121 |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39557
diff
changeset
|
122 |
fun msg d s = Output.urgent_message (Symbol.spaces d ^ s); |
13402 | 123 |
|
28812
413695e07bd4
Frees in PThms are now quantified in the order of their appearance in the
berghofe
parents:
28805
diff
changeset
|
124 |
fun vars_of t = map Var (rev (Term.add_vars t [])); |
413695e07bd4
Frees in PThms are now quantified in the order of their appearance in the
berghofe
parents:
28805
diff
changeset
|
125 |
fun frees_of t = map Free (rev (Term.add_frees t [])); |
413695e07bd4
Frees in PThms are now quantified in the order of their appearance in the
berghofe
parents:
28805
diff
changeset
|
126 |
fun vfs_of t = vars_of t @ frees_of t; |
13402 | 127 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
128 |
val mkabs = fold_rev (fn v => fn t => Abs ("x", fastype_of v, abstract_over (v, t))); |
13402 | 129 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
130 |
val mkabsp = fold_rev (fn t => fn prf => AbsP ("H", SOME t, prf)); |
13402 | 131 |
|
13732 | 132 |
fun strip_abs 0 t = t |
133 |
| strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t |
|
134 |
| strip_abs _ _ = error "strip_abs: not an abstraction"; |
|
135 |
||
37310 | 136 |
val prf_subst_TVars = Proofterm.map_proof_types o typ_subst_TVars; |
13402 | 137 |
|
40844 | 138 |
fun relevant_vars types prop = |
139 |
List.foldr |
|
140 |
(fn (Var ((a, _), T), vs) => |
|
141 |
(case body_type T of |
|
142 |
Type (s, _) => if member (op =) types s then a :: vs else vs |
|
143 |
| _ => vs) |
|
144 |
| (_, vs) => vs) [] (vars_of prop); |
|
13402 | 145 |
|
13732 | 146 |
fun tname_of (Type (s, _)) = s |
147 |
| tname_of _ = ""; |
|
148 |
||
149 |
fun get_var_type t = |
|
150 |
let |
|
16865 | 151 |
val vs = Term.add_vars t []; |
152 |
val fs = Term.add_frees t []; |
|
13732 | 153 |
in fn |
17232 | 154 |
Var (ixn, _) => (case AList.lookup (op =) vs ixn of |
15531 | 155 |
NONE => error "get_var_type: no such variable in term" |
156 |
| SOME T => Var (ixn, T)) |
|
17232 | 157 |
| Free (s, _) => (case AList.lookup (op =) fs s of |
15531 | 158 |
NONE => error "get_var_type: no such variable in term" |
159 |
| SOME T => Free (s, T)) |
|
13732 | 160 |
| _ => error "get_var_type: not a variable" |
161 |
end; |
|
162 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
163 |
fun read_term thy T s = |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
164 |
let |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
165 |
val ctxt = ProofContext.init_global thy |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
166 |
|> Proof_Syntax.strip_sorts_consttypes |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
167 |
|> ProofContext.set_defsort []; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
168 |
val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; |
39288 | 169 |
in parse ctxt s |> Type.constraint T |> Syntax.check_term ctxt end; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
170 |
|
13402 | 171 |
|
172 |
(**** theory data ****) |
|
173 |
||
22846 | 174 |
(* theory data *) |
13402 | 175 |
|
33522 | 176 |
structure ExtractionData = Theory_Data |
22846 | 177 |
( |
13402 | 178 |
type T = |
179 |
{realizes_eqns : rules, |
|
180 |
typeof_eqns : rules, |
|
13732 | 181 |
types : (string * ((term -> term option) list * |
182 |
(term -> typ -> term -> typ -> term) option)) list, |
|
13402 | 183 |
realizers : (string list * (term * proof)) list Symtab.table, |
184 |
defs : thm list, |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
185 |
expand : string list, |
16458 | 186 |
prep : (theory -> proof -> proof) option} |
13402 | 187 |
|
188 |
val empty = |
|
189 |
{realizes_eqns = empty_rules, |
|
190 |
typeof_eqns = empty_rules, |
|
191 |
types = [], |
|
192 |
realizers = Symtab.empty, |
|
193 |
defs = [], |
|
194 |
expand = [], |
|
15531 | 195 |
prep = NONE}; |
16458 | 196 |
val extend = I; |
13402 | 197 |
|
33522 | 198 |
fun merge |
199 |
({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1, |
|
13402 | 200 |
realizers = realizers1, defs = defs1, expand = expand1, prep = prep1}, |
201 |
{realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2, |
|
33522 | 202 |
realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T = |
13402 | 203 |
{realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2, |
204 |
typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2, |
|
22717 | 205 |
types = AList.merge (op =) (K true) (types1, types2), |
33038 | 206 |
realizers = Symtab.merge_list (eq_set (op =) o pairself #1) (realizers1, realizers2), |
22662 | 207 |
defs = Library.merge Thm.eq_thm (defs1, defs2), |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
208 |
expand = Library.merge (op =) (expand1, expand2), |
38761 | 209 |
prep = if is_some prep1 then prep1 else prep2}; |
22846 | 210 |
); |
13402 | 211 |
|
212 |
fun read_condeq thy = |
|
16458 | 213 |
let val thy' = add_syntax thy |
13402 | 214 |
in fn s => |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
215 |
let val t = Logic.varify_global (read_term thy' propT s) |
35424 | 216 |
in |
217 |
(map Logic.dest_equals (Logic.strip_imp_prems t), |
|
218 |
Logic.dest_equals (Logic.strip_imp_concl t)) |
|
219 |
handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s) |
|
220 |
end |
|
13402 | 221 |
end; |
222 |
||
223 |
(** preprocessor **) |
|
224 |
||
225 |
fun set_preprocessor prep thy = |
|
226 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
227 |
defs, expand, ...} = ExtractionData.get thy |
|
228 |
in |
|
229 |
ExtractionData.put |
|
230 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
15531 | 231 |
realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy |
13402 | 232 |
end; |
233 |
||
234 |
(** equations characterizing realizability **) |
|
235 |
||
236 |
fun gen_add_realizes_eqns prep_eq eqns thy = |
|
237 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
238 |
defs, expand, prep} = ExtractionData.get thy; |
|
239 |
in |
|
240 |
ExtractionData.put |
|
33337 | 241 |
{realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns, |
13402 | 242 |
typeof_eqns = typeof_eqns, types = types, realizers = realizers, |
243 |
defs = defs, expand = expand, prep = prep} thy |
|
244 |
end |
|
245 |
||
246 |
val add_realizes_eqns_i = gen_add_realizes_eqns (K I); |
|
247 |
val add_realizes_eqns = gen_add_realizes_eqns read_condeq; |
|
248 |
||
249 |
(** equations characterizing type of extracted program **) |
|
250 |
||
251 |
fun gen_add_typeof_eqns prep_eq eqns thy = |
|
252 |
let |
|
253 |
val {realizes_eqns, typeof_eqns, types, realizers, |
|
254 |
defs, expand, prep} = ExtractionData.get thy; |
|
13732 | 255 |
val eqns' = map (prep_eq thy) eqns |
13402 | 256 |
in |
257 |
ExtractionData.put |
|
258 |
{realizes_eqns = realizes_eqns, realizers = realizers, |
|
33337 | 259 |
typeof_eqns = fold_rev add_rule eqns' typeof_eqns, |
13732 | 260 |
types = types, defs = defs, expand = expand, prep = prep} thy |
13402 | 261 |
end |
262 |
||
263 |
val add_typeof_eqns_i = gen_add_typeof_eqns (K I); |
|
264 |
val add_typeof_eqns = gen_add_typeof_eqns read_condeq; |
|
265 |
||
266 |
fun thaw (T as TFree (a, S)) = |
|
28375 | 267 |
if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T |
13402 | 268 |
| thaw (Type (a, Ts)) = Type (a, map thaw Ts) |
269 |
| thaw T = T; |
|
270 |
||
271 |
fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S) |
|
272 |
| freeze (Type (a, Ts)) = Type (a, map freeze Ts) |
|
273 |
| freeze T = T; |
|
274 |
||
275 |
fun freeze_thaw f x = |
|
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
19482
diff
changeset
|
276 |
map_types thaw (f (map_types freeze x)); |
13402 | 277 |
|
16458 | 278 |
fun etype_of thy vs Ts t = |
13402 | 279 |
let |
16458 | 280 |
val {typeof_eqns, ...} = ExtractionData.get thy; |
13402 | 281 |
fun err () = error ("Unable to determine type of extracted program for\n" ^ |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26653
diff
changeset
|
282 |
Syntax.string_of_term_global thy t) |
16458 | 283 |
in case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns) |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
284 |
[typeof_proc [] vs]) (list_abs (map (pair "x") (rev Ts), |
13402 | 285 |
Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of |
286 |
Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ()) |
|
287 |
| _ => err () |
|
288 |
end; |
|
289 |
||
290 |
(** realizers for axioms / theorems, together with correctness proofs **) |
|
291 |
||
292 |
fun gen_add_realizers prep_rlz rs thy = |
|
293 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
294 |
defs, expand, prep} = ExtractionData.get thy |
|
295 |
in |
|
296 |
ExtractionData.put |
|
297 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
25389 | 298 |
realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers, |
13402 | 299 |
defs = defs, expand = expand, prep = prep} thy |
300 |
end |
|
301 |
||
302 |
fun prep_realizer thy = |
|
303 |
let |
|
13732 | 304 |
val {realizes_eqns, typeof_eqns, defs, types, ...} = |
13402 | 305 |
ExtractionData.get thy; |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
306 |
val procs = maps (fst o snd) types; |
13732 | 307 |
val rtypes = map fst types; |
16800 | 308 |
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
13402 | 309 |
val thy' = add_syntax thy; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
310 |
val rd = Proof_Syntax.read_proof thy' true false; |
13402 | 311 |
in fn (thm, (vs, s1, s2)) => |
312 |
let |
|
36744
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents:
36620
diff
changeset
|
313 |
val name = Thm.derivation_name thm; |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
314 |
val _ = name <> "" orelse error "add_realizers: unnamed theorem"; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
315 |
val prop = Thm.unconstrainT thm |> prop_of |> |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
316 |
Pattern.rewrite_term thy' (map (Logic.dest_equals o prop_of) defs) []; |
13402 | 317 |
val vars = vars_of prop; |
13732 | 318 |
val vars' = filter_out (fn v => |
20664 | 319 |
member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
320 |
val shyps = maps (fn Var ((x, i), _) => |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
321 |
if member (op =) vs x then Logic.mk_of_sort |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
322 |
(TVar (("'" ^ x, i), []), Sign.defaultS thy') |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
323 |
else []) vars; |
16458 | 324 |
val T = etype_of thy' vs [] prop; |
33832 | 325 |
val (T', thw) = Type.legacy_freeze_thaw_type |
13732 | 326 |
(if T = nullT then nullT else map fastype_of vars' ---> T); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
327 |
val t = map_types thw (read_term thy' T' s1); |
16458 | 328 |
val r' = freeze_thaw (condrew thy' eqns |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
329 |
(procs @ [typeof_proc [] vs, rlz_proc])) |
13402 | 330 |
(Const ("realizes", T --> propT --> propT) $ |
13732 | 331 |
(if T = nullT then t else list_comb (t, vars')) $ prop); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
332 |
val r = Logic.list_implies (shyps, |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
333 |
fold_rev Logic.all (map (get_var_type r') vars) r'); |
16458 | 334 |
val prf = Reconstruct.reconstruct_proof thy' r (rd s2); |
13402 | 335 |
in (name, (vs, (t, prf))) end |
336 |
end; |
|
337 |
||
338 |
val add_realizers_i = gen_add_realizers |
|
339 |
(fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf)))); |
|
340 |
val add_realizers = gen_add_realizers prep_realizer; |
|
341 |
||
13714 | 342 |
fun realizes_of thy vs t prop = |
343 |
let |
|
344 |
val thy' = add_syntax thy; |
|
13732 | 345 |
val {realizes_eqns, typeof_eqns, defs, types, ...} = |
13714 | 346 |
ExtractionData.get thy'; |
22717 | 347 |
val procs = maps (rev o fst o snd) types; |
16800 | 348 |
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
17203 | 349 |
val prop' = Pattern.rewrite_term thy' |
13714 | 350 |
(map (Logic.dest_equals o prop_of) defs) [] prop; |
16458 | 351 |
in freeze_thaw (condrew thy' eqns |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
352 |
(procs @ [typeof_proc [] vs, rlz_proc])) |
13714 | 353 |
(Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop') |
354 |
end; |
|
355 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
356 |
fun abs_corr_shyps thy thm vs xs prf = |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
357 |
let |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
358 |
val S = Sign.defaultS thy; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
359 |
val ((atyp_map, constraints, _), prop') = |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
360 |
Logic.unconstrainT (#shyps (rep_thm thm)) (prop_of thm); |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
361 |
val atyps = fold_types (fold_atyps (insert (op =))) (prop_of thm) []; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
362 |
val Ts = map_filter (fn ((v, i), _) => if member (op =) vs v then |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
363 |
SOME (TVar (("'" ^ v, i), [])) else NONE) |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
364 |
(rev (Term.add_vars prop' [])); |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
365 |
val cs = maps (fn T => map (pair T) S) Ts; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
366 |
val constraints' = map Logic.mk_of_class cs; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
367 |
val cs' = rev (cs @ map (Logic.dest_of_class o snd) constraints); |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
368 |
fun typ_map T = Type.strip_sorts |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
369 |
(map_atyps (fn U => if member (op =) atyps U then atyp_map U else U) T); |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
370 |
fun mk_hyp (T, c) = Hyp (Logic.mk_of_class (typ_map T, c)); |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
371 |
val xs' = map (map_types typ_map) xs |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
372 |
in |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
373 |
prf |> |
37310 | 374 |
Same.commit (Proofterm.map_proof_same (map_types typ_map) typ_map mk_hyp) |> |
375 |
fold_rev Proofterm.implies_intr_proof' (map snd constraints) |> |
|
376 |
fold_rev Proofterm.forall_intr_proof' xs' |> |
|
377 |
fold_rev Proofterm.implies_intr_proof' constraints' |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
378 |
end; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
379 |
|
13402 | 380 |
(** expanding theorems / definitions **) |
381 |
||
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
382 |
fun add_expand_thm is_def thm thy = |
13402 | 383 |
let |
384 |
val {realizes_eqns, typeof_eqns, types, realizers, |
|
385 |
defs, expand, prep} = ExtractionData.get thy; |
|
386 |
||
36744
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents:
36620
diff
changeset
|
387 |
val name = Thm.derivation_name thm; |
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
388 |
val _ = name <> "" orelse error "add_expand_thm: unnamed theorem"; |
13402 | 389 |
in |
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
390 |
thy |> ExtractionData.put |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
391 |
(if is_def then |
13402 | 392 |
{realizes_eqns = realizes_eqns, |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
393 |
typeof_eqns = add_rule ([], Logic.dest_equals (map_types |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
394 |
Type.strip_sorts (prop_of (Drule.abs_def thm)))) typeof_eqns, |
13402 | 395 |
types = types, |
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
21858
diff
changeset
|
396 |
realizers = realizers, defs = insert Thm.eq_thm thm defs, |
13402 | 397 |
expand = expand, prep = prep} |
398 |
else |
|
399 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
400 |
realizers = realizers, defs = defs, |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
401 |
expand = insert (op =) name expand, prep = prep}) |
13402 | 402 |
end; |
403 |
||
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
404 |
fun extraction_expand is_def = |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
405 |
Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm is_def th) I); |
13402 | 406 |
|
15801 | 407 |
|
13732 | 408 |
(** types with computational content **) |
409 |
||
410 |
fun add_types tys thy = |
|
22717 | 411 |
ExtractionData.map |
412 |
(fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} => |
|
13732 | 413 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, |
22717 | 414 |
types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types, |
415 |
realizers = realizers, defs = defs, expand = expand, prep = prep}) |
|
416 |
thy; |
|
13732 | 417 |
|
13402 | 418 |
|
15801 | 419 |
(** Pure setup **) |
420 |
||
26463 | 421 |
val _ = Context.>> (Context.map_theory |
18708 | 422 |
(add_types [("prop", ([], NONE))] #> |
15801 | 423 |
|
424 |
add_typeof_eqns |
|
425 |
["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \ |
|
426 |
\ (typeof (PROP Q)) == (Type (TYPE('Q))) ==> \ |
|
427 |
\ (typeof (PROP P ==> PROP Q)) == (Type (TYPE('Q)))", |
|
428 |
||
429 |
"(typeof (PROP Q)) == (Type (TYPE(Null))) ==> \ |
|
430 |
\ (typeof (PROP P ==> PROP Q)) == (Type (TYPE(Null)))", |
|
431 |
||
432 |
"(typeof (PROP P)) == (Type (TYPE('P))) ==> \ |
|
433 |
\ (typeof (PROP Q)) == (Type (TYPE('Q))) ==> \ |
|
434 |
\ (typeof (PROP P ==> PROP Q)) == (Type (TYPE('P => 'Q)))", |
|
435 |
||
436 |
"(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \ |
|
437 |
\ (typeof (!!x. PROP P (x))) == (Type (TYPE(Null)))", |
|
438 |
||
439 |
"(%x. typeof (PROP P (x))) == (%x. Type (TYPE('P))) ==> \ |
|
440 |
\ (typeof (!!x::'a. PROP P (x))) == (Type (TYPE('a => 'P)))", |
|
441 |
||
442 |
"(%x. typeof (f (x))) == (%x. Type (TYPE('f))) ==> \ |
|
18708 | 443 |
\ (typeof (f)) == (Type (TYPE('f)))"] #> |
15801 | 444 |
|
445 |
add_realizes_eqns |
|
446 |
["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \ |
|
447 |
\ (realizes (r) (PROP P ==> PROP Q)) == \ |
|
448 |
\ (PROP realizes (Null) (PROP P) ==> PROP realizes (r) (PROP Q))", |
|
449 |
||
450 |
"(typeof (PROP P)) == (Type (TYPE('P))) ==> \ |
|
451 |
\ (typeof (PROP Q)) == (Type (TYPE(Null))) ==> \ |
|
452 |
\ (realizes (r) (PROP P ==> PROP Q)) == \ |
|
453 |
\ (!!x::'P. PROP realizes (x) (PROP P) ==> PROP realizes (Null) (PROP Q))", |
|
454 |
||
455 |
"(realizes (r) (PROP P ==> PROP Q)) == \ |
|
456 |
\ (!!x. PROP realizes (x) (PROP P) ==> PROP realizes (r (x)) (PROP Q))", |
|
457 |
||
458 |
"(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \ |
|
459 |
\ (realizes (r) (!!x. PROP P (x))) == \ |
|
460 |
\ (!!x. PROP realizes (Null) (PROP P (x)))", |
|
461 |
||
462 |
"(realizes (r) (!!x. PROP P (x))) == \ |
|
18708 | 463 |
\ (!!x. PROP realizes (r (x)) (PROP P (x)))"] #> |
15801 | 464 |
|
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
465 |
Attrib.setup (Binding.name "extraction_expand") (Scan.succeed (extraction_expand false)) |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
466 |
"specify theorems to be expanded during extraction" #> |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
467 |
Attrib.setup (Binding.name "extraction_expand_def") (Scan.succeed (extraction_expand true)) |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
468 |
"specify definitions to be expanded during extraction")); |
15801 | 469 |
|
470 |
||
13402 | 471 |
(**** extract program ****) |
472 |
||
473 |
val dummyt = Const ("dummy", dummyT); |
|
474 |
||
475 |
fun extract thms thy = |
|
476 |
let |
|
16458 | 477 |
val thy' = add_syntax thy; |
13402 | 478 |
val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = |
479 |
ExtractionData.get thy; |
|
22717 | 480 |
val procs = maps (rev o fst o snd) types; |
13732 | 481 |
val rtypes = map fst types; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
482 |
val typroc = typeof_proc []; |
19466 | 483 |
val prep = the_default (K I) prep thy' o ProofRewriteRules.elim_defs thy' false defs o |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
484 |
Reconstruct.expand_proof thy' (map (rpair NONE) ("" :: expand)); |
16800 | 485 |
val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
13402 | 486 |
|
487 |
fun find_inst prop Ts ts vs = |
|
488 |
let |
|
13732 | 489 |
val rvs = relevant_vars rtypes prop; |
13402 | 490 |
val vars = vars_of prop; |
491 |
val n = Int.min (length vars, length ts); |
|
492 |
||
33337 | 493 |
fun add_args (Var ((a, i), _), t) (vs', tye) = |
20664 | 494 |
if member (op =) rvs a then |
16458 | 495 |
let val T = etype_of thy' vs Ts t |
13402 | 496 |
in if T = nullT then (vs', tye) |
497 |
else (a :: vs', (("'" ^ a, i), T) :: tye) |
|
498 |
end |
|
499 |
else (vs', tye) |
|
500 |
||
33957 | 501 |
in fold_rev add_args (take n vars ~~ take n ts) ([], []) end; |
13402 | 502 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
503 |
fun mk_shyps tye = maps (fn (ixn, _) => |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
504 |
Logic.mk_of_sort (TVar (ixn, []), Sign.defaultS thy)) tye; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
505 |
|
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
506 |
fun mk_sprfs cs tye = maps (fn (_, T) => |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
507 |
ProofRewriteRules.mk_of_sort_proof thy (map SOME cs) |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
508 |
(T, Sign.defaultS thy)) tye; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
509 |
|
33038 | 510 |
fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); |
28375 | 511 |
fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE); |
13402 | 512 |
|
13732 | 513 |
fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw |
16458 | 514 |
(condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (list_abs |
13732 | 515 |
(map (pair "x") (rev Ts), t))); |
516 |
||
517 |
fun realizes_null vs prop = app_rlz_rews [] vs |
|
518 |
(Const ("realizes", nullT --> propT --> propT) $ nullt $ prop); |
|
13402 | 519 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
520 |
fun corr d defs vs ts Ts hs cs (PBound i) _ _ = (defs, PBound i) |
13402 | 521 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
522 |
| corr d defs vs ts Ts hs cs (Abst (s, SOME T, prf)) (Abst (_, _, prf')) t = |
13402 | 523 |
let val (defs', corr_prf) = corr d defs vs [] (T :: Ts) |
37310 | 524 |
(dummyt :: hs) cs prf (Proofterm.incr_pboundvars 1 0 prf') |
15531 | 525 |
(case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE) |
526 |
in (defs', Abst (s, SOME T, corr_prf)) end |
|
13402 | 527 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
528 |
| corr d defs vs ts Ts hs cs (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) t = |
13402 | 529 |
let |
16458 | 530 |
val T = etype_of thy' vs Ts prop; |
13402 | 531 |
val u = if T = nullT then |
15531 | 532 |
(case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE) |
533 |
else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE); |
|
37310 | 534 |
val (defs', corr_prf) = |
535 |
corr d defs vs [] (T :: Ts) (prop :: hs) |
|
536 |
(prop :: cs) (Proofterm.incr_pboundvars 0 1 prf) |
|
537 |
(Proofterm.incr_pboundvars 0 1 prf') u; |
|
13402 | 538 |
val rlz = Const ("realizes", T --> propT --> propT) |
539 |
in (defs', |
|
13732 | 540 |
if T = nullT then AbsP ("R", |
15531 | 541 |
SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)), |
37310 | 542 |
Proofterm.prf_subst_bounds [nullt] corr_prf) |
15531 | 543 |
else Abst (s, SOME T, AbsP ("R", |
544 |
SOME (app_rlz_rews (T :: Ts) vs |
|
13732 | 545 |
(rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf))) |
13402 | 546 |
end |
547 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
548 |
| corr d defs vs ts Ts hs cs (prf % SOME t) (prf' % _) t' = |
13732 | 549 |
let |
550 |
val (Us, T) = strip_type (fastype_of1 (Ts, t)); |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
551 |
val (defs', corr_prf) = corr d defs vs (t :: ts) Ts hs cs prf prf' |
20664 | 552 |
(if member (op =) rtypes (tname_of T) then t' |
15531 | 553 |
else (case t' of SOME (u $ _) => SOME u | _ => NONE)); |
20664 | 554 |
val u = if not (member (op =) rtypes (tname_of T)) then t else |
13732 | 555 |
let |
16458 | 556 |
val eT = etype_of thy' vs Ts t; |
13732 | 557 |
val (r, Us') = if eT = nullT then (nullt, Us) else |
558 |
(Bound (length Us), eT :: Us); |
|
559 |
val u = list_comb (incr_boundvars (length Us') t, |
|
560 |
map Bound (length Us - 1 downto 0)); |
|
17271 | 561 |
val u' = (case AList.lookup (op =) types (tname_of T) of |
15531 | 562 |
SOME ((_, SOME f)) => f r eT u T |
13732 | 563 |
| _ => Const ("realizes", eT --> T --> T) $ r $ u) |
564 |
in app_rlz_rews Ts vs (list_abs (map (pair "x") Us', u')) end |
|
15531 | 565 |
in (defs', corr_prf % SOME u) end |
13402 | 566 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
567 |
| corr d defs vs ts Ts hs cs (prf1 %% prf2) (prf1' %% prf2') t = |
13402 | 568 |
let |
569 |
val prop = Reconstruct.prop_of' hs prf2'; |
|
16458 | 570 |
val T = etype_of thy' vs Ts prop; |
15531 | 571 |
val (defs1, f, u) = if T = nullT then (defs, t, NONE) else |
13402 | 572 |
(case t of |
15531 | 573 |
SOME (f $ u) => (defs, SOME f, SOME u) |
13402 | 574 |
| _ => |
575 |
let val (defs1, u) = extr d defs vs [] Ts hs prf2' |
|
15531 | 576 |
in (defs1, NONE, SOME u) end) |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
577 |
val (defs2, corr_prf1) = corr d defs1 vs [] Ts hs cs prf1 prf1' f; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
578 |
val (defs3, corr_prf2) = corr d defs2 vs [] Ts hs cs prf2 prf2' u; |
13402 | 579 |
in |
580 |
if T = nullT then (defs3, corr_prf1 %% corr_prf2) else |
|
581 |
(defs3, corr_prf1 % u %% corr_prf2) |
|
582 |
end |
|
583 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
584 |
| corr d defs vs ts Ts hs cs (prf0 as PThm (_, ((name, prop, SOME Ts'), body))) _ _ = |
13402 | 585 |
let |
37310 | 586 |
val prf = Proofterm.join_proof body; |
13402 | 587 |
val (vs', tye) = find_inst prop Ts ts vs; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
588 |
val shyps = mk_shyps tye; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
589 |
val sprfs = mk_sprfs cs tye; |
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
590 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye; |
16458 | 591 |
val T = etype_of thy' vs' [] prop; |
13402 | 592 |
val defs' = if T = nullT then defs |
593 |
else fst (extr d defs vs ts Ts hs prf0) |
|
594 |
in |
|
13609
73c3915553b4
Added check for axioms with "realizes Null A = A".
berghofe
parents:
13417
diff
changeset
|
595 |
if T = nullT andalso realizes_null vs' prop aconv prop then (defs, prf0) |
17412 | 596 |
else case Symtab.lookup realizers name of |
15531 | 597 |
NONE => (case find vs' (find' name defs') of |
598 |
NONE => |
|
13402 | 599 |
let |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
600 |
val _ = T = nullT orelse error "corr: internal error"; |
13402 | 601 |
val _ = msg d ("Building correctness proof for " ^ quote name ^ |
602 |
(if null vs' then "" |
|
603 |
else " (relevant variables: " ^ commas_quote vs' ^ ")")); |
|
16458 | 604 |
val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
605 |
val (defs'', corr_prf0) = corr (d + 1) defs' vs' [] [] [] |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
606 |
(rev shyps) prf' prf' NONE; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
607 |
val corr_prf = mkabsp shyps corr_prf0; |
13732 | 608 |
val corr_prop = Reconstruct.prop_of corr_prf; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
609 |
val corr_prf' = |
37310 | 610 |
Proofterm.proof_combP (Proofterm.proof_combt |
28805 | 611 |
(PThm (serial (), |
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
612 |
((corr_name name vs', corr_prop, SOME (map TVar (Term.add_tvars corr_prop [] |> rev))), |
37310 | 613 |
Future.value (Proofterm.approximate_proof_body corr_prf))), |
614 |
vfs_of corr_prop), |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
615 |
map PBound (length shyps - 1 downto 0)) |> |
37310 | 616 |
fold_rev Proofterm.forall_intr_proof' |
617 |
(map (get_var_type corr_prop) (vfs_of prop)) |> |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
618 |
mkabsp shyps |
13402 | 619 |
in |
13732 | 620 |
((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'', |
37310 | 621 |
Proofterm.proof_combP (prf_subst_TVars tye' corr_prf', sprfs)) |
13402 | 622 |
end |
37310 | 623 |
| SOME (_, (_, prf')) => |
624 |
(defs', Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs))) |
|
15531 | 625 |
| SOME rs => (case find vs' rs of |
37310 | 626 |
SOME (_, prf') => (defs', Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs)) |
15531 | 627 |
| NONE => error ("corr: no realizer for instance of theorem " ^ |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26653
diff
changeset
|
628 |
quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm |
37310 | 629 |
(Reconstruct.prop_of (Proofterm.proof_combt (prf0, ts)))))) |
13402 | 630 |
end |
631 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
632 |
| corr d defs vs ts Ts hs cs (prf0 as PAxm (s, prop, SOME Ts')) _ _ = |
13402 | 633 |
let |
634 |
val (vs', tye) = find_inst prop Ts ts vs; |
|
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
635 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 636 |
in |
16458 | 637 |
if etype_of thy' vs' [] prop = nullT andalso |
13609
73c3915553b4
Added check for axioms with "realizes Null A = A".
berghofe
parents:
13417
diff
changeset
|
638 |
realizes_null vs' prop aconv prop then (defs, prf0) |
18956 | 639 |
else case find vs' (Symtab.lookup_list realizers s) of |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
640 |
SOME (_, prf) => (defs, |
37310 | 641 |
Proofterm.proof_combP (prf_subst_TVars tye' prf, mk_sprfs cs tye)) |
15531 | 642 |
| NONE => error ("corr: no realizer for instance of axiom " ^ |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26653
diff
changeset
|
643 |
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm |
37310 | 644 |
(Reconstruct.prop_of (Proofterm.proof_combt (prf0, ts))))) |
13402 | 645 |
end |
646 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
647 |
| corr d defs vs ts Ts hs _ _ _ _ = error "corr: bad proof" |
13402 | 648 |
|
649 |
and extr d defs vs ts Ts hs (PBound i) = (defs, Bound i) |
|
650 |
||
15531 | 651 |
| extr d defs vs ts Ts hs (Abst (s, SOME T, prf)) = |
13402 | 652 |
let val (defs', t) = extr d defs vs [] |
37310 | 653 |
(T :: Ts) (dummyt :: hs) (Proofterm.incr_pboundvars 1 0 prf) |
13402 | 654 |
in (defs', Abs (s, T, t)) end |
655 |
||
15531 | 656 |
| extr d defs vs ts Ts hs (AbsP (s, SOME t, prf)) = |
13402 | 657 |
let |
16458 | 658 |
val T = etype_of thy' vs Ts t; |
37310 | 659 |
val (defs', t) = |
660 |
extr d defs vs [] (T :: Ts) (t :: hs) (Proofterm.incr_pboundvars 0 1 prf) |
|
13402 | 661 |
in (defs', |
662 |
if T = nullT then subst_bound (nullt, t) else Abs (s, T, t)) |
|
663 |
end |
|
664 |
||
15531 | 665 |
| extr d defs vs ts Ts hs (prf % SOME t) = |
13402 | 666 |
let val (defs', u) = extr d defs vs (t :: ts) Ts hs prf |
13732 | 667 |
in (defs', |
20664 | 668 |
if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u |
13732 | 669 |
else u $ t) |
670 |
end |
|
13402 | 671 |
|
672 |
| extr d defs vs ts Ts hs (prf1 %% prf2) = |
|
673 |
let |
|
674 |
val (defs', f) = extr d defs vs [] Ts hs prf1; |
|
675 |
val prop = Reconstruct.prop_of' hs prf2; |
|
16458 | 676 |
val T = etype_of thy' vs Ts prop |
13402 | 677 |
in |
678 |
if T = nullT then (defs', f) else |
|
679 |
let val (defs'', t) = extr d defs' vs [] Ts hs prf2 |
|
680 |
in (defs'', f $ t) end |
|
681 |
end |
|
682 |
||
28805 | 683 |
| extr d defs vs ts Ts hs (prf0 as PThm (_, ((s, prop, SOME Ts'), body))) = |
13402 | 684 |
let |
37310 | 685 |
val prf = Proofterm.join_proof body; |
13402 | 686 |
val (vs', tye) = find_inst prop Ts ts vs; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
687 |
val shyps = mk_shyps tye; |
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
688 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 689 |
in |
17412 | 690 |
case Symtab.lookup realizers s of |
15531 | 691 |
NONE => (case find vs' (find' s defs) of |
692 |
NONE => |
|
13402 | 693 |
let |
694 |
val _ = msg d ("Extracting " ^ quote s ^ |
|
695 |
(if null vs' then "" |
|
696 |
else " (relevant variables: " ^ commas_quote vs' ^ ")")); |
|
16458 | 697 |
val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf); |
13402 | 698 |
val (defs', t) = extr (d + 1) defs vs' [] [] [] prf'; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
699 |
val (defs'', corr_prf) = corr (d + 1) defs' vs' [] [] [] |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
700 |
(rev shyps) prf' prf' (SOME t); |
13402 | 701 |
|
702 |
val nt = Envir.beta_norm t; |
|
20664 | 703 |
val args = filter_out (fn v => member (op =) rtypes |
704 |
(tname_of (body_type (fastype_of v)))) (vfs_of prop); |
|
33317 | 705 |
val args' = filter (fn v => Logic.occs (v, nt)) args; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
706 |
val t' = mkabs args' nt; |
13402 | 707 |
val T = fastype_of t'; |
13732 | 708 |
val cname = extr_name s vs'; |
13402 | 709 |
val c = Const (cname, T); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
710 |
val u = mkabs args (list_comb (c, args')); |
13402 | 711 |
val eqn = Logic.mk_equals (c, t'); |
712 |
val rlz = |
|
713 |
Const ("realizes", fastype_of nt --> propT --> propT); |
|
13732 | 714 |
val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop); |
715 |
val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop); |
|
716 |
val f = app_rlz_rews [] vs' |
|
717 |
(Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop)); |
|
13402 | 718 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
719 |
val corr_prf' = mkabsp shyps |
37310 | 720 |
(chtype [] Proofterm.equal_elim_axm %> lhs %> rhs %% |
721 |
(chtype [propT] Proofterm.symmetric_axm %> rhs %> lhs %% |
|
722 |
(chtype [T, propT] Proofterm.combination_axm %> f %> f %> c %> t' %% |
|
723 |
(chtype [T --> propT] Proofterm.reflexive_axm %> f) %% |
|
13732 | 724 |
PAxm (cname ^ "_def", eqn, |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
725 |
SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf); |
13732 | 726 |
val corr_prop = Reconstruct.prop_of corr_prf'; |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
727 |
val corr_prf'' = |
37310 | 728 |
Proofterm.proof_combP (Proofterm.proof_combt |
28805 | 729 |
(PThm (serial (), |
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
730 |
((corr_name s vs', corr_prop, SOME (map TVar (Term.add_tvars corr_prop [] |> rev))), |
37310 | 731 |
Future.value (Proofterm.approximate_proof_body corr_prf'))), |
732 |
vfs_of corr_prop), |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
733 |
map PBound (length shyps - 1 downto 0)) |> |
37310 | 734 |
fold_rev Proofterm.forall_intr_proof' |
735 |
(map (get_var_type corr_prop) (vfs_of prop)) |> |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
736 |
mkabsp shyps |
13402 | 737 |
in |
13732 | 738 |
((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'', |
13402 | 739 |
subst_TVars tye' u) |
740 |
end |
|
15531 | 741 |
| SOME ((_, u), _) => (defs, subst_TVars tye' u)) |
742 |
| SOME rs => (case find vs' rs of |
|
743 |
SOME (t, _) => (defs, subst_TVars tye' t) |
|
744 |
| NONE => error ("extr: no realizer for instance of theorem " ^ |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26653
diff
changeset
|
745 |
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm |
37310 | 746 |
(Reconstruct.prop_of (Proofterm.proof_combt (prf0, ts)))))) |
13402 | 747 |
end |
748 |
||
15531 | 749 |
| extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) = |
13402 | 750 |
let |
751 |
val (vs', tye) = find_inst prop Ts ts vs; |
|
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
35985
diff
changeset
|
752 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 753 |
in |
18956 | 754 |
case find vs' (Symtab.lookup_list realizers s) of |
15531 | 755 |
SOME (t, _) => (defs, subst_TVars tye' t) |
756 |
| NONE => error ("extr: no realizer for instance of axiom " ^ |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26653
diff
changeset
|
757 |
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm |
37310 | 758 |
(Reconstruct.prop_of (Proofterm.proof_combt (prf0, ts))))) |
13402 | 759 |
end |
760 |
||
761 |
| extr d defs vs ts Ts hs _ = error "extr: bad proof"; |
|
762 |
||
13732 | 763 |
fun prep_thm (thm, vs) = |
13402 | 764 |
let |
26626
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26481
diff
changeset
|
765 |
val thy = Thm.theory_of_thm thm; |
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26481
diff
changeset
|
766 |
val prop = Thm.prop_of thm; |
28814 | 767 |
val prf = Thm.proof_of thm; |
36744
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents:
36620
diff
changeset
|
768 |
val name = Thm.derivation_name thm; |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
769 |
val _ = name <> "" orelse error "extraction: unnamed theorem"; |
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
770 |
val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^ |
13402 | 771 |
quote name ^ " has no computational content") |
22596 | 772 |
in (Reconstruct.reconstruct_proof thy prop prf, vs) end; |
13402 | 773 |
|
33245 | 774 |
val defs = |
775 |
fold (fn (prf, vs) => fn defs => fst (extr 0 defs vs [] [] [] prf)) |
|
776 |
(map prep_thm thms) []; |
|
13402 | 777 |
|
16149 | 778 |
fun add_def (s, (vs, ((t, u), (prf, _)))) thy = |
16458 | 779 |
(case Sign.const_type thy (extr_name s vs) of |
15531 | 780 |
NONE => |
13732 | 781 |
let |
782 |
val corr_prop = Reconstruct.prop_of prf; |
|
33832 | 783 |
val ft = Type.legacy_freeze t; |
784 |
val fu = Type.legacy_freeze u; |
|
22750 | 785 |
val (def_thms, thy') = if t = nullt then ([], thy) else |
786 |
thy |
|
30435
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents:
30364
diff
changeset
|
787 |
|> Sign.add_consts_i [(Binding.qualified_name (extr_name s vs), fastype_of ft, NoSyn)] |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
788 |
|> Global_Theory.add_defs false [((Binding.qualified_name (extr_name s vs ^ "_def"), |
22750 | 789 |
Logic.mk_equals (head_of (strip_abs_body fu), ft)), [])] |
13732 | 790 |
in |
22750 | 791 |
thy' |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
792 |
|> Global_Theory.store_thm (Binding.qualified_name (corr_name s vs), |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
793 |
Thm.varifyT_global (funpow (length (vars_of corr_prop)) |
35985
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
wenzelm
parents:
35845
diff
changeset
|
794 |
(Thm.forall_elim_var 0) (Thm.forall_intr_frees |
22750 | 795 |
(ProofChecker.thm_of_proof thy' |
26481 | 796 |
(fst (Proofterm.freeze_thaw_prf prf)))))) |
22750 | 797 |
|> snd |
28370 | 798 |
|> fold Code.add_default_eqn def_thms |
13732 | 799 |
end |
15531 | 800 |
| SOME _ => thy); |
13402 | 801 |
|
16149 | 802 |
in |
803 |
thy |
|
30435
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents:
30364
diff
changeset
|
804 |
|> Sign.root_path |
16149 | 805 |
|> fold_rev add_def defs |
22796 | 806 |
|> Sign.restore_naming thy |
13402 | 807 |
end; |
808 |
||
809 |
||
810 |
(**** interface ****) |
|
811 |
||
36950 | 812 |
val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") []; |
13732 | 813 |
|
24867 | 814 |
val _ = |
36953
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents:
36950
diff
changeset
|
815 |
Outer_Syntax.command "realizers" |
13402 | 816 |
"specify realizers for primitive axioms / theorems, together with correctness proof" |
36950 | 817 |
Keyword.thy_decl |
818 |
(Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >> |
|
13402 | 819 |
(fn xs => Toplevel.theory (fn thy => add_realizers |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
820 |
(map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy))); |
13402 | 821 |
|
24867 | 822 |
val _ = |
36953
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents:
36950
diff
changeset
|
823 |
Outer_Syntax.command "realizability" |
36950 | 824 |
"add equations characterizing realizability" Keyword.thy_decl |
825 |
(Scan.repeat1 Parse.string >> (Toplevel.theory o add_realizes_eqns)); |
|
13402 | 826 |
|
24867 | 827 |
val _ = |
36953
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents:
36950
diff
changeset
|
828 |
Outer_Syntax.command "extract_type" |
36950 | 829 |
"add equations characterizing type of extracted program" Keyword.thy_decl |
830 |
(Scan.repeat1 Parse.string >> (Toplevel.theory o add_typeof_eqns)); |
|
13402 | 831 |
|
24867 | 832 |
val _ = |
36953
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents:
36950
diff
changeset
|
833 |
Outer_Syntax.command "extract" "extract terms from proofs" Keyword.thy_decl |
36950 | 834 |
(Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy => |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
835 |
extract (map (apfst (Global_Theory.get_thm thy)) xs) thy))); |
13402 | 836 |
|
16458 | 837 |
val etype_of = etype_of o add_syntax; |
13714 | 838 |
|
13402 | 839 |
end; |