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