1 (* Title: Pure/Proof/extraction.ML
3 Author: Stefan Berghofer, TU Muenchen
5 Extraction of programs from proofs.
10 val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory
11 val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
12 val add_realizes_eqns : string list -> theory -> theory
13 val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory
14 val add_typeof_eqns : string list -> theory -> theory
15 val add_realizers_i : (string * (string list * term * Proofterm.proof)) list
17 val add_realizers : (thm * (string list * string * string)) list
19 val add_expand_thms : thm list -> theory -> theory
20 val add_types : (xstring * ((term -> term option) list *
21 (term -> typ -> term -> typ -> term) option)) list -> theory -> theory
22 val extract : (thm * string list) list -> theory -> theory
25 val mk_typ : typ -> term
26 val etype_of : theory -> string list -> typ list -> term -> typ
27 val realizes_of: theory -> string list -> term -> term -> term
30 structure Extraction : EXTRACTION =
41 |> Sign.add_types [("Type", 0, NoSyn), ("Null", 0, NoSyn)]
43 [("typeof", "'b::{} => Type", NoSyn),
44 ("Type", "'a::{} itself => Type", NoSyn),
45 ("Null", "Null", NoSyn),
46 ("realizes", "'a::{} => 'b::{} => 'b", NoSyn)];
48 val nullT = Type ("Null", []);
49 val nullt = Const ("Null", nullT);
52 Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T;
54 fun typeof_proc defaultS vs (Const ("typeof", _) $ u) =
55 SOME (mk_typ (case strip_comb u of
56 (Var ((a, i), _), _) =>
57 if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS)
60 if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT
62 | typeof_proc _ _ _ = NONE;
64 fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = SOME t
65 | rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) =
67 (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts))
68 | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts))
72 val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o
73 take_prefix (not o equal ":") o explode;
76 {next: int, rs: ((term * term) list * (term * term)) list,
77 net: (int * ((term * term) list * (term * term))) Net.net};
79 val empty_rules : rules = {next = 0, rs = [], net = Net.empty};
81 fun add_rule (r as (_, (lhs, _)), {next, rs, net} : rules) =
82 {next = next - 1, rs = r :: rs, net = Net.insert_term (K false)
83 (Envir.eta_contract lhs, (next, r)) net};
86 ({next, rs = rs1, net} : rules) ({next = next2, rs = rs2, ...} : rules) =
87 List.foldr add_rule {next = next, rs = rs1, net = net} (subtract (op =) rs1 rs2);
89 fun condrew thy rules procs =
92 Pattern.rewrite_term thy [] (condrew' :: procs) tm
95 val cache = ref ([] : (term * term) list);
96 fun lookup f x = (case AList.lookup (op =) (!cache) x of
99 in (cache := (x, y) :: !cache; y) end
102 get_first (fn (_, (prems, (tm1, tm2))) =>
104 fun ren t = the_default t (Term.rename_abs tm1 tm t);
105 val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1);
106 val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty);
107 val prems' = map (pairself (Envir.subst_vars env o inc o ren)) prems;
108 val env' = Envir.Envir
109 {maxidx = Library.foldl Int.max
110 (~1, map (Int.max o pairself maxidx_of_term) prems'),
111 iTs = Tenv, asol = tenv};
112 val env'' = fold (Pattern.unify thy o pairself (lookup rew)) prems' env';
113 in SOME (Envir.norm_term env'' (inc (ren tm2)))
114 end handle Pattern.MATCH => NONE | Pattern.Unif => NONE)
115 (sort (int_ord o pairself fst)
116 (Net.match_term rules (Envir.eta_contract tm)))
121 val chtype = change_type o SOME;
123 fun extr_name s vs = NameSpace.append "extr" (space_implode "_" (s :: vs));
124 fun corr_name s vs = extr_name s vs ^ "_correctness";
126 fun msg d s = priority (Symbol.spaces d ^ s);
128 fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []);
129 fun vfs_of t = vars_of t @ sort Term.term_ord (term_frees t);
131 fun forall_intr_prf (t, prf) =
132 let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
133 in Abst (a, SOME T, prf_abstract_over t prf) end;
135 val mkabs = List.foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t)));
137 fun strip_abs 0 t = t
138 | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t
139 | strip_abs _ _ = error "strip_abs: not an abstraction";
141 fun prf_subst_TVars tye =
142 map_proof_terms (subst_TVars tye) (typ_subst_TVars tye);
144 fun relevant_vars types prop = List.foldr (fn
145 (Var ((a, i), T), vs) => (case strip_type T of
146 (_, Type (s, _)) => if member (op =) types s then a :: vs else vs
148 | (_, vs) => vs) [] (vars_of prop);
150 fun tname_of (Type (s, _)) = s
155 val vs = Term.add_vars t [];
156 val fs = Term.add_frees t [];
158 Var (ixn, _) => (case AList.lookup (op =) vs ixn of
159 NONE => error "get_var_type: no such variable in term"
160 | SOME T => Var (ixn, T))
161 | Free (s, _) => (case AList.lookup (op =) fs s of
162 NONE => error "get_var_type: no such variable in term"
163 | SOME T => Free (s, T))
164 | _ => error "get_var_type: not a variable"
168 (**** theory data ****)
172 structure ExtractionData = TheoryDataFun
175 {realizes_eqns : rules,
177 types : (string * ((term -> term option) list *
178 (term -> typ -> term -> typ -> term) option)) list,
179 realizers : (string list * (term * proof)) list Symtab.table,
181 expand : (string * term) list,
182 prep : (theory -> proof -> proof) option}
185 {realizes_eqns = empty_rules,
186 typeof_eqns = empty_rules,
188 realizers = Symtab.empty,
196 (({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1,
197 realizers = realizers1, defs = defs1, expand = expand1, prep = prep1},
198 {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2,
199 realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T * T) =
200 {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2,
201 typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2,
202 types = AList.merge (op =) (K true) (types1, types2),
203 realizers = Symtab.merge_list (gen_eq_set (op =) o pairself #1) (realizers1, realizers2),
204 defs = Library.merge Thm.eq_thm (defs1, defs2),
205 expand = Library.merge (op =) (expand1, expand2),
206 prep = (case prep1 of NONE => prep2 | _ => prep1)};
209 fun read_condeq thy =
210 let val thy' = add_syntax thy
212 let val t = Logic.varify (Syntax.read_prop_global thy' s)
213 in (map Logic.dest_equals (Logic.strip_imp_prems t),
214 Logic.dest_equals (Logic.strip_imp_concl t))
215 end handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s)
220 fun set_preprocessor prep thy =
221 let val {realizes_eqns, typeof_eqns, types, realizers,
222 defs, expand, ...} = ExtractionData.get thy
225 {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
226 realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy
229 (** equations characterizing realizability **)
231 fun gen_add_realizes_eqns prep_eq eqns thy =
232 let val {realizes_eqns, typeof_eqns, types, realizers,
233 defs, expand, prep} = ExtractionData.get thy;
236 {realizes_eqns = List.foldr add_rule realizes_eqns (map (prep_eq thy) eqns),
237 typeof_eqns = typeof_eqns, types = types, realizers = realizers,
238 defs = defs, expand = expand, prep = prep} thy
241 val add_realizes_eqns_i = gen_add_realizes_eqns (K I);
242 val add_realizes_eqns = gen_add_realizes_eqns read_condeq;
244 (** equations characterizing type of extracted program **)
246 fun gen_add_typeof_eqns prep_eq eqns thy =
248 val {realizes_eqns, typeof_eqns, types, realizers,
249 defs, expand, prep} = ExtractionData.get thy;
250 val eqns' = map (prep_eq thy) eqns
253 {realizes_eqns = realizes_eqns, realizers = realizers,
254 typeof_eqns = List.foldr add_rule typeof_eqns eqns',
255 types = types, defs = defs, expand = expand, prep = prep} thy
258 val add_typeof_eqns_i = gen_add_typeof_eqns (K I);
259 val add_typeof_eqns = gen_add_typeof_eqns read_condeq;
261 fun thaw (T as TFree (a, S)) =
262 if exists_string (equal ":") a then TVar (unpack_ixn a, S) else T
263 | thaw (Type (a, Ts)) = Type (a, map thaw Ts)
266 fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S)
267 | freeze (Type (a, Ts)) = Type (a, map freeze Ts)
270 fun freeze_thaw f x =
271 map_types thaw (f (map_types freeze x));
273 fun etype_of thy vs Ts t =
275 val {typeof_eqns, ...} = ExtractionData.get thy;
276 fun err () = error ("Unable to determine type of extracted program for\n" ^
277 Syntax.string_of_term_global thy t)
278 in case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns)
279 [typeof_proc (Sign.defaultS thy) vs]) (list_abs (map (pair "x") (rev Ts),
280 Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of
281 Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ())
285 (** realizers for axioms / theorems, together with correctness proofs **)
287 fun gen_add_realizers prep_rlz rs thy =
288 let val {realizes_eqns, typeof_eqns, types, realizers,
289 defs, expand, prep} = ExtractionData.get thy
292 {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
293 realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers,
294 defs = defs, expand = expand, prep = prep} thy
297 fun prep_realizer thy =
299 val {realizes_eqns, typeof_eqns, defs, types, ...} =
300 ExtractionData.get thy;
301 val procs = maps (fst o snd) types;
302 val rtypes = map fst types;
303 val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
304 val thy' = add_syntax thy;
305 val rd = ProofSyntax.read_proof thy' false
306 in fn (thm, (vs, s1, s2)) =>
308 val name = Thm.get_name thm;
309 val _ = name <> "" orelse error "add_realizers: unnamed theorem";
310 val prop = Pattern.rewrite_term thy'
311 (map (Logic.dest_equals o prop_of) defs) [] (prop_of thm);
312 val vars = vars_of prop;
313 val vars' = filter_out (fn v =>
314 member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars;
315 val T = etype_of thy' vs [] prop;
316 val (T', thw) = Type.freeze_thaw_type
317 (if T = nullT then nullT else map fastype_of vars' ---> T);
318 val t = map_types thw (OldGoals.simple_read_term thy' T' s1);
319 val r' = freeze_thaw (condrew thy' eqns
320 (procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))
321 (Const ("realizes", T --> propT --> propT) $
322 (if T = nullT then t else list_comb (t, vars')) $ prop);
323 val r = fold_rev Logic.all (map (get_var_type r') vars) r';
324 val prf = Reconstruct.reconstruct_proof thy' r (rd s2);
325 in (name, (vs, (t, prf))) end
328 val add_realizers_i = gen_add_realizers
329 (fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf))));
330 val add_realizers = gen_add_realizers prep_realizer;
332 fun realizes_of thy vs t prop =
334 val thy' = add_syntax thy;
335 val {realizes_eqns, typeof_eqns, defs, types, ...} =
336 ExtractionData.get thy';
337 val procs = maps (rev o fst o snd) types;
338 val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
339 val prop' = Pattern.rewrite_term thy'
340 (map (Logic.dest_equals o prop_of) defs) [] prop;
341 in freeze_thaw (condrew thy' eqns
342 (procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))
343 (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop')
346 (** expanding theorems / definitions **)
348 fun add_expand_thm thm thy =
350 val {realizes_eqns, typeof_eqns, types, realizers,
351 defs, expand, prep} = ExtractionData.get thy;
353 val name = Thm.get_name thm;
354 val _ = name <> "" orelse error "add_expand_thms: unnamed theorem";
357 (case strip_comb (fst (Logic.dest_equals (prop_of thm))) of
358 (Const _, ts) => forall is_Var ts andalso not (has_duplicates (op =) ts)
359 andalso (PureThy.get_kind thm = Thm.definitionK orelse can (Thm.get_axiom_i thy) name)
360 | _ => false) handle TERM _ => false;
362 (ExtractionData.put (if is_def then
363 {realizes_eqns = realizes_eqns,
364 typeof_eqns = add_rule (([],
365 Logic.dest_equals (prop_of (Drule.abs_def thm))), typeof_eqns),
367 realizers = realizers, defs = insert Thm.eq_thm thm defs,
368 expand = expand, prep = prep}
370 {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
371 realizers = realizers, defs = defs,
372 expand = insert (op =) (name, prop_of thm) expand, prep = prep}) thy)
375 val add_expand_thms = fold add_expand_thm;
377 val extraction_expand =
378 Attrib.no_args (Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm th) I));
381 (** types with computational content **)
383 fun add_types tys thy =
385 (fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =>
386 {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns,
387 types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types,
388 realizers = realizers, defs = defs, expand = expand, prep = prep})
394 val _ = Context.>> (Context.map_theory
395 (add_types [("prop", ([], NONE))] #>
398 ["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \
399 \ (typeof (PROP Q)) == (Type (TYPE('Q))) ==> \
400 \ (typeof (PROP P ==> PROP Q)) == (Type (TYPE('Q)))",
402 "(typeof (PROP Q)) == (Type (TYPE(Null))) ==> \
403 \ (typeof (PROP P ==> PROP Q)) == (Type (TYPE(Null)))",
405 "(typeof (PROP P)) == (Type (TYPE('P))) ==> \
406 \ (typeof (PROP Q)) == (Type (TYPE('Q))) ==> \
407 \ (typeof (PROP P ==> PROP Q)) == (Type (TYPE('P => 'Q)))",
409 "(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \
410 \ (typeof (!!x. PROP P (x))) == (Type (TYPE(Null)))",
412 "(%x. typeof (PROP P (x))) == (%x. Type (TYPE('P))) ==> \
413 \ (typeof (!!x::'a. PROP P (x))) == (Type (TYPE('a => 'P)))",
415 "(%x. typeof (f (x))) == (%x. Type (TYPE('f))) ==> \
416 \ (typeof (f)) == (Type (TYPE('f)))"] #>
419 ["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \
420 \ (realizes (r) (PROP P ==> PROP Q)) == \
421 \ (PROP realizes (Null) (PROP P) ==> PROP realizes (r) (PROP Q))",
423 "(typeof (PROP P)) == (Type (TYPE('P))) ==> \
424 \ (typeof (PROP Q)) == (Type (TYPE(Null))) ==> \
425 \ (realizes (r) (PROP P ==> PROP Q)) == \
426 \ (!!x::'P. PROP realizes (x) (PROP P) ==> PROP realizes (Null) (PROP Q))",
428 "(realizes (r) (PROP P ==> PROP Q)) == \
429 \ (!!x. PROP realizes (x) (PROP P) ==> PROP realizes (r (x)) (PROP Q))",
431 "(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \
432 \ (realizes (r) (!!x. PROP P (x))) == \
433 \ (!!x. PROP realizes (Null) (PROP P (x)))",
435 "(realizes (r) (!!x. PROP P (x))) == \
436 \ (!!x. PROP realizes (r (x)) (PROP P (x)))"] #>
438 Attrib.add_attributes
439 [("extraction_expand", extraction_expand,
440 "specify theorems / definitions to be expanded during extraction")]));
443 (**** extract program ****)
445 val dummyt = Const ("dummy", dummyT);
447 fun extract thms thy =
449 val thy' = add_syntax thy;
450 val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =
451 ExtractionData.get thy;
452 val procs = maps (rev o fst o snd) types;
453 val rtypes = map fst types;
454 val typroc = typeof_proc (Sign.defaultS thy');
455 val prep = the_default (K I) prep thy' o ProofRewriteRules.elim_defs thy' false defs o
456 Reconstruct.expand_proof thy' (("", NONE) :: map (apsnd SOME) expand);
457 val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
459 fun find_inst prop Ts ts vs =
461 val rvs = relevant_vars rtypes prop;
462 val vars = vars_of prop;
463 val n = Int.min (length vars, length ts);
465 fun add_args ((Var ((a, i), _), t), (vs', tye)) =
466 if member (op =) rvs a then
467 let val T = etype_of thy' vs Ts t
468 in if T = nullT then (vs', tye)
469 else (a :: vs', (("'" ^ a, i), T) :: tye)
473 in List.foldr add_args ([], []) (Library.take (n, vars) ~~ Library.take (n, ts)) end;
475 fun find (vs: string list) = Option.map snd o find_first (curry (gen_eq_set (op =)) vs o fst);
476 fun find' s = map snd o List.filter (equal s o fst)
478 fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw
479 (condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (list_abs
480 (map (pair "x") (rev Ts), t)));
482 fun realizes_null vs prop = app_rlz_rews [] vs
483 (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop);
485 fun corr d defs vs ts Ts hs (PBound i) _ _ = (defs, PBound i)
487 | corr d defs vs ts Ts hs (Abst (s, SOME T, prf)) (Abst (_, _, prf')) t =
488 let val (defs', corr_prf) = corr d defs vs [] (T :: Ts)
489 (dummyt :: hs) prf (incr_pboundvars 1 0 prf')
490 (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE)
491 in (defs', Abst (s, SOME T, corr_prf)) end
493 | corr d defs vs ts Ts hs (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) t =
495 val T = etype_of thy' vs Ts prop;
496 val u = if T = nullT then
497 (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE)
498 else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE);
499 val (defs', corr_prf) = corr d defs vs [] (T :: Ts) (prop :: hs)
500 (incr_pboundvars 0 1 prf) (incr_pboundvars 0 1 prf') u;
501 val rlz = Const ("realizes", T --> propT --> propT)
503 if T = nullT then AbsP ("R",
504 SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)),
505 prf_subst_bounds [nullt] corr_prf)
506 else Abst (s, SOME T, AbsP ("R",
507 SOME (app_rlz_rews (T :: Ts) vs
508 (rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)))
511 | corr d defs vs ts Ts hs (prf % SOME t) (prf' % _) t' =
513 val (Us, T) = strip_type (fastype_of1 (Ts, t));
514 val (defs', corr_prf) = corr d defs vs (t :: ts) Ts hs prf prf'
515 (if member (op =) rtypes (tname_of T) then t'
516 else (case t' of SOME (u $ _) => SOME u | _ => NONE));
517 val u = if not (member (op =) rtypes (tname_of T)) then t else
519 val eT = etype_of thy' vs Ts t;
520 val (r, Us') = if eT = nullT then (nullt, Us) else
521 (Bound (length Us), eT :: Us);
522 val u = list_comb (incr_boundvars (length Us') t,
523 map Bound (length Us - 1 downto 0));
524 val u' = (case AList.lookup (op =) types (tname_of T) of
525 SOME ((_, SOME f)) => f r eT u T
526 | _ => Const ("realizes", eT --> T --> T) $ r $ u)
527 in app_rlz_rews Ts vs (list_abs (map (pair "x") Us', u')) end
528 in (defs', corr_prf % SOME u) end
530 | corr d defs vs ts Ts hs (prf1 %% prf2) (prf1' %% prf2') t =
532 val prop = Reconstruct.prop_of' hs prf2';
533 val T = etype_of thy' vs Ts prop;
534 val (defs1, f, u) = if T = nullT then (defs, t, NONE) else
536 SOME (f $ u) => (defs, SOME f, SOME u)
538 let val (defs1, u) = extr d defs vs [] Ts hs prf2'
539 in (defs1, NONE, SOME u) end)
540 val (defs2, corr_prf1) = corr d defs1 vs [] Ts hs prf1 prf1' f;
541 val (defs3, corr_prf2) = corr d defs2 vs [] Ts hs prf2 prf2' u;
543 if T = nullT then (defs3, corr_prf1 %% corr_prf2) else
544 (defs3, corr_prf1 % u %% corr_prf2)
547 | corr d defs vs ts Ts hs (prf0 as PThm (name, prf, prop, SOME Ts')) _ _ =
549 val (vs', tye) = find_inst prop Ts ts vs;
550 val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;
551 val T = etype_of thy' vs' [] prop;
552 val defs' = if T = nullT then defs
553 else fst (extr d defs vs ts Ts hs prf0)
555 if T = nullT andalso realizes_null vs' prop aconv prop then (defs, prf0)
556 else case Symtab.lookup realizers name of
557 NONE => (case find vs' (find' name defs') of
560 val _ = T = nullT orelse error "corr: internal error";
561 val _ = msg d ("Building correctness proof for " ^ quote name ^
563 else " (relevant variables: " ^ commas_quote vs' ^ ")"));
564 val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);
565 val (defs'', corr_prf) =
566 corr (d + 1) defs' vs' [] [] [] prf' prf' NONE;
567 val corr_prop = Reconstruct.prop_of corr_prf;
568 val corr_prf' = List.foldr forall_intr_prf
570 (PThm (corr_name name vs', corr_prf, corr_prop,
571 SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop))
572 (map (get_var_type corr_prop) (vfs_of prop))
574 ((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',
575 prf_subst_TVars tye' corr_prf')
577 | SOME (_, (_, prf')) => (defs', prf_subst_TVars tye' prf'))
578 | SOME rs => (case find vs' rs of
579 SOME (_, prf') => (defs', prf_subst_TVars tye' prf')
580 | NONE => error ("corr: no realizer for instance of theorem " ^
581 quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
582 (Reconstruct.prop_of (proof_combt (prf0, ts))))))
585 | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =
587 val (vs', tye) = find_inst prop Ts ts vs;
588 val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
590 if etype_of thy' vs' [] prop = nullT andalso
591 realizes_null vs' prop aconv prop then (defs, prf0)
592 else case find vs' (Symtab.lookup_list realizers s) of
593 SOME (_, prf) => (defs, prf_subst_TVars tye' prf)
594 | NONE => error ("corr: no realizer for instance of axiom " ^
595 quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
596 (Reconstruct.prop_of (proof_combt (prf0, ts)))))
599 | corr d defs vs ts Ts hs _ _ _ = error "corr: bad proof"
601 and extr d defs vs ts Ts hs (PBound i) = (defs, Bound i)
603 | extr d defs vs ts Ts hs (Abst (s, SOME T, prf)) =
604 let val (defs', t) = extr d defs vs []
605 (T :: Ts) (dummyt :: hs) (incr_pboundvars 1 0 prf)
606 in (defs', Abs (s, T, t)) end
608 | extr d defs vs ts Ts hs (AbsP (s, SOME t, prf)) =
610 val T = etype_of thy' vs Ts t;
611 val (defs', t) = extr d defs vs [] (T :: Ts) (t :: hs)
612 (incr_pboundvars 0 1 prf)
614 if T = nullT then subst_bound (nullt, t) else Abs (s, T, t))
617 | extr d defs vs ts Ts hs (prf % SOME t) =
618 let val (defs', u) = extr d defs vs (t :: ts) Ts hs prf
620 if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u
624 | extr d defs vs ts Ts hs (prf1 %% prf2) =
626 val (defs', f) = extr d defs vs [] Ts hs prf1;
627 val prop = Reconstruct.prop_of' hs prf2;
628 val T = etype_of thy' vs Ts prop
630 if T = nullT then (defs', f) else
631 let val (defs'', t) = extr d defs' vs [] Ts hs prf2
632 in (defs'', f $ t) end
635 | extr d defs vs ts Ts hs (prf0 as PThm (s, prf, prop, SOME Ts')) =
637 val (vs', tye) = find_inst prop Ts ts vs;
638 val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
640 case Symtab.lookup realizers s of
641 NONE => (case find vs' (find' s defs) of
644 val _ = msg d ("Extracting " ^ quote s ^
646 else " (relevant variables: " ^ commas_quote vs' ^ ")"));
647 val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);
648 val (defs', t) = extr (d + 1) defs vs' [] [] [] prf';
649 val (defs'', corr_prf) =
650 corr (d + 1) defs' vs' [] [] [] prf' prf' (SOME t);
652 val nt = Envir.beta_norm t;
653 val args = filter_out (fn v => member (op =) rtypes
654 (tname_of (body_type (fastype_of v)))) (vfs_of prop);
655 val args' = List.filter (fn v => Logic.occs (v, nt)) args;
656 val t' = mkabs nt args';
657 val T = fastype_of t';
658 val cname = extr_name s vs';
659 val c = Const (cname, T);
660 val u = mkabs (list_comb (c, args')) args;
661 val eqn = Logic.mk_equals (c, t');
663 Const ("realizes", fastype_of nt --> propT --> propT);
664 val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop);
665 val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop);
666 val f = app_rlz_rews [] vs'
667 (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop));
670 chtype [] equal_elim_axm %> lhs %> rhs %%
671 (chtype [propT] symmetric_axm %> rhs %> lhs %%
672 (chtype [propT, T] combination_axm %> f %> f %> c %> t' %%
673 (chtype [T --> propT] reflexive_axm %> f) %%
674 PAxm (cname ^ "_def", eqn,
675 SOME (map TVar (term_tvars eqn))))) %% corr_prf;
676 val corr_prop = Reconstruct.prop_of corr_prf';
677 val corr_prf'' = List.foldr forall_intr_prf
679 (PThm (corr_name s vs', corr_prf', corr_prop,
680 SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop))
681 (map (get_var_type corr_prop) (vfs_of prop));
683 ((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',
686 | SOME ((_, u), _) => (defs, subst_TVars tye' u))
687 | SOME rs => (case find vs' rs of
688 SOME (t, _) => (defs, subst_TVars tye' t)
689 | NONE => error ("extr: no realizer for instance of theorem " ^
690 quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
691 (Reconstruct.prop_of (proof_combt (prf0, ts))))))
694 | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =
696 val (vs', tye) = find_inst prop Ts ts vs;
697 val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
699 case find vs' (Symtab.lookup_list realizers s) of
700 SOME (t, _) => (defs, subst_TVars tye' t)
701 | NONE => error ("extr: no realizer for instance of axiom " ^
702 quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm
703 (Reconstruct.prop_of (proof_combt (prf0, ts)))))
706 | extr d defs vs ts Ts hs _ = error "extr: bad proof";
708 fun prep_thm (thm, vs) =
710 val thy = Thm.theory_of_thm thm;
711 val prop = Thm.prop_of thm;
712 val prf = Thm.proof_of thm;
713 val name = Thm.get_name thm;
714 val _ = name <> "" orelse error "extraction: unnamed theorem";
715 val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^
716 quote name ^ " has no computational content")
717 in (Reconstruct.reconstruct_proof thy prop prf, vs) end;
719 val defs = Library.foldl (fn (defs, (prf, vs)) =>
720 fst (extr 0 defs vs [] [] [] prf)) ([], map prep_thm thms);
722 fun add_def (s, (vs, ((t, u), (prf, _)))) thy =
723 (case Sign.const_type thy (extr_name s vs) of
726 val corr_prop = Reconstruct.prop_of prf;
727 val ft = Type.freeze t;
728 val fu = Type.freeze u;
729 val (def_thms, thy') = if t = nullt then ([], thy) else
731 |> Sign.add_consts_i [(extr_name s vs, fastype_of ft, NoSyn)]
732 |> PureThy.add_defs false [((extr_name s vs ^ "_def",
733 Logic.mk_equals (head_of (strip_abs_body fu), ft)), [])]
736 |> PureThy.store_thm (corr_name s vs,
737 Thm.varifyT (funpow (length (term_vars corr_prop))
738 (Thm.forall_elim_var 0) (forall_intr_frees
739 (ProofChecker.thm_of_proof thy'
740 (fst (Proofterm.freeze_thaw_prf prf))))))
742 |> fold Code.add_default_func def_thms
748 |> Sign.absolute_path
749 |> fold_rev add_def defs
750 |> Sign.restore_naming thy
754 (**** interface ****)
756 structure P = OuterParse and K = OuterKeyword;
758 val parse_vars = Scan.optional (P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") [];
761 OuterSyntax.command "realizers"
762 "specify realizers for primitive axioms / theorems, together with correctness proof"
764 (Scan.repeat1 (P.xname -- parse_vars --| P.$$$ ":" -- P.string -- P.string) >>
765 (fn xs => Toplevel.theory (fn thy => add_realizers
766 (map (fn (((a, vs), s1), s2) => (PureThy.get_thm thy a, (vs, s1, s2))) xs) thy)));
769 OuterSyntax.command "realizability"
770 "add equations characterizing realizability" K.thy_decl
771 (Scan.repeat1 P.string >> (Toplevel.theory o add_realizes_eqns));
774 OuterSyntax.command "extract_type"
775 "add equations characterizing type of extracted program" K.thy_decl
776 (Scan.repeat1 P.string >> (Toplevel.theory o add_typeof_eqns));
779 OuterSyntax.command "extract" "extract terms from proofs" K.thy_decl
780 (Scan.repeat1 (P.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>
781 extract (map (apfst (PureThy.get_thm thy)) xs) thy)));
783 val etype_of = etype_of o add_syntax;