| author | wenzelm | 
| Wed, 31 Dec 2008 00:01:07 +0100 | |
| changeset 29262 | 3ee4656b9e0c | 
| parent 28814 | 463c9e9111ae | 
| child 29265 | 5b4247055bd7 | 
| permissions | -rw-r--r-- | 
| 13402 | 1  | 
(* Title: Pure/Proof/extraction.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Stefan Berghofer, TU Muenchen  | 
|
4  | 
||
5  | 
Extraction of programs from proofs.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature EXTRACTION =  | 
|
9  | 
sig  | 
|
| 16458 | 10  | 
val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory  | 
| 13402 | 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  | 
|
16  | 
-> theory -> theory  | 
|
17  | 
val add_realizers : (thm * (string list * string * string)) list  | 
|
18  | 
-> theory -> theory  | 
|
19  | 
val add_expand_thms : thm list -> theory -> theory  | 
|
| 13732 | 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  | 
|
| 13402 | 23  | 
val nullT : typ  | 
24  | 
val nullt : term  | 
|
| 13714 | 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  | 
|
| 13402 | 28  | 
end;  | 
29  | 
||
30  | 
structure Extraction : EXTRACTION =  | 
|
31  | 
struct  | 
|
32  | 
||
33  | 
open Proofterm;  | 
|
34  | 
||
35  | 
(**** tools ****)  | 
|
36  | 
||
37  | 
fun add_syntax thy =  | 
|
38  | 
thy  | 
|
39  | 
|> Theory.copy  | 
|
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
24707 
diff
changeset
 | 
40  | 
|> Sign.root_path  | 
| 22796 | 41  | 
  |> Sign.add_types [("Type", 0, NoSyn), ("Null", 0, NoSyn)]
 | 
42  | 
|> Sign.add_consts  | 
|
| 14854 | 43  | 
      [("typeof", "'b::{} => Type", NoSyn),
 | 
44  | 
       ("Type", "'a::{} itself => Type", NoSyn),
 | 
|
| 13402 | 45  | 
       ("Null", "Null", NoSyn),
 | 
| 14854 | 46  | 
       ("realizes", "'a::{} => 'b::{} => 'b", NoSyn)];
 | 
| 13402 | 47  | 
|
48  | 
val nullT = Type ("Null", []);
 | 
|
49  | 
val nullt = Const ("Null", nullT);
 | 
|
50  | 
||
51  | 
fun mk_typ T =  | 
|
| 19391 | 52  | 
  Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T;
 | 
| 13402 | 53  | 
|
54  | 
fun typeof_proc defaultS vs (Const ("typeof", _) $ u) =
 | 
|
| 15531 | 55  | 
SOME (mk_typ (case strip_comb u of  | 
| 13402 | 56  | 
(Var ((a, i), _), _) =>  | 
| 20664 | 57  | 
            if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS)
 | 
| 13402 | 58  | 
else nullT  | 
59  | 
| (Free (a, _), _) =>  | 
|
| 20664 | 60  | 
            if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT
 | 
| 13402 | 61  | 
| _ => nullT))  | 
| 15531 | 62  | 
| typeof_proc _ _ _ = NONE;  | 
| 13402 | 63  | 
|
| 15531 | 64  | 
fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = SOME t
 | 
| 13732 | 65  | 
  | rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) =
 | 
66  | 
(case strip_comb t of  | 
|
| 15531 | 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))  | 
|
69  | 
| _ => NONE)  | 
|
70  | 
| rlz_proc _ = NONE;  | 
|
| 13402 | 71  | 
|
72  | 
val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o  | 
|
| 28375 | 73  | 
take_prefix (fn s => s <> ":") o explode;  | 
| 13402 | 74  | 
|
75  | 
type rules =  | 
|
76  | 
  {next: int, rs: ((term * term) list * (term * term)) list,
 | 
|
77  | 
net: (int * ((term * term) list * (term * term))) Net.net};  | 
|
78  | 
||
79  | 
val empty_rules : rules = {next = 0, rs = [], net = Net.empty};
 | 
|
80  | 
||
81  | 
fun add_rule (r as (_, (lhs, _)), {next, rs, net} : rules) =
 | 
|
| 16800 | 82  | 
  {next = next - 1, rs = r :: rs, net = Net.insert_term (K false)
 | 
| 18956 | 83  | 
(Envir.eta_contract lhs, (next, r)) net};  | 
| 13402 | 84  | 
|
| 
13417
 
12cc77f90811
Tuned type constraint of function merge_rules to make smlnj happy.
 
berghofe 
parents: 
13402 
diff
changeset
 | 
85  | 
fun merge_rules  | 
| 
 
12cc77f90811
Tuned type constraint of function merge_rules to make smlnj happy.
 
berghofe 
parents: 
13402 
diff
changeset
 | 
86  | 
  ({next, rs = rs1, net} : rules) ({next = next2, rs = rs2, ...} : rules) =
 | 
| 23178 | 87  | 
  List.foldr add_rule {next = next, rs = rs1, net = net} (subtract (op =) rs1 rs2);
 | 
| 13402 | 88  | 
|
| 16458 | 89  | 
fun condrew thy rules procs =  | 
| 13402 | 90  | 
let  | 
91  | 
fun rew tm =  | 
|
| 17203 | 92  | 
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
 | 
93  | 
and condrew' tm =  | 
| 13402 | 94  | 
let  | 
| 
15399
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
95  | 
val cache = ref ([] : (term * term) list);  | 
| 17232 | 96  | 
fun lookup f x = (case AList.lookup (op =) (!cache) x of  | 
| 15531 | 97  | 
NONE =>  | 
| 
15399
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
98  | 
let val y = f x  | 
| 
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
99  | 
in (cache := (x, y) :: !cache; y) end  | 
| 15531 | 100  | 
| SOME y => y);  | 
| 
15399
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
101  | 
in  | 
| 
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
102  | 
get_first (fn (_, (prems, (tm1, tm2))) =>  | 
| 
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
103  | 
let  | 
| 19466 | 104  | 
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
 | 
105  | 
val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1);  | 
| 18184 | 106  | 
val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty);  | 
| 
15798
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
107  | 
val prems' = map (pairself (Envir.subst_vars 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
 | 
