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