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