108  | 
val env' = Envir.Envir  | 
| 15570 | 109  | 
            {maxidx = Library.foldl Int.max
 | 
| 
15399
 
683d83051d6a
Added term cache to function condrew in order to speed up rewriting.
 
berghofe 
parents: 
14981 
diff
changeset
 | 
110  | 
(~1, map (Int.max o pairself maxidx_of_term) prems'),  | 
| 
15798
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
111  | 
iTs = Tenv, asol = tenv};  | 
| 18184 | 112  | 
val env'' = fold (Pattern.unify thy o pairself (lookup rew)) prems' env';  | 
| 15531 | 113  | 
in SOME (Envir.norm_term env'' (inc (ren tm2)))  | 
114  | 
end handle Pattern.MATCH => NONE | Pattern.Unif => NONE)  | 
|
| 16486 | 115  | 
(sort (int_ord o pairself fst)  | 
| 18956 | 116  | 
(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
 | 
117  | 
end;  | 
| 13402 | 118  | 
|
119  | 
in rew end;  | 
|
120  | 
||
| 15531 | 121  | 
val chtype = change_type o SOME;  | 
| 13402 | 122  | 
|
| 16195 | 123  | 
fun extr_name s vs = NameSpace.append "extr" (space_implode "_" (s :: vs));  | 
124  | 
fun corr_name s vs = extr_name s vs ^ "_correctness";  | 
|
| 13402 | 125  | 
|
| 16195 | 126  | 
fun msg d s = priority (Symbol.spaces d ^ s);  | 
| 13402 | 127  | 
|
| 
28812
 
413695e07bd4
Frees in PThms are now quantified in the order of their appearance in the
 
berghofe 
parents: 
28805 
diff
changeset
 | 
128  | 
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
 | 
129  | 
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
 | 
130  | 
fun vfs_of t = vars_of t @ frees_of t;  | 
| 13402 | 131  | 
|
132  | 
fun forall_intr_prf (t, prf) =  | 
|
133  | 
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)  | 
|
| 15531 | 134  | 
in Abst (a, SOME T, prf_abstract_over t prf) end;  | 
| 13402 | 135  | 
|
| 23178 | 136  | 
val mkabs = List.foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t)));
 | 
| 13402 | 137  | 
|
| 13732 | 138  | 
fun strip_abs 0 t = t  | 
139  | 
| strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t  | 
|
140  | 
| strip_abs _ _ = error "strip_abs: not an abstraction";  | 
|
141  | 
||
| 13402 | 142  | 
fun prf_subst_TVars tye =  | 
143  | 
map_proof_terms (subst_TVars tye) (typ_subst_TVars tye);  | 
|
144  | 
||
| 23178 | 145  | 
fun relevant_vars types prop = List.foldr (fn  | 
| 13402 | 146  | 
(Var ((a, i), T), vs) => (case strip_type T of  | 
| 20664 | 147  | 
(_, Type (s, _)) => if member (op =) types s then a :: vs else vs  | 
| 13402 | 148  | 
| _ => vs)  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
149  | 
| (_, vs) => vs) [] (vars_of prop);  | 
| 13402 | 150  | 
|
| 13732 | 151  | 
fun tname_of (Type (s, _)) = s  | 
152  | 
| tname_of _ = "";  | 
|
153  | 
||
154  | 
fun get_var_type t =  | 
|
155  | 
let  | 
|
| 16865 | 156  | 
val vs = Term.add_vars t [];  | 
157  | 
val fs = Term.add_frees t [];  | 
|
| 13732 | 158  | 
in fn  | 
| 17232 | 159  | 
Var (ixn, _) => (case AList.lookup (op =) vs ixn of  | 
| 15531 | 160  | 
NONE => error "get_var_type: no such variable in term"  | 
161  | 
| SOME T => Var (ixn, T))  | 
|
| 17232 | 162  | 
| Free (s, _) => (case AList.lookup (op =) fs s of  | 
| 15531 | 163  | 
NONE => error "get_var_type: no such variable in term"  | 
164  | 
| SOME T => Free (s, T))  | 
|
| 13732 | 165  | 
| _ => error "get_var_type: not a variable"  | 
166  | 
end;  | 
|
167  | 
||
| 13402 | 168  | 
|
169  | 
(**** theory data ****)  | 
|
170  | 
||
| 22846 | 171  | 
(* theory data *)  | 
| 13402 | 172  | 
|
| 16458 | 173  | 
structure ExtractionData = TheoryDataFun  | 
| 22846 | 174  | 
(  | 
| 13402 | 175  | 
type T =  | 
176  | 
    {realizes_eqns : rules,
 | 
|
177  | 
typeof_eqns : rules,  | 
|
| 13732 | 178  | 
types : (string * ((term -> term option) list *  | 
179  | 
(term -> typ -> term -> typ -> term) option)) list,  | 
|
| 13402 | 180  | 
realizers : (string list * (term * proof)) list Symtab.table,  | 
181  | 
defs : thm list,  | 
|
182  | 
expand : (string * term) list,  | 
|
| 16458 | 183  | 
prep : (theory -> proof -> proof) option}  | 
| 13402 | 184  | 
|
185  | 
val empty =  | 
|
186  | 
    {realizes_eqns = empty_rules,
 | 
|
187  | 
typeof_eqns = empty_rules,  | 
|
188  | 
types = [],  | 
|
189  | 
realizers = Symtab.empty,  | 
|
190  | 
defs = [],  | 
|
191  | 
expand = [],  | 
|
| 15531 | 192  | 
prep = NONE};  | 
| 13402 | 193  | 
val copy = I;  | 
| 16458 | 194  | 
val extend = I;  | 
| 13402 | 195  | 
|
| 16458 | 196  | 
fun merge _  | 
| 13402 | 197  | 
    (({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1,
 | 
198  | 
realizers = realizers1, defs = defs1, expand = expand1, prep = prep1},  | 
|
199  | 
      {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2,
 | 
|
200  | 
realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T * T) =  | 
|
201  | 
    {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2,
 | 
|
202  | 
typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2,  | 
|
| 22717 | 203  | 
types = AList.merge (op =) (K true) (types1, types2),  | 
| 19305 | 204  | 
realizers = Symtab.merge_list (gen_eq_set (op =) o pairself #1) (realizers1, realizers2),  | 
| 22662 | 205  | 
defs = Library.merge Thm.eq_thm (defs1, defs2),  | 
206  | 
expand = Library.merge (op =) (expand1, expand2),  | 
|
| 15531 | 207  | 
prep = (case prep1 of NONE => prep2 | _ => prep1)};  | 
| 22846 | 208  | 
);  | 
| 13402 | 209  | 
|
210  | 
fun read_condeq thy =  | 
|
| 16458 | 211  | 
let val thy' = add_syntax thy  | 
| 13402 | 212  | 
in fn s =>  | 
| 24707 | 213  | 
let val t = Logic.varify (Syntax.read_prop_global thy' s)  | 
| 13402 | 214  | 
in (map Logic.dest_equals (Logic.strip_imp_prems t),  | 
215  | 
Logic.dest_equals (Logic.strip_imp_concl t))  | 
|
216  | 
    end handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s)
 | 
|
217  | 
end;  | 
|
218  | 
||
219  | 
(** preprocessor **)  | 
|
220  | 
||
221  | 
fun set_preprocessor prep thy =  | 
|
222  | 
  let val {realizes_eqns, typeof_eqns, types, realizers,
 | 
|
223  | 
defs, expand, ...} = ExtractionData.get thy  | 
|
224  | 
in  | 
|
225  | 
ExtractionData.put  | 
|
226  | 
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
 | 
|
| 15531 | 227  | 
realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy  | 
| 13402 | 228  | 
end;  | 
229  | 
||
230  | 
(** equations characterizing realizability **)  | 
|
231  | 
||
232  | 
fun gen_add_realizes_eqns prep_eq eqns thy =  | 
|
233  | 
  let val {realizes_eqns, typeof_eqns, types, realizers,
 | 
|
234  | 
defs, expand, prep} = ExtractionData.get thy;  | 
|
235  | 
in  | 
|
236  | 
ExtractionData.put  | 
|
| 23178 | 237  | 
      {realizes_eqns = List.foldr add_rule realizes_eqns (map (prep_eq thy) eqns),
 | 
| 13402 | 238  | 
typeof_eqns = typeof_eqns, types = types, realizers = realizers,  | 
239  | 
defs = defs, expand = expand, prep = prep} thy  | 
|
240  | 
end  | 
|
241  | 
||
242  | 
val add_realizes_eqns_i = gen_add_realizes_eqns (K I);  | 
|
243  | 
val add_realizes_eqns = gen_add_realizes_eqns read_condeq;  | 
|
244  | 
||
245  | 
(** equations characterizing type of extracted program **)  | 
|
246  | 
||
247  | 
fun gen_add_typeof_eqns prep_eq eqns thy =  | 
|
248  | 
let  | 
|
249  | 
    val {realizes_eqns, typeof_eqns, types, realizers,
 | 
|
250  | 
defs, expand, prep} = ExtractionData.get thy;  | 
|
| 13732 | 251  | 
val eqns' = map (prep_eq thy) eqns  | 
| 13402 | 252  | 
in  | 
253  | 
ExtractionData.put  | 
|
254  | 
      {realizes_eqns = realizes_eqns, realizers = realizers,
 | 
|
| 23178 | 255  | 
typeof_eqns = List.foldr add_rule typeof_eqns eqns',  | 
| 13732 | 256  | 
types = types, defs = defs, expand = expand, prep = prep} thy  | 
| 13402 | 257  | 
end  | 
258  | 
||
259  | 
val add_typeof_eqns_i = gen_add_typeof_eqns (K I);  | 
|
260  | 
val add_typeof_eqns = gen_add_typeof_eqns read_condeq;  | 
|
261  | 
||
262  | 
fun thaw (T as TFree (a, S)) =  | 
|
| 28375 | 263  | 
if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T  | 
| 13402 | 264  | 
| thaw (Type (a, Ts)) = Type (a, map thaw Ts)  | 
265  | 
| thaw T = T;  | 
|
266  | 
||
267  | 
fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S)  | 
|
268  | 
| freeze (Type (a, Ts)) = Type (a, map freeze Ts)  | 
|
269  | 
| freeze T = T;  | 
|
270  | 
||
271  | 
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
 | 
272  | 
map_types thaw (f (map_types freeze x));  | 
| 13402 | 273  | 
|
| 16458 | 274  | 
fun etype_of thy vs Ts t =  | 
| 13402 | 275  | 
let  | 
| 16458 | 276  | 
    val {typeof_eqns, ...} = ExtractionData.get thy;
 | 
| 13402 | 277  | 
    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
 | 
278  | 
Syntax.string_of_term_global thy t)  | 
| 16458 | 279  | 
in case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns)  | 
280  | 
[typeof_proc (Sign.defaultS thy) vs]) (list_abs (map (pair "x") (rev Ts),  | 
|
| 13402 | 281  | 
      Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of
 | 
282  | 
      Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ())
 | 
|
283  | 
| _ => err ()  | 
|
284  | 
end;  | 
|
285  | 
||
286  | 
(** realizers for axioms / theorems, together with correctness proofs **)  | 
|
287  | 
||
288  | 
fun gen_add_realizers prep_rlz rs thy =  | 
|
289  | 
  let val {realizes_eqns, typeof_eqns, types, realizers,
 | 
|
290  | 
defs, expand, prep} = ExtractionData.get thy  | 
|
291  | 
in  | 
|
292  | 
ExtractionData.put  | 
|
293  | 
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
 | 
|
| 25389 | 294  | 
realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers,  | 
| 13402 | 295  | 
defs = defs, expand = expand, prep = prep} thy  | 
296  | 
end  | 
|
297  | 
||
298  | 
fun prep_realizer thy =  | 
|
299  | 
let  | 
|
| 13732 | 300  | 
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
 | 
