author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 81534 | c32ebdcbe8ca |
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 |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
14 |
val add_realizers_i : (Thm_Name.T * (string list * term * Proofterm.proof)) list |
13402 | 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 |
79242 | 103 |
fun ren t = perhaps (Term.rename_abs tm1 tm) t; |
79232 | 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 |
||
70806 | 119 |
val change_types = Proofterm.change_types o SOME; |
13402 | 120 |
|
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
121 |
fun extr_name thm_name vs = |
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
122 |
Long_Name.append "extr" (space_implode "_" (Thm_Name.short thm_name :: vs)); |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
123 |
|
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
124 |
fun corr_name thm_name vs = |
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
125 |
extr_name thm_name vs ^ "_correctness"; |
13402 | 126 |
|
61865 | 127 |
fun msg d s = writeln (Symbol.spaces d ^ s); |
13402 | 128 |
|
74235 | 129 |
fun vars_of t = map Var (build_rev (Term.add_vars t)); |
130 |
fun frees_of t = map Free (build_rev (Term.add_frees t)); |
|
28812
413695e07bd4
Frees in PThms are now quantified in the order of their appearance in the
berghofe
parents:
28805
diff
changeset
|
131 |
fun vfs_of t = vars_of t @ frees_of t; |
13402 | 132 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
133 |
val mkabs = fold_rev (fn v => fn t => Abs ("x", fastype_of v, abstract_over (v, t))); |
13402 | 134 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
135 |
val mkabsp = fold_rev (fn t => fn prf => AbsP ("H", SOME t, prf)); |
13402 | 136 |
|
13732 | 137 |
fun strip_abs 0 t = t |
138 |
| strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t |
|
139 |
| strip_abs _ _ = error "strip_abs: not an abstraction"; |
|
140 |
||
37310 | 141 |
val prf_subst_TVars = Proofterm.map_proof_types o typ_subst_TVars; |
13402 | 142 |
|
40844 | 143 |
fun relevant_vars types prop = |
144 |
List.foldr |
|
145 |
(fn (Var ((a, _), T), vs) => |
|
146 |
(case body_type T of |
|
147 |
Type (s, _) => if member (op =) types s then a :: vs else vs |
|
148 |
| _ => vs) |
|
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 []; |
|
42406
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents:
42375
diff
changeset
|
158 |
in |
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents:
42375
diff
changeset
|
159 |
fn Var (ixn, _) => |
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents:
42375
diff
changeset
|
160 |
(case AList.lookup (op =) vs ixn of |
15531 | 161 |
NONE => error "get_var_type: no such variable in term" |
162 |
| SOME T => Var (ixn, T)) |
|
42406
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents:
42375
diff
changeset
|
163 |
| Free (s, _) => |
05f2468d6b36
eliminated obsolete Proof_Syntax.strip_sorts_consttypes;
wenzelm
parents:
42375
diff
changeset
|
164 |
(case AList.lookup (op =) fs s of |
15531 | 165 |
NONE => error "get_var_type: no such variable in term" |
166 |
| SOME T => Free (s, T)) |
|
13732 | 167 |
| _ => error "get_var_type: not a variable" |
168 |
end; |
|
169 |
||
62922 | 170 |
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
|
171 |
let |
62922 | 172 |
val ctxt' = ctxt |
62958
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62922
diff
changeset
|
173 |
|> Proof_Context.set_defsort [] |
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62922
diff
changeset
|
174 |
|> Config.put Type_Infer.object_logic false |
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62922
diff
changeset
|
175 |
|> 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
|
176 |
val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; |
62922 | 177 |
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
|
178 |
|
70589 | 179 |
fun make_proof_body prf = |
180 |
let |
|
181 |
val (oracles, thms) = |
|
182 |
([prf], ([], [])) |-> Proofterm.fold_proof_atoms false |
|
71465
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents:
71214
diff
changeset
|
183 |
(fn Oracle (name, prop, _) => apfst (cons ((name, Position.none), SOME prop)) |
70589 | 184 |
| PThm (header, thm_body) => apsnd (cons (Proofterm.make_thm header thm_body)) |
185 |
| _ => I); |
|
77866 | 186 |
val body = |
187 |
PBody |
|
77867 | 188 |
{oracles = Ord_List.make Proofterm.oracle_ord oracles, |
77866 | 189 |
thms = Ord_List.make Proofterm.thm_ord thms, |
79279 | 190 |
zboxes = [], |
80560 | 191 |
zproof = ZNop, |
79220 | 192 |
proof = prf}; |
70589 | 193 |
in Proofterm.thm_body body end; |
194 |
||
195 |
||
13402 | 196 |
|
197 |
(**** theory data ****) |
|
198 |
||
22846 | 199 |
(* theory data *) |
13402 | 200 |
|
33522 | 201 |
structure ExtractionData = Theory_Data |
22846 | 202 |
( |
13402 | 203 |
type T = |
204 |
{realizes_eqns : rules, |
|
205 |
typeof_eqns : rules, |
|
13732 | 206 |
types : (string * ((term -> term option) list * |
207 |
(term -> typ -> term -> typ -> term) option)) list, |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
208 |
realizers : (string list * (term * proof)) list Thm_Name.Table.table, |
13402 | 209 |
defs : thm list, |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
210 |
expand : Thm_Name.T list, |
16458 | 211 |
prep : (theory -> proof -> proof) option} |
13402 | 212 |
|
213 |
val empty = |
|
214 |
{realizes_eqns = empty_rules, |
|
215 |
typeof_eqns = empty_rules, |
|
216 |
types = [], |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
217 |
realizers = Thm_Name.Table.empty, |
13402 | 218 |
defs = [], |
219 |
expand = [], |
|
15531 | 220 |
prep = NONE}; |
13402 | 221 |
|
33522 | 222 |
fun merge |
223 |
({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1, |
|
13402 | 224 |
realizers = realizers1, defs = defs1, expand = expand1, prep = prep1}, |
225 |
{realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2, |
|
33522 | 226 |
realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T = |
13402 | 227 |
{realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2, |
228 |
typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2, |
|
22717 | 229 |
types = AList.merge (op =) (K true) (types1, types2), |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
230 |
realizers = Thm_Name.Table.merge_list (eq_set (op =) o apply2 #1) (realizers1, realizers2), |
22662 | 231 |
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
|
232 |
expand = Library.merge (op =) (expand1, expand2), |
38761 | 233 |
prep = if is_some prep1 then prep1 else prep2}; |
22846 | 234 |
); |
13402 | 235 |
|
236 |
fun read_condeq thy = |
|
62922 | 237 |
let val ctxt' = Proof_Context.init_global (add_syntax thy) |
13402 | 238 |
in fn s => |
62922 | 239 |
let val t = Logic.varify_global (read_term ctxt' propT s) |
35424 | 240 |
in |
241 |
(map Logic.dest_equals (Logic.strip_imp_prems t), |
|
242 |
Logic.dest_equals (Logic.strip_imp_concl t)) |
|
243 |
handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s) |
|
244 |
end |
|
13402 | 245 |
end; |
246 |
||
247 |
(** preprocessor **) |
|
248 |
||
249 |
fun set_preprocessor prep thy = |
|
250 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
251 |
defs, expand, ...} = ExtractionData.get thy |
|
252 |
in |
|
253 |
ExtractionData.put |
|
254 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
15531 | 255 |
realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy |
13402 | 256 |
end; |
257 |
||
258 |
(** equations characterizing realizability **) |
|
259 |
||
260 |
fun gen_add_realizes_eqns prep_eq eqns thy = |
|
261 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
262 |
defs, expand, prep} = ExtractionData.get thy; |
|
263 |
in |
|
264 |
ExtractionData.put |
|
33337 | 265 |
{realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns, |
13402 | 266 |
typeof_eqns = typeof_eqns, types = types, realizers = realizers, |
267 |
defs = defs, expand = expand, prep = prep} thy |
|
268 |
end |
|
269 |
||
270 |
val add_realizes_eqns_i = gen_add_realizes_eqns (K I); |
|
271 |
val add_realizes_eqns = gen_add_realizes_eqns read_condeq; |
|
272 |
||
273 |
(** equations characterizing type of extracted program **) |
|
274 |
||
275 |
fun gen_add_typeof_eqns prep_eq eqns thy = |
|
276 |
let |
|
277 |
val {realizes_eqns, typeof_eqns, types, realizers, |
|
278 |
defs, expand, prep} = ExtractionData.get thy; |
|
13732 | 279 |
val eqns' = map (prep_eq thy) eqns |
13402 | 280 |
in |
281 |
ExtractionData.put |
|
282 |
{realizes_eqns = realizes_eqns, realizers = realizers, |
|
33337 | 283 |
typeof_eqns = fold_rev add_rule eqns' typeof_eqns, |
13732 | 284 |
types = types, defs = defs, expand = expand, prep = prep} thy |
13402 | 285 |
end |
286 |
||
287 |
val add_typeof_eqns_i = gen_add_typeof_eqns (K I); |
|
288 |
val add_typeof_eqns = gen_add_typeof_eqns read_condeq; |
|
289 |
||
290 |
fun thaw (T as TFree (a, S)) = |
|
77851 | 291 |
if member_string a ":" then TVar (unpack_ixn a, S) else T |
13402 | 292 |
| thaw (Type (a, Ts)) = Type (a, map thaw Ts) |
293 |
| thaw T = T; |
|
294 |
||
295 |
fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S) |
|
296 |
| freeze (Type (a, Ts)) = Type (a, map freeze Ts) |
|
297 |
| freeze T = T; |
|
298 |
||
299 |
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
|
300 |
map_types thaw (f (map_types freeze x)); |
13402 | 301 |
|
16458 | 302 |
fun etype_of thy vs Ts t = |
13402 | 303 |
let |
16458 | 304 |
val {typeof_eqns, ...} = ExtractionData.get thy; |
13402 | 305 |
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
|
306 |
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
|
307 |
in |
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
44057
diff
changeset
|
308 |
(case |
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
44057
diff
changeset
|
309 |
strip_abs_body |
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
44057
diff
changeset
|
310 |
(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
|
311 |
(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
|
312 |
(Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of |
13402 | 313 |
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
|
314 |
| _ => err ()) |
13402 | 315 |
end; |
316 |
||
317 |
(** realizers for axioms / theorems, together with correctness proofs **) |
|
318 |
||
319 |
fun gen_add_realizers prep_rlz rs thy = |
|
320 |
let val {realizes_eqns, typeof_eqns, types, realizers, |
|
321 |
defs, expand, prep} = ExtractionData.get thy |
|
322 |
in |
|
323 |
ExtractionData.put |
|
324 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
325 |
realizers = fold (Thm_Name.Table.cons_list o prep_rlz thy) rs realizers, |
13402 | 326 |
defs = defs, expand = expand, prep = prep} thy |
327 |
end |
|
328 |
||
329 |
fun prep_realizer thy = |
|
330 |
let |
|
13732 | 331 |
val {realizes_eqns, typeof_eqns, defs, types, ...} = |
13402 | 332 |
ExtractionData.get thy; |
19482
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
wenzelm
parents:
19466
diff
changeset
|
333 |
val procs = maps (fst o snd) types; |
13732 | 334 |
val rtypes = map fst types; |
16800 | 335 |
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
13402 | 336 |
val thy' = add_syntax thy; |
62922 | 337 |
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
|
338 |
val rd = Proof_Syntax.read_proof thy' true false; |
13402 | 339 |
in fn (thm, (vs, s1, s2)) => |
340 |
let |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
341 |
val thm_name = Thm.derivation_name thm; |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
342 |
val _ = if Thm_Name.is_empty thm_name then error "add_realizers: unnamed theorem" else (); |
59582 | 343 |
val prop = Thm.unconstrainT thm |> Thm.prop_of |> |
344 |
Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) []; |
|
13402 | 345 |
val vars = vars_of prop; |
13732 | 346 |
val vars' = filter_out (fn v => |
20664 | 347 |
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
|
348 |
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
|
349 |
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
|
350 |
(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
|
351 |
else []) vars; |
16458 | 352 |
val T = etype_of thy' vs [] prop; |
33832 | 353 |
val (T', thw) = Type.legacy_freeze_thaw_type |
13732 | 354 |
(if T = nullT then nullT else map fastype_of vars' ---> T); |
62922 | 355 |
val t = map_types thw (read_term ctxt' T' s1); |
16458 | 356 |
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
|
357 |
(procs @ [typeof_proc [] vs, rlz_proc])) |
13402 | 358 |
(Const ("realizes", T --> propT --> propT) $ |
13732 | 359 |
(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
|
360 |
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
|
361 |
fold_rev Logic.all (map (get_var_type r') vars) r'); |
70449 | 362 |
val prf = Proofterm.reconstruct_proof thy' r (rd s2); |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
363 |
in (thm_name, (vs, (t, prf))) end |
13402 | 364 |
end; |
365 |
||
366 |
val add_realizers_i = gen_add_realizers |
|
367 |
(fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf)))); |
|
368 |
val add_realizers = gen_add_realizers prep_realizer; |
|
369 |
||
13714 | 370 |
fun realizes_of thy vs t prop = |
371 |
let |
|
372 |
val thy' = add_syntax thy; |
|
13732 | 373 |
val {realizes_eqns, typeof_eqns, defs, types, ...} = |
13714 | 374 |
ExtractionData.get thy'; |
22717 | 375 |
val procs = maps (rev o fst o snd) types; |
16800 | 376 |
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
17203 | 377 |
val prop' = Pattern.rewrite_term thy' |
59582 | 378 |
(map (Logic.dest_equals o Thm.prop_of) defs) [] prop; |
16458 | 379 |
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
|
380 |
(procs @ [typeof_proc [] vs, rlz_proc])) |
13714 | 381 |
(Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop') |
382 |
end; |
|
383 |
||
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
384 |
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
|
385 |
let |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
386 |
val S = Sign.defaultS thy; |
70436 | 387 |
val (ucontext, prop') = |
61039 | 388 |
Logic.unconstrainT (Thm.shyps_of thm) (Thm.prop_of thm); |
59582 | 389 |
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
|
390 |
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
|
391 |
SOME (TVar (("'" ^ v, i), [])) else NONE) |
74235 | 392 |
(build_rev (Term.add_vars prop')); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
393 |
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
|
394 |
val constraints' = map Logic.mk_of_class cs; |
79411 | 395 |
val typ_map = Term.strip_sortsT o |
79409 | 396 |
Term.map_atyps (fn U => |
79391 | 397 |
if member (op =) atyps U |
79407 | 398 |
then #typ_operation ucontext {strip_sorts = false} U |
79399 | 399 |
else raise Same.SAME); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
400 |
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
|
401 |
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
|
402 |
in |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
403 |
prf |> |
37310 | 404 |
Same.commit (Proofterm.map_proof_same (map_types typ_map) typ_map mk_hyp) |> |
70436 | 405 |
fold_rev Proofterm.implies_intr_proof' (map snd (#constraints ucontext)) |> |
37310 | 406 |
fold_rev Proofterm.forall_intr_proof' xs' |> |
407 |
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
|
408 |
end; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
409 |
|
13402 | 410 |
(** expanding theorems / definitions **) |
411 |
||
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
412 |
fun add_expand_thm is_def thm thy = |
13402 | 413 |
let |
414 |
val {realizes_eqns, typeof_eqns, types, realizers, |
|
415 |
defs, expand, prep} = ExtractionData.get thy; |
|
416 |
||
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
417 |
val thm_name = Thm.derivation_name thm; |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
418 |
val _ = if Thm_Name.is_empty thm_name then error "add_expand_thm: unnamed theorem" else (); |
13402 | 419 |
in |
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
420 |
thy |> ExtractionData.put |
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
421 |
(if is_def then |
13402 | 422 |
{realizes_eqns = realizes_eqns, |
79411 | 423 |
typeof_eqns = typeof_eqns |
424 |
|> add_rule ([], Logic.dest_equals (Term.strip_sorts (Thm.prop_of (Drule.abs_def thm)))), |
|
13402 | 425 |
types = types, |
61094 | 426 |
realizers = realizers, defs = insert Thm.eq_thm_prop (Thm.trim_context thm) defs, |
13402 | 427 |
expand = expand, prep = prep} |
428 |
else |
|
429 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, |
|
430 |
realizers = realizers, defs = defs, |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
431 |
expand = insert (op =) thm_name expand, prep = prep}) |
13402 | 432 |
end; |
433 |
||
33704
6aeb8454efc1
add_expand_thm: explicit indication of is_def instead of fragile heuristic, tuned signature;
wenzelm
parents:
33522
diff
changeset
|
434 |
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
|
435 |
Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm is_def th) I); |
13402 | 436 |
|
15801 | 437 |
|
13732 | 438 |
(** types with computational content **) |
439 |
||
440 |
fun add_types tys thy = |
|
22717 | 441 |
ExtractionData.map |
442 |
(fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} => |
|
13732 | 443 |
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, |
22717 | 444 |
types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types, |
445 |
realizers = realizers, defs = defs, expand = expand, prep = prep}) |
|
446 |
thy; |
|
13732 | 447 |
|
13402 | 448 |
|
15801 | 449 |
(** Pure setup **) |
450 |
||
53171 | 451 |
val _ = Theory.setup |
18708 | 452 |
(add_types [("prop", ([], NONE))] #> |
15801 | 453 |
|
454 |
add_typeof_eqns |
|
67721 | 455 |
["(typeof (PROP P)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow> \ |
456 |
\ (typeof (PROP Q)) \<equiv> (Type (TYPE('Q))) \<Longrightarrow> \ |
|
457 |
\ (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE('Q)))", |
|
15801 | 458 |
|
67721 | 459 |
"(typeof (PROP Q)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow> \ |
460 |
\ (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE(Null)))", |
|
15801 | 461 |
|
67721 | 462 |
"(typeof (PROP P)) \<equiv> (Type (TYPE('P))) \<Longrightarrow> \ |
463 |
\ (typeof (PROP Q)) \<equiv> (Type (TYPE('Q))) \<Longrightarrow> \ |
|
464 |
\ (typeof (PROP P \<Longrightarrow> PROP Q)) \<equiv> (Type (TYPE('P \<Rightarrow> 'Q)))", |
|
15801 | 465 |
|
67721 | 466 |
"(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE(Null))) \<Longrightarrow> \ |
467 |
\ (typeof (\<And>x. PROP P (x))) \<equiv> (Type (TYPE(Null)))", |
|
15801 | 468 |
|
67721 | 469 |
"(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE('P))) \<Longrightarrow> \ |
470 |
\ (typeof (\<And>x::'a. PROP P (x))) \<equiv> (Type (TYPE('a \<Rightarrow> 'P)))", |
|
15801 | 471 |
|
67721 | 472 |
"(\<lambda>x. typeof (f (x))) \<equiv> (\<lambda>x. Type (TYPE('f))) \<Longrightarrow> \ |
473 |
\ (typeof (f)) \<equiv> (Type (TYPE('f)))"] #> |
|
15801 | 474 |
|
475 |
add_realizes_eqns |
|
67721 | 476 |
["(typeof (PROP P)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow> \ |
477 |
\ (realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv> \ |
|
478 |
\ (PROP realizes (Null) (PROP P) \<Longrightarrow> PROP realizes (r) (PROP Q))", |
|
15801 | 479 |
|
67721 | 480 |
"(typeof (PROP P)) \<equiv> (Type (TYPE('P))) \<Longrightarrow> \ |
481 |
\ (typeof (PROP Q)) \<equiv> (Type (TYPE(Null))) \<Longrightarrow> \ |
|
482 |
\ (realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv> \ |
|
483 |
\ (\<And>x::'P. PROP realizes (x) (PROP P) \<Longrightarrow> PROP realizes (Null) (PROP Q))", |
|
15801 | 484 |
|
67721 | 485 |
"(realizes (r) (PROP P \<Longrightarrow> PROP Q)) \<equiv> \ |
486 |
\ (\<And>x. PROP realizes (x) (PROP P) \<Longrightarrow> PROP realizes (r (x)) (PROP Q))", |
|
15801 | 487 |
|
67721 | 488 |
"(\<lambda>x. typeof (PROP P (x))) \<equiv> (\<lambda>x. Type (TYPE(Null))) \<Longrightarrow> \ |
489 |
\ (realizes (r) (\<And>x. PROP P (x))) \<equiv> \ |
|
490 |
\ (\<And>x. PROP realizes (Null) (PROP P (x)))", |
|
15801 | 491 |
|
67721 | 492 |
"(realizes (r) (\<And>x. PROP P (x))) \<equiv> \ |
493 |
\ (\<And>x. PROP realizes (r (x)) (PROP P (x)))"] #> |
|
15801 | 494 |
|
67147 | 495 |
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
|
496 |
"specify theorems to be expanded during extraction" #> |
67147 | 497 |
Attrib.setup \<^binding>\<open>extraction_expand_def\<close> (Scan.succeed (extraction_expand true)) |
53171 | 498 |
"specify definitions to be expanded during extraction"); |
15801 | 499 |
|
500 |
||
13402 | 501 |
(**** extract program ****) |
502 |
||
503 |
val dummyt = Const ("dummy", dummyT); |
|
504 |
||
58436 | 505 |
fun extract thm_vss thy = |
13402 | 506 |
let |
16458 | 507 |
val thy' = add_syntax thy; |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
508 |
val global_ctxt = Syntax.init_pretty_global thy'; |
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
509 |
val print_thm_name = Global_Theory.print_thm_name global_ctxt; |
62922 | 510 |
|
13402 | 511 |
val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = |
512 |
ExtractionData.get thy; |
|
22717 | 513 |
val procs = maps (rev o fst o snd) types; |
13732 | 514 |
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
|
515 |
val typroc = typeof_proc []; |
80590 | 516 |
fun expand_name ({thm_name = (thm_name, _), ...}: Proofterm.thm_header) = |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
517 |
if Thm_Name.is_empty thm_name orelse member (op =) expand thm_name |
80590 | 518 |
then SOME Thm_Name.none else NONE; |
61094 | 519 |
val prep = the_default (K I) prep thy' o |
70840 | 520 |
Proof_Rewrite_Rules.elim_defs thy' false (map (Thm.transfer thy) defs) o |
70915
bd4d37edfee4
clarified expand_proof/expand_name: allow more detailed control via thm_header;
wenzelm
parents:
70840
diff
changeset
|
521 |
Proofterm.expand_proof thy' expand_name; |
16800 | 522 |
val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); |
13402 | 523 |
|
524 |
fun find_inst prop Ts ts vs = |
|
525 |
let |
|
13732 | 526 |
val rvs = relevant_vars rtypes prop; |
13402 | 527 |
val vars = vars_of prop; |
528 |
val n = Int.min (length vars, length ts); |
|
529 |
||
33337 | 530 |
fun add_args (Var ((a, i), _), t) (vs', tye) = |
20664 | 531 |
if member (op =) rvs a then |
16458 | 532 |
let val T = etype_of thy' vs Ts t |
13402 | 533 |
in if T = nullT then (vs', tye) |
534 |
else (a :: vs', (("'" ^ a, i), T) :: tye) |
|
535 |
end |
|
536 |
else (vs', tye) |
|
537 |
||
33957 | 538 |
in fold_rev add_args (take n vars ~~ take n ts) ([], []) end; |
13402 | 539 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
540 |
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
|
541 |
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
|
542 |
|
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
543 |
fun mk_sprfs cs tye = maps (fn (_, T) => |
70840 | 544 |
Proof_Rewrite_Rules.expand_of_sort_proof thy' (map SOME cs) |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
545 |
(T, Sign.defaultS thy)) tye; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
546 |
|
33038 | 547 |
fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
548 |
fun filter_thm_name (a: Thm_Name.T) = |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
549 |
map_filter (fn (b, x) => if a = b then SOME x else NONE); |
13402 | 550 |
|
46219
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
44057
diff
changeset
|
551 |
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
|
552 |
strip_abs (length Ts) |
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents:
44057
diff
changeset
|
553 |
(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
|
554 |
(fold (Term.abs o pair "x") Ts t)); |
13732 | 555 |
|
556 |
fun realizes_null vs prop = app_rlz_rews [] vs |
|
557 |
(Const ("realizes", nullT --> propT --> propT) $ nullt $ prop); |
|
13402 | 558 |
|
49960 | 559 |
fun corr d vs ts Ts hs cs _ (PBound i) _ defs = (PBound i, defs) |
13402 | 560 |
|
49960 | 561 |
| corr d vs ts Ts hs cs t (Abst (s, SOME T, prf)) (Abst (_, _, prf')) defs = |
562 |
let val (corr_prf, defs') = corr d vs [] (T :: Ts) |
|
563 |
(dummyt :: hs) cs (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE) |
|
79175
04dfecb9343a
misc tuning and clarification, following Term.incr_bv / Term.incr_boundvars;
wenzelm
parents:
79167
diff
changeset
|
564 |
prf (Proofterm.incr_boundvars 1 0 prf') defs |
49960 | 565 |
in (Abst (s, SOME T, corr_prf), defs') end |
13402 | 566 |
|
49960 | 567 |
| corr d vs ts Ts hs cs t (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) defs = |
13402 | 568 |
let |
16458 | 569 |
val T = etype_of thy' vs Ts prop; |
42407
5b9dd52f5dca
prefer internal types, via Simple_Syntax.read_typ;
wenzelm
parents:
42406
diff
changeset
|
570 |
val u = if T = nullT then |
15531 | 571 |
(case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE) |
572 |
else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE); |
|
49960 | 573 |
val (corr_prf, defs') = |
574 |
corr d vs [] (T :: Ts) (prop :: hs) |
|
79175
04dfecb9343a
misc tuning and clarification, following Term.incr_bv / Term.incr_boundvars;
wenzelm
parents:
79167
diff
changeset
|
575 |
(prop :: cs) u (Proofterm.incr_boundvars 0 1 prf) |
04dfecb9343a
misc tuning and clarification, following Term.incr_bv / Term.incr_boundvars;
wenzelm
parents:
79167
diff
changeset
|
576 |
(Proofterm.incr_boundvars 0 1 prf') defs; |
13402 | 577 |
val rlz = Const ("realizes", T --> propT --> propT) |
49960 | 578 |
in ( |
13732 | 579 |
if T = nullT then AbsP ("R", |
15531 | 580 |
SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)), |
79167 | 581 |
Proofterm.subst_bounds [nullt] corr_prf) |
15531 | 582 |
else Abst (s, SOME T, AbsP ("R", |
583 |
SOME (app_rlz_rews (T :: Ts) vs |
|
49960 | 584 |
(rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)), defs') |
13402 | 585 |
end |
586 |
||
49960 | 587 |
| corr d vs ts Ts hs cs t' (prf % SOME t) (prf' % _) defs = |
13732 | 588 |
let |
589 |
val (Us, T) = strip_type (fastype_of1 (Ts, t)); |
|
49960 | 590 |
val (corr_prf, defs') = corr d vs (t :: ts) Ts hs cs |
20664 | 591 |
(if member (op =) rtypes (tname_of T) then t' |
49960 | 592 |
else (case t' of SOME (u $ _) => SOME u | _ => NONE)) |
593 |
prf prf' defs; |
|
20664 | 594 |
val u = if not (member (op =) rtypes (tname_of T)) then t else |
13732 | 595 |
let |
16458 | 596 |
val eT = etype_of thy' vs Ts t; |
13732 | 597 |
val (r, Us') = if eT = nullT then (nullt, Us) else |
598 |
(Bound (length Us), eT :: Us); |
|
599 |
val u = list_comb (incr_boundvars (length Us') t, |
|
600 |
map Bound (length Us - 1 downto 0)); |
|
17271 | 601 |
val u' = (case AList.lookup (op =) types (tname_of T) of |
15531 | 602 |
SOME ((_, SOME f)) => f r eT u T |
13732 | 603 |
| _ => 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
|
604 |
in app_rlz_rews Ts vs (fold_rev (Term.abs o pair "x") Us' u') end |
49960 | 605 |
in (corr_prf % SOME u, defs') end |
13402 | 606 |
|
49960 | 607 |
| corr d vs ts Ts hs cs t (prf1 %% prf2) (prf1' %% prf2') defs = |
13402 | 608 |
let |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
609 |
val prop = Proofterm.prop_of' hs prf2'; |
16458 | 610 |
val T = etype_of thy' vs Ts prop; |
49960 | 611 |
val (f, u, defs1) = if T = nullT then (t, NONE, defs) else |
13402 | 612 |
(case t of |
49960 | 613 |
SOME (f $ u) => (SOME f, SOME u, defs) |
13402 | 614 |
| _ => |
49960 | 615 |
let val (u, defs1) = extr d vs [] Ts hs prf2' defs |
616 |
in (NONE, SOME u, defs1) end) |
|
617 |
val ((corr_prf1, corr_prf2), defs2) = |
|
618 |
defs1 |
|
619 |
|> corr d vs [] Ts hs cs f prf1 prf1' |
|
620 |
||>> corr d vs [] Ts hs cs u prf2 prf2'; |
|
13402 | 621 |
in |
49960 | 622 |
if T = nullT then (corr_prf1 %% corr_prf2, defs2) else |
623 |
(corr_prf1 % u %% corr_prf2, defs2) |
|
13402 | 624 |
end |
625 |
||
70494 | 626 |
| corr d vs ts Ts hs cs _ (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) _ defs = |
13402 | 627 |
let |
80590 | 628 |
val {command_pos, theory_name, thm_name = (thm_name, thm_pos), prop, ...} = thm_header; |
70493 | 629 |
val prf = Proofterm.thm_body_proof_open thm_body; |
13402 | 630 |
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
|
631 |
val shyps = mk_shyps tye; |
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
632 |
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
|
633 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye; |
16458 | 634 |
val T = etype_of thy' vs' [] prop; |
13402 | 635 |
val defs' = if T = nullT then defs |
49960 | 636 |
else snd (extr d vs ts Ts hs prf0 defs) |
13402 | 637 |
in |
49960 | 638 |
if T = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs) |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
639 |
else (case Thm_Name.Table.lookup realizers thm_name of |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
640 |
NONE => (case find vs' (filter_thm_name thm_name defs') of |
15531 | 641 |
NONE => |
13402 | 642 |
let |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
643 |
val _ = T = nullT orelse error "corr: internal error"; |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
644 |
val _ = |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
645 |
msg d ("Building correctness proof for " ^ quote (print_thm_name thm_name) ^ |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
646 |
(if null vs' then "" |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
647 |
else " (relevant variables: " ^ commas_quote vs' ^ ")")); |
70449 | 648 |
val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); |
49960 | 649 |
val (corr_prf0, defs'') = corr (d + 1) vs' [] [] [] |
650 |
(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
|
651 |
val corr_prf = mkabsp shyps corr_prf0; |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
652 |
val corr_prop = Proofterm.prop_of corr_prf; |
70493 | 653 |
val corr_header = |
80590 | 654 |
Proofterm.thm_header (serial ()) command_pos theory_name |
655 |
((corr_name thm_name vs', 0), thm_pos) corr_prop |
|
70538
fc9ba6fe367f
clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents:
70501
diff
changeset
|
656 |
(SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
657 |
val corr_prf' = |
70493 | 658 |
Proofterm.proof_combP |
659 |
(Proofterm.proof_combt |
|
70589 | 660 |
(PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), |
70493 | 661 |
map PBound (length shyps - 1 downto 0)) |> |
37310 | 662 |
fold_rev Proofterm.forall_intr_proof' |
663 |
(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
|
664 |
mkabsp shyps |
13402 | 665 |
in |
49960 | 666 |
(Proofterm.proof_combP (prf_subst_TVars tye' corr_prf', sprfs), |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
667 |
(thm_name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'') |
13402 | 668 |
end |
37310 | 669 |
| SOME (_, (_, prf')) => |
49960 | 670 |
(Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs')) |
15531 | 671 |
| SOME rs => (case find vs' rs of |
49960 | 672 |
SOME (_, prf') => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs') |
15531 | 673 |
| NONE => error ("corr: no realizer for instance of theorem " ^ |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
674 |
quote (print_thm_name thm_name) ^ ":\n" ^ |
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
675 |
Syntax.string_of_term global_ctxt |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
676 |
(Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))))) |
13402 | 677 |
end |
678 |
||
49960 | 679 |
| corr d vs ts Ts hs cs _ (prf0 as PAxm (s, prop, SOME Ts')) _ defs = |
13402 | 680 |
let |
681 |
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
|
682 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 683 |
in |
16458 | 684 |
if etype_of thy' vs' [] prop = nullT andalso |
49960 | 685 |
realizes_null vs' prop aconv prop then (prf0, defs) |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
686 |
else case find vs' (Thm_Name.Table.lookup_list realizers (s, 0)) of |
49960 | 687 |
SOME (_, prf) => (Proofterm.proof_combP (prf_subst_TVars tye' prf, mk_sprfs cs tye), |
688 |
defs) |
|
15531 | 689 |
| NONE => error ("corr: no realizer for instance of axiom " ^ |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
690 |
quote s ^ ":\n" ^ Syntax.string_of_term global_ctxt (Envir.beta_norm |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
691 |
(Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) |
13402 | 692 |
end |
693 |
||
49960 | 694 |
| corr d vs ts Ts hs _ _ _ _ defs = error "corr: bad proof" |
13402 | 695 |
|
49960 | 696 |
and extr d vs ts Ts hs (PBound i) defs = (Bound i, defs) |
13402 | 697 |
|
49960 | 698 |
| extr d vs ts Ts hs (Abst (s, SOME T, prf)) defs = |
699 |
let val (t, defs') = extr d vs [] |
|
79175
04dfecb9343a
misc tuning and clarification, following Term.incr_bv / Term.incr_boundvars;
wenzelm
parents:
79167
diff
changeset
|
700 |
(T :: Ts) (dummyt :: hs) (Proofterm.incr_boundvars 1 0 prf) defs |
49960 | 701 |
in (Abs (s, T, t), defs') end |
13402 | 702 |
|
49960 | 703 |
| extr d vs ts Ts hs (AbsP (s, SOME t, prf)) defs = |
13402 | 704 |
let |
16458 | 705 |
val T = etype_of thy' vs Ts t; |
49960 | 706 |
val (t, defs') = |
79175
04dfecb9343a
misc tuning and clarification, following Term.incr_bv / Term.incr_boundvars;
wenzelm
parents:
79167
diff
changeset
|
707 |
extr d vs [] (T :: Ts) (t :: hs) (Proofterm.incr_boundvars 0 1 prf) defs |
49960 | 708 |
in |
709 |
(if T = nullT then subst_bound (nullt, t) else Abs (s, T, t), defs') |
|
13402 | 710 |
end |
711 |
||
49960 | 712 |
| extr d vs ts Ts hs (prf % SOME t) defs = |
713 |
let val (u, defs') = extr d vs (t :: ts) Ts hs prf defs |
|
714 |
in (if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u |
|
715 |
else u $ t, defs') |
|
13732 | 716 |
end |
13402 | 717 |
|
49960 | 718 |
| extr d vs ts Ts hs (prf1 %% prf2) defs = |
13402 | 719 |
let |
49960 | 720 |
val (f, defs') = extr d vs [] Ts hs prf1 defs; |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
721 |
val prop = Proofterm.prop_of' hs prf2; |
16458 | 722 |
val T = etype_of thy' vs Ts prop |
13402 | 723 |
in |
49960 | 724 |
if T = nullT then (f, defs') else |
725 |
let val (t, defs'') = extr d vs [] Ts hs prf2 defs' |
|
726 |
in (f $ t, defs'') end |
|
13402 | 727 |
end |
728 |
||
70494 | 729 |
| extr d vs ts Ts hs (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) defs = |
13402 | 730 |
let |
80590 | 731 |
val {command_pos, theory_name, thm_name = (thm_name, thm_pos), prop, ...} = thm_header; |
70493 | 732 |
val prf = Proofterm.thm_body_proof_open thm_body; |
13402 | 733 |
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
|
734 |
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
|
735 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 736 |
in |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
737 |
case Thm_Name.Table.lookup realizers thm_name of |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
738 |
NONE => (case find vs' (filter_thm_name thm_name defs) of |
15531 | 739 |
NONE => |
13402 | 740 |
let |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
741 |
val _ = |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
742 |
msg d ("Extracting " ^ quote (print_thm_name thm_name) ^ |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
743 |
(if null vs' then "" |
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
744 |
else " (relevant variables: " ^ commas_quote vs' ^ ")")); |
70449 | 745 |
val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); |
49960 | 746 |
val (t, defs') = extr (d + 1) vs' [] [] [] prf' defs; |
747 |
val (corr_prf, defs'') = corr (d + 1) vs' [] [] [] |
|
748 |
(rev shyps) (SOME t) prf' prf' defs'; |
|
13402 | 749 |
|
750 |
val nt = Envir.beta_norm t; |
|
20664 | 751 |
val args = filter_out (fn v => member (op =) rtypes |
752 |
(tname_of (body_type (fastype_of v)))) (vfs_of prop); |
|
33317 | 753 |
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
|
754 |
val t' = mkabs args' nt; |
13402 | 755 |
val T = fastype_of t'; |
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
756 |
val cname = extr_name thm_name vs'; |
13402 | 757 |
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
|
758 |
val u = mkabs args (list_comb (c, args')); |
13402 | 759 |
val eqn = Logic.mk_equals (c, t'); |
760 |
val rlz = |
|
761 |
Const ("realizes", fastype_of nt --> propT --> propT); |
|
13732 | 762 |
val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop); |
763 |
val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop); |
|
764 |
val f = app_rlz_rews [] vs' |
|
765 |
(Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop)); |
|
13402 | 766 |
|
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
767 |
val corr_prf' = mkabsp shyps |
70806 | 768 |
(change_types [] Proofterm.equal_elim_axm %> lhs %> rhs %% |
769 |
(change_types [propT] Proofterm.symmetric_axm %> rhs %> lhs %% |
|
770 |
(change_types [T, propT] Proofterm.combination_axm %> f %> f %> c %> t' %% |
|
771 |
(change_types [T --> propT] Proofterm.reflexive_axm %> f) %% |
|
46909 | 772 |
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
|
773 |
SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf); |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
774 |
val corr_prop = Proofterm.prop_of corr_prf'; |
70493 | 775 |
val corr_header = |
80590 | 776 |
Proofterm.thm_header (serial ()) command_pos theory_name |
777 |
((corr_name thm_name vs', 0), thm_pos) corr_prop |
|
70538
fc9ba6fe367f
clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents:
70501
diff
changeset
|
778 |
(SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
779 |
val corr_prf'' = |
37310 | 780 |
Proofterm.proof_combP (Proofterm.proof_combt |
70589 | 781 |
(PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), |
37233
b78f31ca4675
Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents:
36953
diff
changeset
|
782 |
map PBound (length shyps - 1 downto 0)) |> |
37310 | 783 |
fold_rev Proofterm.forall_intr_proof' |
784 |
(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
|
785 |
mkabsp shyps |
13402 | 786 |
in |
49960 | 787 |
(subst_TVars tye' u, |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
788 |
(thm_name, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'') |
13402 | 789 |
end |
49960 | 790 |
| SOME ((_, u), _) => (subst_TVars tye' u, defs)) |
15531 | 791 |
| SOME rs => (case find vs' rs of |
49960 | 792 |
SOME (t, _) => (subst_TVars tye' t, defs) |
15531 | 793 |
| NONE => error ("extr: no realizer for instance of theorem " ^ |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
794 |
quote (print_thm_name thm_name) ^ ":\n" ^ |
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
795 |
Syntax.string_of_term global_ctxt (Envir.beta_norm |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
796 |
(Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))) |
13402 | 797 |
end |
798 |
||
49960 | 799 |
| extr d vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) defs = |
13402 | 800 |
let |
801 |
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
|
802 |
val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye |
13402 | 803 |
in |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
804 |
case find vs' (Thm_Name.Table.lookup_list realizers (s, 0)) of |
49960 | 805 |
SOME (t, _) => (subst_TVars tye' t, defs) |
15531 | 806 |
| NONE => error ("extr: no realizer for instance of axiom " ^ |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
807 |
quote s ^ ":\n" ^ Syntax.string_of_term global_ctxt (Envir.beta_norm |
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
808 |
(Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) |
13402 | 809 |
end |
810 |
||
49960 | 811 |
| extr d vs ts Ts hs _ defs = error "extr: bad proof"; |
13402 | 812 |
|
60826 | 813 |
fun prep_thm vs raw_thm = |
13402 | 814 |
let |
60826 | 815 |
val thm = Thm.transfer thy raw_thm; |
26626
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents:
26481
diff
changeset
|
816 |
val prop = Thm.prop_of thm; |
28814 | 817 |
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
|
818 |
val name = Thm.derivation_name thm; |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
819 |
val _ = if Thm_Name.is_empty name then error "extraction: unnamed theorem" else (); |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21646
diff
changeset
|
820 |
val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^ |
81534
c32ebdcbe8ca
proper context for extern operation: observe local options;
wenzelm
parents:
80590
diff
changeset
|
821 |
quote (print_thm_name name) ^ " has no computational content") |
70449 | 822 |
in Proofterm.reconstruct_proof thy' prop prf end; |
13402 | 823 |
|
33245 | 824 |
val defs = |
58436 | 825 |
fold (fn (thm, vs) => snd o (extr 0 vs [] [] [] o prep_thm vs) thm) |
826 |
thm_vss []; |
|
13402 | 827 |
|
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
828 |
fun add_def (name, (vs, ((t, u)))) thy = |
66145 | 829 |
let |
830 |
val ft = Type.legacy_freeze t; |
|
831 |
val fu = Type.legacy_freeze u; |
|
66146
fd3e128b174f
register equations stemming from extracted proofs as specification rules
haftmann
parents:
66145
diff
changeset
|
832 |
val head = head_of (strip_abs_body fu); |
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
833 |
val b = Binding.qualified_name (extr_name name vs); |
66145 | 834 |
in |
835 |
thy |
|
71179 | 836 |
|> Sign.add_consts [(b, fastype_of ft, NoSyn)] |
79336
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents:
79279
diff
changeset
|
837 |
|> Global_Theory.add_def |
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
838 |
(Binding.qualified_name (Thm.def_name (extr_name name vs)), |
80295
8a9588ffc133
more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents:
79411
diff
changeset
|
839 |
Logic.mk_equals (head, ft)) |
79336
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents:
79279
diff
changeset
|
840 |
|-> (fn def_thm => |
71214
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
wenzelm
parents:
71206
diff
changeset
|
841 |
Spec_Rules.add_global b Spec_Rules.equational |
71206
20dce31fe7f4
proper spec rules via resulting def_thm, e.g. relevant for "isabelle build -o export_theory";
wenzelm
parents:
71179
diff
changeset
|
842 |
[Thm.term_of (Thm.lhs_of def_thm)] [def_thm] |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66146
diff
changeset
|
843 |
#> Code.declare_default_eqns_global [(def_thm, true)]) |
66145 | 844 |
end; |
845 |
||
846 |
fun add_corr (s, (vs, prf)) thy = |
|
847 |
let |
|
70447
755d58b48cec
clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents:
70436
diff
changeset
|
848 |
val corr_prop = Proofterm.prop_of prf; |
66145 | 849 |
in |
850 |
thy |
|
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
851 |
|> Global_Theory.store_thm (Binding.qualified_name (corr_name s vs), |
66145 | 852 |
Thm.varifyT_global (funpow (length (vars_of corr_prop)) |
853 |
(Thm.forall_elim_var 0) (Thm.forall_intr_frees |
|
854 |
(Proof_Checker.thm_of_proof thy |
|
855 |
(fst (Proofterm.freeze_thaw_prf prf)))))) |
|
856 |
|> snd |
|
857 |
end; |
|
858 |
||
859 |
fun add_def_and_corr (s, (vs, ((t, u), (prf, _)))) thy = |
|
80310
6d091c0c252e
clarified treatment of Thm_Name.T (again, see also 8a9588ffc133);
wenzelm
parents:
80299
diff
changeset
|
860 |
if is_none (Sign.const_type thy (extr_name s vs)) |
66145 | 861 |
then |
862 |
thy |
|
863 |
|> not (t = nullt) ? add_def (s, (vs, ((t, u)))) |
|
864 |
|> add_corr (s, (vs, prf)) |
|
865 |
else |
|
866 |
thy; |
|
13402 | 867 |
|
16149 | 868 |
in |
869 |
thy |
|
30435
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents:
30364
diff
changeset
|
870 |
|> Sign.root_path |
66145 | 871 |
|> fold_rev add_def_and_corr defs |
22796 | 872 |
|> Sign.restore_naming thy |
13402 | 873 |
end; |
874 |
||
16458 | 875 |
val etype_of = etype_of o add_syntax; |
13714 | 876 |
|
13402 | 877 |
end; |