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