| 13402 | 301  | 
ExtractionData.get thy;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19466 
diff
changeset
 | 
302  | 
val procs = maps (fst o snd) types;  | 
| 13732 | 303  | 
val rtypes = map fst types;  | 
| 16800 | 304  | 
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);  | 
| 13402 | 305  | 
val thy' = add_syntax thy;  | 
306  | 
val rd = ProofSyntax.read_proof thy' false  | 
|
307  | 
in fn (thm, (vs, s1, s2)) =>  | 
|
308  | 
let  | 
|
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
20897 
diff
changeset
 | 
309  | 
val name = Thm.get_name thm;  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21646 
diff
changeset
 | 
310  | 
val _ = name <> "" orelse error "add_realizers: unnamed theorem";  | 
| 17203 | 311  | 
val prop = Pattern.rewrite_term thy'  | 
| 13402 | 312  | 
(map (Logic.dest_equals o prop_of) defs) [] (prop_of thm);  | 
313  | 
val vars = vars_of prop;  | 
|
| 13732 | 314  | 
val vars' = filter_out (fn v =>  | 
| 20664 | 315  | 
member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars;  | 
| 16458 | 316  | 
val T = etype_of thy' vs [] prop;  | 
| 13402 | 317  | 
val (T', thw) = Type.freeze_thaw_type  | 
| 13732 | 318  | 
(if T = nullT then nullT else map fastype_of vars' ---> T);  | 
| 27251 | 319  | 
val t = map_types thw (OldGoals.simple_read_term thy' T' s1);  | 
| 16458 | 320  | 
val r' = freeze_thaw (condrew thy' eqns  | 
321  | 
(procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))  | 
|
| 13402 | 322  | 
          (Const ("realizes", T --> propT --> propT) $
 | 
| 13732 | 323  | 
(if T = nullT then t else list_comb (t, vars')) $ prop);  | 
| 27330 | 324  | 
val r = fold_rev Logic.all (map (get_var_type r') vars) r';  | 
| 16458 | 325  | 
val prf = Reconstruct.reconstruct_proof thy' r (rd s2);  | 
| 13402 | 326  | 
in (name, (vs, (t, prf))) end  | 
327  | 
end;  | 
|
328  | 
||
329  | 
val add_realizers_i = gen_add_realizers  | 
|
330  | 
(fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf))));  | 
|
331  | 
val add_realizers = gen_add_realizers prep_realizer;  | 
|
332  | 
||
| 13714 | 333  | 
fun realizes_of thy vs t prop =  | 
334  | 
let  | 
|
335  | 
val thy' = add_syntax thy;  | 
|
| 13732 | 336  | 
    val {realizes_eqns, typeof_eqns, defs, types, ...} =
 | 
| 13714 | 337  | 
ExtractionData.get thy';  | 
| 22717 | 338  | 
val procs = maps (rev o fst o snd) types;  | 
| 16800 | 339  | 
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);  | 
| 17203 | 340  | 
val prop' = Pattern.rewrite_term thy'  | 
| 13714 | 341  | 
(map (Logic.dest_equals o prop_of) defs) [] prop;  | 
| 16458 | 342  | 
in freeze_thaw (condrew thy' eqns  | 
343  | 
(procs @ [typeof_proc (Sign.defaultS thy') vs, rlz_proc]))  | 
|
| 13714 | 344  | 
      (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop')
 | 
345  | 
end;  | 
|
346  | 
||
| 13402 | 347  | 
(** expanding theorems / definitions **)  | 
348  | 
||
| 18728 | 349  | 
fun add_expand_thm thm thy =  | 
| 13402 | 350  | 
let  | 
351  | 
    val {realizes_eqns, typeof_eqns, types, realizers,
 | 
|
352  | 
defs, expand, prep} = ExtractionData.get thy;  | 
|
353  | 
||
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
20897 
diff
changeset
 | 
354  | 
val name = Thm.get_name thm;  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21646 
diff
changeset
 | 
355  | 
val _ = name <> "" orelse error "add_expand_thms: unnamed theorem";  | 
| 13402 | 356  | 
|
357  | 
val is_def =  | 
|
358  | 
(case strip_comb (fst (Logic.dest_equals (prop_of thm))) of  | 
|
| 
18928
 
042608ffa2ec
subsituted gen_duplicates / has_duplicates for duplicates whenever appropriate
 
haftmann 
parents: 
18921 
diff
changeset
 | 
359  | 
(Const _, ts) => forall is_Var ts andalso not (has_duplicates (op =) ts)  | 
| 28674 | 360  | 
andalso (Thm.get_kind thm = Thm.definitionK orelse can (Thm.axiom thy) name)  | 
| 13402 | 361  | 
| _ => false) handle TERM _ => false;  | 
362  | 
in  | 
|
363  | 
(ExtractionData.put (if is_def then  | 
|
364  | 
        {realizes_eqns = realizes_eqns,
 | 
|
365  | 
typeof_eqns = add_rule (([],  | 
|
366  | 
Logic.dest_equals (prop_of (Drule.abs_def thm))), typeof_eqns),  | 
|
367  | 
types = types,  | 
|
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
21858 
diff
changeset
 | 
368  | 
realizers = realizers, defs = insert Thm.eq_thm thm defs,  | 
| 13402 | 369  | 
expand = expand, prep = prep}  | 
370  | 
else  | 
|
371  | 
        {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
 | 
|
372  | 
realizers = realizers, defs = defs,  | 
|
| 20854 | 373  | 
expand = insert (op =) (name, prop_of thm) expand, prep = prep}) thy)  | 
| 13402 | 374  | 
end;  | 
375  | 
||
| 18728 | 376  | 
val add_expand_thms = fold add_expand_thm;  | 
377  | 
||
378  | 
val extraction_expand =  | 
|
| 20897 | 379  | 
Attrib.no_args (Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm th) I));  | 
| 13402 | 380  | 
|
| 15801 | 381  | 
|
| 13732 | 382  | 
(** types with computational content **)  | 
383  | 
||
384  | 
fun add_types tys thy =  | 
|
| 22717 | 385  | 
ExtractionData.map  | 
386  | 
    (fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =>
 | 
|
| 13732 | 387  | 
      {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns,
 | 
| 22717 | 388  | 
types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types,  | 
389  | 
realizers = realizers, defs = defs, expand = expand, prep = prep})  | 
|
390  | 
thy;  | 
|
| 13732 | 391  | 
|
| 13402 | 392  | 
|
| 15801 | 393  | 
(** Pure setup **)  | 
394  | 
||
| 26463 | 395  | 
val _ = Context.>> (Context.map_theory  | 
| 18708 | 396  | 
  (add_types [("prop", ([], NONE))] #>
 | 
| 15801 | 397  | 
|
398  | 
add_typeof_eqns  | 
|
399  | 
["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \  | 
|
400  | 
    \  (typeof (PROP Q)) == (Type (TYPE('Q))) ==>  \
 | 
|
401  | 
    \    (typeof (PROP P ==> PROP Q)) == (Type (TYPE('Q)))",
 | 
|
402  | 
||
403  | 
"(typeof (PROP Q)) == (Type (TYPE(Null))) ==> \  | 
|
404  | 
\ (typeof (PROP P ==> PROP Q)) == (Type (TYPE(Null)))",  | 
|
405  | 
||
406  | 
      "(typeof (PROP P)) == (Type (TYPE('P))) ==>  \
 | 
|
407  | 
    \  (typeof (PROP Q)) == (Type (TYPE('Q))) ==>  \
 | 
|
408  | 
    \    (typeof (PROP P ==> PROP Q)) == (Type (TYPE('P => 'Q)))",
 | 
|
409  | 
||
410  | 
"(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \  | 
|
411  | 
\ (typeof (!!x. PROP P (x))) == (Type (TYPE(Null)))",  | 
|
412  | 
||
413  | 
      "(%x. typeof (PROP P (x))) == (%x. Type (TYPE('P))) ==>  \
 | 
|
414  | 
    \    (typeof (!!x::'a. PROP P (x))) == (Type (TYPE('a => 'P)))",
 | 
|
415  | 
||
416  | 
      "(%x. typeof (f (x))) == (%x. Type (TYPE('f))) ==>  \
 | 
|
| 18708 | 417  | 
    \    (typeof (f)) == (Type (TYPE('f)))"] #>
 | 
| 15801 | 418  | 
|
419  | 
add_realizes_eqns  | 
|
420  | 
["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \  | 
|
421  | 
\ (realizes (r) (PROP P ==> PROP Q)) == \  | 
|
422  | 
\ (PROP realizes (Null) (PROP P) ==> PROP realizes (r) (PROP Q))",  | 
|
423  | 
||
424  | 
      "(typeof (PROP P)) == (Type (TYPE('P))) ==>  \
 | 
|
425  | 
\ (typeof (PROP Q)) == (Type (TYPE(Null))) ==> \  | 
|
426  | 
\ (realizes (r) (PROP P ==> PROP Q)) == \  | 
|
427  | 
\ (!!x::'P. PROP realizes (x) (PROP P) ==> PROP realizes (Null) (PROP Q))",  | 
|
428  | 
||
429  | 
"(realizes (r) (PROP P ==> PROP Q)) == \  | 
|
430  | 
\ (!!x. PROP realizes (x) (PROP P) ==> PROP realizes (r (x)) (PROP Q))",  | 
|
431  | 
||
432  | 
"(%x. typeof (PROP P (x))) == (%x. Type (TYPE(Null))) ==> \  | 
|
433  | 
\ (realizes (r) (!!x. PROP P (x))) == \  | 
|
434  | 
\ (!!x. PROP realizes (Null) (PROP P (x)))",  | 
|
435  | 
||
436  | 
"(realizes (r) (!!x. PROP P (x))) == \  | 
|
| 18708 | 437  | 
\ (!!x. PROP realizes (r (x)) (PROP P (x)))"] #>  | 
| 15801 | 438  | 
|
439  | 
Attrib.add_attributes  | 
|
| 18728 | 440  | 
     [("extraction_expand", extraction_expand,
 | 
| 26463 | 441  | 
"specify theorems / definitions to be expanded during extraction")]));  | 
| 15801 | 442  | 
|
443  | 
||
| 13402 | 444  | 
(**** extract program ****)  | 
445  | 
||
446  | 
val dummyt = Const ("dummy", dummyT);
 | 
|
447  | 
||
448  | 
fun extract thms thy =  | 
|
449  | 
let  | 
|
| 16458 | 450  | 
val thy' = add_syntax thy;  | 
| 13402 | 451  | 
    val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} =
 | 
452  | 
ExtractionData.get thy;  | 
|
| 22717 | 453  | 
val procs = maps (rev o fst o snd) types;  | 
| 13732 | 454  | 
val rtypes = map fst types;  | 
| 16458 | 455  | 
val typroc = typeof_proc (Sign.defaultS thy');  | 
| 19466 | 456  | 
val prep = the_default (K I) prep thy' o ProofRewriteRules.elim_defs thy' false defs o  | 
| 16458 | 457  | 
      Reconstruct.expand_proof thy' (("", NONE) :: map (apsnd SOME) expand);
 | 
| 16800 | 458  | 
val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);  | 
| 13402 | 459  | 
|
460  | 
fun find_inst prop Ts ts vs =  | 
|
461  | 
let  | 
|
| 13732 | 462  | 
val rvs = relevant_vars rtypes prop;  | 
| 13402 | 463  | 
val vars = vars_of prop;  | 
464  | 
val n = Int.min (length vars, length ts);  | 
|
465  | 
||
466  | 
fun add_args ((Var ((a, i), _), t), (vs', tye)) =  | 
|
| 20664 | 467  | 
if member (op =) rvs a then  | 
| 16458 | 468  | 
let val T = etype_of thy' vs Ts t  | 
| 13402 | 469  | 
in if T = nullT then (vs', tye)  | 
470  | 
               else (a :: vs', (("'" ^ a, i), T) :: tye)
 | 
|
471  | 
end  | 
|
472  | 
else (vs', tye)  | 
|
473  | 
||
| 23178 | 474  | 
in List.foldr add_args ([], []) (Library.take (n, vars) ~~ Library.take (n, ts)) end;  | 
| 13402 | 475  | 
|
| 20664 | 476  | 
fun find (vs: string list) = Option.map snd o find_first (curry (gen_eq_set (op =)) vs o fst);  | 
| 28375 | 477  | 
fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE);  | 
| 13402 | 478  | 
|
| 13732 | 479  | 
fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw  | 
| 16458 | 480  | 
(condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (list_abs  | 
| 13732 | 481  | 
(map (pair "x") (rev Ts), t)));  | 
482  | 
||
483  | 
fun realizes_null vs prop = app_rlz_rews [] vs  | 
|
484  | 
      (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop);
 | 
|
| 13402 | 485  | 
|
486  | 
fun corr d defs vs ts Ts hs (PBound i) _ _ = (defs, PBound i)  | 
|
487  | 
||
| 15531 | 488  | 
| corr d defs vs ts Ts hs (Abst (s, SOME T, prf)) (Abst (_, _, prf')) t =  | 
| 13402 | 489  | 
let val (defs', corr_prf) = corr d defs vs [] (T :: Ts)  | 
490  | 
(dummyt :: hs) prf (incr_pboundvars 1 0 prf')  | 
|
| 15531 | 491  | 
(case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE)  | 
492  | 
in (defs', Abst (s, SOME T, corr_prf)) end  | 
|
| 13402 | 493  | 
|
| 15531 | 494  | 
| corr d defs vs ts Ts hs (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) t =  | 
| 13402 | 495  | 
let  | 
| 16458 | 496  | 
val T = etype_of thy' vs Ts prop;  | 
| 13402 | 497  | 
val u = if T = nullT then  | 
| 15531 | 498  | 
(case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE)  | 
499  | 
else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE);  | 
|
| 13402 | 500  | 
val (defs', corr_prf) = corr d defs vs [] (T :: Ts) (prop :: hs)  | 
501  | 
(incr_pboundvars 0 1 prf) (incr_pboundvars 0 1 prf') u;  | 
|
502  | 
            val rlz = Const ("realizes", T --> propT --> propT)
 | 
|
503  | 
in (defs',  | 
|
| 13732 | 504  | 
            if T = nullT then AbsP ("R",
 | 
| 15531 | 505  | 
SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)),  | 
| 13732 | 506  | 
prf_subst_bounds [nullt] corr_prf)  | 
| 15531 | 507  | 
            else Abst (s, SOME T, AbsP ("R",
 | 
508  | 
SOME (app_rlz_rews (T :: Ts) vs  | 
|
| 13732 | 509  | 
(rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)))  | 
| 13402 | 510  | 
end  | 
511  | 
||
| 15531 | 512  | 
| corr d defs vs ts Ts hs (prf % SOME t) (prf' % _) t' =  | 
| 13732 | 513  | 
let  | 
514  | 
val (Us, T) = strip_type (fastype_of1 (Ts, t));  | 
|
515  | 
val (defs', corr_prf) = corr d defs vs (t :: ts) Ts hs prf prf'  | 
|
| 20664 | 516  | 
(if member (op =) rtypes (tname_of T) then t'  | 
| 15531 | 517  | 
else (case t' of SOME (u $ _) => SOME u | _ => NONE));  | 
| 20664 | 518  | 
val u = if not (member (op =) rtypes (tname_of T)) then t else  | 
| 13732 | 519  | 
let  | 
| 16458 | 520  | 
val eT = etype_of thy' vs Ts t;  | 
| 13732 | 521  | 
val (r, Us') = if eT = nullT then (nullt, Us) else  | 
522  | 
(Bound (length Us), eT :: Us);  | 
|
523  | 
val u = list_comb (incr_boundvars (length Us') t,  | 
|
524  | 
map Bound (length Us - 1 downto 0));  | 
|
| 17271 | 525  | 
val u' = (case AList.lookup (op =) types (tname_of T) of  | 
| 15531 | 526  | 
SOME ((_, SOME f)) => f r eT u T  | 
| 13732 | 527  | 
                  | _ => Const ("realizes", eT --> T --> T) $ r $ u)
 | 
528  | 
in app_rlz_rews Ts vs (list_abs (map (pair "x") Us', u')) end  | 
|
| 15531 | 529  | 
in (defs', corr_prf % SOME u) end  | 
| 13402 | 530  | 
|
531  | 
| corr d defs vs ts Ts hs (prf1 %% prf2) (prf1' %% prf2') t =  | 
|
532  | 
let  | 
|
533  | 
val prop = Reconstruct.prop_of' hs prf2';  | 
|
| 16458 | 534  | 
val T = etype_of thy' vs Ts prop;  | 
| 15531 | 535  | 
val (defs1, f, u) = if T = nullT then (defs, t, NONE) else  | 
| 13402 | 536  | 
(case t of  | 
| 15531 | 537  | 
SOME (f $ u) => (defs, SOME f, SOME u)  | 
| 13402 | 538  | 
| _ =>  | 
539  | 
let val (defs1, u) = extr d defs vs [] Ts hs prf2'  | 
|
| 15531 | 540  | 
in (defs1, NONE, SOME u) end)  | 
| 13402 | 541  | 
val (defs2, corr_prf1) = corr d defs1 vs [] Ts hs prf1 prf1' f;  | 
542  | 
val (defs3, corr_prf2) = corr d defs2 vs [] Ts hs prf2 prf2' u;  | 
|
543  | 
in  | 
|
544  | 
if T = nullT then (defs3, corr_prf1 %% corr_prf2) else  | 
|
545  | 
(defs3, corr_prf1 % u %% corr_prf2)  | 
|
546  | 
end  | 
|
547  | 
||
| 28805 | 548  | 
| corr d defs vs ts Ts hs (prf0 as PThm (_, ((name, prop, SOME Ts'), body))) _ _ =  | 
| 13402 | 549  | 
let  | 
| 28805 | 550  | 
val prf = force_proof body;  | 
| 13402 | 551  | 
val (vs', tye) = find_inst prop Ts ts vs;  | 
552  | 
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;  | 
|
| 16458 | 553  | 
val T = etype_of thy' vs' [] prop;  | 
| 13402 | 554  | 
val defs' = if T = nullT then defs  | 
555  | 
else fst (extr d defs vs ts Ts hs prf0)  | 
|
556  | 
in  | 
|
| 
13609
 
73c3915553b4
Added check for axioms with "realizes Null A = A".
 
berghofe 
parents: 
13417 
diff
changeset
 | 
557  | 
if T = nullT andalso realizes_null vs' prop aconv prop then (defs, prf0)  | 
| 17412 | 558  | 
else case Symtab.lookup realizers name of  | 
| 15531 | 559  | 
NONE => (case find vs' (find' name defs') of  | 
560  | 
NONE =>  | 
|
| 13402 | 561  | 
let  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21646 
diff
changeset
 | 
562  | 
val _ = T = nullT orelse error "corr: internal error";  | 
| 13402 | 563  | 
                    val _ = msg d ("Building correctness proof for " ^ quote name ^
 | 
564  | 
(if null vs' then ""  | 
|
565  | 
else " (relevant variables: " ^ commas_quote vs' ^ ")"));  | 
|
| 16458 | 566  | 
val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);  | 
| 13402 | 567  | 
val (defs'', corr_prf) =  | 
| 15531 | 568  | 
corr (d + 1) defs' vs' [] [] [] prf' prf' NONE;  | 
| 13732 | 569  | 
val corr_prop = Reconstruct.prop_of corr_prf;  | 
| 23178 | 570  | 
val corr_prf' = List.foldr forall_intr_prf  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
571  | 
(proof_combt  | 
| 28805 | 572  | 
(PThm (serial (),  | 
573  | 
((corr_name name vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),  | 
|
574  | 
Lazy.value (make_proof_body corr_prf))), vfs_of corr_prop))  | 
|
| 22750 | 575  | 
(map (get_var_type corr_prop) (vfs_of prop))  | 
| 13402 | 576  | 
in  | 
| 13732 | 577  | 
((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',  | 
| 13402 | 578  | 
prf_subst_TVars tye' corr_prf')  | 
579  | 
end  | 
|
| 15531 | 580  | 
| SOME (_, (_, prf')) => (defs', prf_subst_TVars tye' prf'))  | 
581  | 
| SOME rs => (case find vs' rs of  | 
|
582  | 
SOME (_, prf') => (defs', prf_subst_TVars tye' prf')  | 
|
583  | 
              | 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
 | 
584  | 
quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm  | 
| 13402 | 585  | 
(Reconstruct.prop_of (proof_combt (prf0, ts))))))  | 
586  | 
end  | 
|
587  | 
||
| 15531 | 588  | 
| corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =  | 
| 13402 | 589  | 
let  | 
590  | 
val (vs', tye) = find_inst prop Ts ts vs;  | 
|
591  | 
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye  | 
|
592  | 
in  | 
|
| 16458 | 593  | 
if etype_of thy' vs' [] prop = nullT andalso  | 
| 
13609
 
73c3915553b4
Added check for axioms with "realizes Null A = A".
 
berghofe 
parents: 
13417 
diff
changeset
 | 
594  | 
realizes_null vs' prop aconv prop then (defs, prf0)  | 
| 18956 | 595  | 
else case find vs' (Symtab.lookup_list realizers s) of  | 
| 15531 | 596  | 
SOME (_, prf) => (defs, prf_subst_TVars tye' prf)  | 
597  | 
            | 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
 | 
598  | 
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm  | 
| 13402 | 599  | 
(Reconstruct.prop_of (proof_combt (prf0, ts)))))  | 
600  | 
end  | 
|
601  | 
||
602  | 
| corr d defs vs ts Ts hs _ _ _ = error "corr: bad proof"  | 
|
603  | 
||
604  | 
and extr d defs vs ts Ts hs (PBound i) = (defs, Bound i)  | 
|
605  | 
||
| 15531 | 606  | 
| extr d defs vs ts Ts hs (Abst (s, SOME T, prf)) =  | 
| 13402 | 607  | 
let val (defs', t) = extr d defs vs []  | 
608  | 
(T :: Ts) (dummyt :: hs) (incr_pboundvars 1 0 prf)  | 
|
609  | 
in (defs', Abs (s, T, t)) end  | 
|
610  | 
||
| 15531 | 611  | 
| extr d defs vs ts Ts hs (AbsP (s, SOME t, prf)) =  | 
| 13402 | 612  | 
let  | 
| 16458 | 613  | 
val T = etype_of thy' vs Ts t;  | 
| 13402 | 614  | 
val (defs', t) = extr d defs vs [] (T :: Ts) (t :: hs)  | 
615  | 
(incr_pboundvars 0 1 prf)  | 
|
616  | 
in (defs',  | 
|
617  | 
if T = nullT then subst_bound (nullt, t) else Abs (s, T, t))  | 
|
618  | 
end  | 
|
619  | 
||
| 15531 | 620  | 
| extr d defs vs ts Ts hs (prf % SOME t) =  | 
| 13402 | 621  | 
let val (defs', u) = extr d defs vs (t :: ts) Ts hs prf  | 
| 13732 | 622  | 
in (defs',  | 
| 20664 | 623  | 
if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u  | 
| 13732 | 624  | 
else u $ t)  | 
625  | 
end  | 
|
| 13402 | 626  | 
|
627  | 
| extr d defs vs ts Ts hs (prf1 %% prf2) =  | 
|
628  | 
let  | 
|
629  | 
val (defs', f) = extr d defs vs [] Ts hs prf1;  | 
|
630  | 
val prop = Reconstruct.prop_of' hs prf2;  | 
|
| 16458 | 631  | 
val T = etype_of thy' vs Ts prop  | 
| 13402 | 632  | 
in  | 
633  | 
if T = nullT then (defs', f) else  | 
|
634  | 
let val (defs'', t) = extr d defs' vs [] Ts hs prf2  | 
|
635  | 
in (defs'', f $ t) end  | 
|
636  | 
end  | 
|
637  | 
||
| 28805 | 638  | 
| extr d defs vs ts Ts hs (prf0 as PThm (_, ((s, prop, SOME Ts'), body))) =  | 
| 13402 | 639  | 
let  | 
| 28805 | 640  | 
val prf = force_proof body;  | 
| 13402 | 641  | 
val (vs', tye) = find_inst prop Ts ts vs;  | 
642  | 
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye  | 
|
643  | 
in  | 
|
| 17412 | 644  | 
case Symtab.lookup realizers s of  | 
| 15531 | 645  | 
NONE => (case find vs' (find' s defs) of  | 
646  | 
NONE =>  | 
|
| 13402 | 647  | 
let  | 
648  | 
                    val _ = msg d ("Extracting " ^ quote s ^
 | 
|
649  | 
(if null vs' then ""  | 
|
650  | 
else " (relevant variables: " ^ commas_quote vs' ^ ")"));  | 
|
| 16458 | 651  | 
val prf' = prep (Reconstruct.reconstruct_proof thy' prop prf);  | 
| 13402 | 652  | 
val (defs', t) = extr (d + 1) defs vs' [] [] [] prf';  | 
653  | 
val (defs'', corr_prf) =  | 
|
| 15531 | 654  | 
corr (d + 1) defs' vs' [] [] [] prf' prf' (SOME t);  | 
| 13402 | 655  | 
|
656  | 
val nt = Envir.beta_norm t;  | 
|
| 20664 | 657  | 
val args = filter_out (fn v => member (op =) rtypes  | 
658  | 
(tname_of (body_type (fastype_of v)))) (vfs_of prop);  | 
|
| 15570 | 659  | 
val args' = List.filter (fn v => Logic.occs (v, nt)) args;  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
660  | 
val t' = mkabs nt args';  | 
| 13402 | 661  | 
val T = fastype_of t';  | 
| 13732 | 662  | 
val cname = extr_name s vs';  | 
| 13402 | 663  | 
val c = Const (cname, T);  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
664  | 
val u = mkabs (list_comb (c, args')) args;  | 
| 13402 | 665  | 
val eqn = Logic.mk_equals (c, t');  | 
666  | 
val rlz =  | 
|
667  | 
                      Const ("realizes", fastype_of nt --> propT --> propT);
 | 
|
| 13732 | 668  | 
val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop);  | 
669  | 
val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop);  | 
|
670  | 
val f = app_rlz_rews [] vs'  | 
|
671  | 
                      (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop));
 | 
|
| 13402 | 672  | 
|
| 13732 | 673  | 
val corr_prf' =  | 
674  | 
chtype [] equal_elim_axm %> lhs %> rhs %%  | 
|
675  | 
(chtype [propT] symmetric_axm %> rhs %> lhs %%  | 
|
676  | 
(chtype [propT, T] combination_axm %> f %> f %> c %> t' %%  | 
|
677  | 
(chtype [T --> propT] reflexive_axm %> f) %%  | 
|
678  | 
PAxm (cname ^ "_def", eqn,  | 
|
| 15531 | 679  | 
SOME (map TVar (term_tvars eqn))))) %% corr_prf;  | 
| 13732 | 680  | 
val corr_prop = Reconstruct.prop_of corr_prf';  | 
| 23178 | 681  | 
val corr_prf'' = List.foldr forall_intr_prf  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
682  | 
(proof_combt  | 
| 28805 | 683  | 
(PThm (serial (),  | 
684  | 
((corr_name s vs', corr_prop, SOME (map TVar (term_tvars corr_prop))),  | 
|
685  | 
Lazy.value (make_proof_body corr_prf'))), vfs_of corr_prop))  | 
|
| 22750 | 686  | 
(map (get_var_type corr_prop) (vfs_of prop));  | 
| 13402 | 687  | 
in  | 
| 13732 | 688  | 
((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',  | 
| 13402 | 689  | 
subst_TVars tye' u)  | 
690  | 
end  | 
|
| 15531 | 691  | 
| SOME ((_, u), _) => (defs, subst_TVars tye' u))  | 
692  | 
| SOME rs => (case find vs' rs of  | 
|
693  | 
SOME (t, _) => (defs, subst_TVars tye' t)  | 
|
694  | 
              | 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
 | 
695  | 
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm  | 
| 13402 | 696  | 
(Reconstruct.prop_of (proof_combt (prf0, ts))))))  | 
697  | 
end  | 
|
698  | 
||
| 15531 | 699  | 
| extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =  | 
| 13402 | 700  | 
let  | 
701  | 
val (vs', tye) = find_inst prop Ts ts vs;  | 
|
702  | 
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye  | 
|
703  | 
in  | 
|
| 18956 | 704  | 
case find vs' (Symtab.lookup_list realizers s) of  | 
| 15531 | 705  | 
SOME (t, _) => (defs, subst_TVars tye' t)  | 
706  | 
            | 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
 | 
707  | 
quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm  | 
| 13402 | 708  | 
(Reconstruct.prop_of (proof_combt (prf0, ts)))))  | 
709  | 
end  | 
|
710  | 
||
711  | 
| extr d defs vs ts Ts hs _ = error "extr: bad proof";  | 
|
712  | 
||
| 13732 | 713  | 
fun prep_thm (thm, vs) =  | 
| 13402 | 714  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26481 
diff
changeset
 | 
715  | 
val thy = Thm.theory_of_thm thm;  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26481 
diff
changeset
 | 
716  | 
val prop = Thm.prop_of thm;  | 
| 28814 | 717  | 
val prf = Thm.proof_of thm;  | 
| 
21646
 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 
wenzelm 
parents: 
20897 
diff
changeset
 | 
718  | 
val name = Thm.get_name thm;  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21646 
diff
changeset
 | 
719  | 
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
 | 
720  | 
        val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^
 | 
| 13402 | 721  | 
quote name ^ " has no computational content")  | 
| 22596 | 722  | 
in (Reconstruct.reconstruct_proof thy prop prf, vs) end;  | 
| 13402 | 723  | 
|
| 15570 | 724  | 
val defs = Library.foldl (fn (defs, (prf, vs)) =>  | 
| 13732 | 725  | 
fst (extr 0 defs vs [] [] [] prf)) ([], map prep_thm thms);  | 
| 13402 | 726  | 
|
| 16149 | 727  | 
fun add_def (s, (vs, ((t, u), (prf, _)))) thy =  | 
| 16458 | 728  | 
(case Sign.const_type thy (extr_name s vs) of  | 
| 15531 | 729  | 
NONE =>  | 
| 13732 | 730  | 
let  | 
731  | 
val corr_prop = Reconstruct.prop_of prf;  | 
|
| 16287 | 732  | 
val ft = Type.freeze t;  | 
733  | 
val fu = Type.freeze u;  | 
|
| 22750 | 734  | 
val (def_thms, thy') = if t = nullt then ([], thy) else  | 
735  | 
thy  | 
|
| 22796 | 736  | 
|> Sign.add_consts_i [(extr_name s vs, fastype_of ft, NoSyn)]  | 
| 
27691
 
ce171cbd4b93
PureThy: dropped note_thmss_qualified, dropped _i suffix
 
haftmann 
parents: 
27330 
diff
changeset
 | 
737  | 
|> PureThy.add_defs false [((extr_name s vs ^ "_def",  | 
| 22750 | 738  | 
Logic.mk_equals (head_of (strip_abs_body fu), ft)), [])]  | 
| 13732 | 739  | 
in  | 
| 22750 | 740  | 
thy'  | 
| 26481 | 741  | 
|> PureThy.store_thm (corr_name s vs,  | 
| 22750 | 742  | 
Thm.varifyT (funpow (length (term_vars corr_prop))  | 
| 26653 | 743  | 
(Thm.forall_elim_var 0) (forall_intr_frees  | 
| 22750 | 744  | 
(ProofChecker.thm_of_proof thy'  | 
| 26481 | 745  | 
(fst (Proofterm.freeze_thaw_prf prf))))))  | 
| 22750 | 746  | 
|> snd  | 
| 28370 | 747  | 
|> fold Code.add_default_eqn def_thms  | 
| 13732 | 748  | 
end  | 
| 15531 | 749  | 
| SOME _ => thy);  | 
| 13402 | 750  | 
|
| 16149 | 751  | 
in  | 
752  | 
thy  | 
|
| 22796 | 753  | 
|> Sign.absolute_path  | 
| 16149 | 754  | 
|> fold_rev add_def defs  | 
| 22796 | 755  | 
|> Sign.restore_naming thy  | 
| 13402 | 756  | 
end;  | 
757  | 
||
758  | 
||
759  | 
(**** interface ****)  | 
|
760  | 
||
| 17057 | 761  | 
structure P = OuterParse and K = OuterKeyword;  | 
| 13402 | 762  | 
|
| 13732 | 763  | 
val parse_vars = Scan.optional (P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") [];
 | 
764  | 
||
| 24867 | 765  | 
val _ =  | 
| 13402 | 766  | 
OuterSyntax.command "realizers"  | 
767  | 
"specify realizers for primitive axioms / theorems, together with correctness proof"  | 
|
768  | 
K.thy_decl  | 
|
| 13732 | 769  | 
(Scan.repeat1 (P.xname -- parse_vars --| P.$$$ ":" -- P.string -- P.string) >>  | 
| 13402 | 770  | 
(fn xs => Toplevel.theory (fn thy => add_realizers  | 
| 
26343
 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
771  | 
(map (fn (((a, vs), s1), s2) => (PureThy.get_thm thy a, (vs, s1, s2))) xs) thy)));  | 
| 13402 | 772  | 
|
| 24867 | 773  | 
val _ =  | 
| 13402 | 774  | 
OuterSyntax.command "realizability"  | 
775  | 
"add equations characterizing realizability" K.thy_decl  | 
|
776  | 
(Scan.repeat1 P.string >> (Toplevel.theory o add_realizes_eqns));  | 
|
777  | 
||
| 24867 | 778  | 
val _ =  | 
| 13402 | 779  | 
OuterSyntax.command "extract_type"  | 
780  | 
"add equations characterizing type of extracted program" K.thy_decl  | 
|
781  | 
(Scan.repeat1 P.string >> (Toplevel.theory o add_typeof_eqns));  | 
|
782  | 
||
| 24867 | 783  | 
val _ =  | 
| 13402 | 784  | 
OuterSyntax.command "extract" "extract terms from proofs" K.thy_decl  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25389 
diff
changeset
 | 
785  | 
(Scan.repeat1 (P.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>  | 
| 
26343
 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
786  | 
extract (map (apfst (PureThy.get_thm thy)) xs) thy)));  | 
| 13402 | 787  | 
|
| 16458 | 788  | 
val etype_of = etype_of o add_syntax;  | 
| 13714 | 789  | 
|
| 13402 | 790  | 
end;  |